Definition: pi_prefix
pi_prefix() ==  lbl:Atom × if lbl =a "send" then chan:Name × Name if lbl =a "rcv" then chan:Name × Name else Void fi 

Lemma: pi_prefix_wf
pi_prefix() ∈ Type

Lemma: pi_prefix-ext
pi_prefix() ≡ lbl:Atom × if lbl =a "send" then chan:Name × Name
                         if lbl =a "rcv" then chan:Name × Name
                         else Void
                         fi 

Definition: pisend
pisend(chan;var) ==  <"send", chan, var>

Lemma: pisend_wf
[chan,var:Name].  (pisend(chan;var) ∈ pi_prefix())

Definition: pircv
pircv(chan;var) ==  <"rcv", chan, var>

Lemma: pircv_wf
[chan,var:Name].  (pircv(chan;var) ∈ pi_prefix())

Definition: pisend?
pisend?(v) ==  fst(v) =a "send"

Lemma: pisend?_wf
[v:pi_prefix()]. (pisend?(v) ∈ 𝔹)

Definition: pisend-chan
pisend-chan(v) ==  fst(snd(v))

Lemma: pisend-chan_wf
[v:pi_prefix()]. pisend-chan(v) ∈ Name supposing ↑pisend?(v)

Definition: pisend-var
pisend-var(v) ==  snd(snd(v))

Lemma: pisend-var_wf
[v:pi_prefix()]. pisend-var(v) ∈ Name supposing ↑pisend?(v)

Definition: pircv?
pircv?(v) ==  fst(v) =a "rcv"

Lemma: pircv?_wf
[v:pi_prefix()]. (pircv?(v) ∈ 𝔹)

Definition: pircv-chan
pircv-chan(v) ==  fst(snd(v))

Lemma: pircv-chan_wf
[v:pi_prefix()]. pircv-chan(v) ∈ Name supposing ↑pircv?(v)

Definition: pircv-var
pircv-var(v) ==  snd(snd(v))

Lemma: pircv-var_wf
[v:pi_prefix()]. pircv-var(v) ∈ Name supposing ↑pircv?(v)

Lemma: pi_prefix-induction
[P:pi_prefix() ⟶ ℙ]
  ((∀chan,var:Name.  P[pisend(chan;var)])  (∀chan,var:Name.  P[pircv(chan;var)])  {∀v:pi_prefix(). P[v]})

Lemma: pi_prefix-definition
[A:Type]. ∀[R:A ⟶ pi_prefix() ⟶ ℙ].
  ((∀chan,var:Name.  {x:A| R[x;pisend(chan;var)]} )
   (∀chan,var:Name.  {x:A| R[x;pircv(chan;var)]} )
   {∀v:pi_prefix(). {x:A| R[x;v]} })

Definition: pi_prefix_ind
pi_prefix_ind(v;
              pisend(chan,var) send[chan; var];
              pircv(chan,var) rcv[chan; var])  ==
  let lbl,v1 
  in if lbl="send" then let chan,v2 v1 in send[chan; v2]
     if lbl="rcv" then let chan,v2 v1 in rcv[chan; v2]
     else Ax
     fi 

Lemma: pi_prefix_ind_wf
[A:Type]. ∀[R:A ⟶ pi_prefix() ⟶ ℙ]. ∀[v:pi_prefix()]. ∀[send:chan:Name ⟶ var:Name ⟶ {x:A| R[x;pisend(chan;var)]} ].
[rcv:chan:Name ⟶ var:Name ⟶ {x:A| R[x;pircv(chan;var)]} ].
  (pi_prefix_ind(v;
                 pisend(chan,var) send[chan;var];
                 pircv(chan,var) rcv[chan;var])  ∈ {x:A| R[x;v]} )

Lemma: pi_prefix_ind_wf_simple
[A:Type]. ∀[v:pi_prefix()]. ∀[send,rcv:chan:Name ⟶ var:Name ⟶ A].
  (pi_prefix_ind(v;
                 pisend(chan,var) send[chan;var];
                 pircv(chan,var) rcv[chan;var])  ∈ A)

Definition: pi_termco
pi_termco() ==
  corec(X.lbl:Atom × if lbl =a "zero" then Unit
                     if lbl =a "comm" then pre:pi_prefix() × X
                     if lbl =a "option" then left:X × X
                     if lbl =a "par" then left:X × X
                     if lbl =a "rep" then X
                     if lbl =a "new" then name:Name × X
                     else Void
                     fi )

Lemma: pi_termco_wf
pi_termco() ∈ Type

Lemma: pi_termco-ext
pi_termco() ≡ lbl:Atom × if lbl =a "zero" then Unit
                         if lbl =a "comm" then pre:pi_prefix() × pi_termco()
                         if lbl =a "option" then left:pi_termco() × pi_termco()
                         if lbl =a "par" then left:pi_termco() × pi_termco()
                         if lbl =a "rep" then pi_termco()
                         if lbl =a "new" then name:Name × pi_termco()
                         else Void
                         fi 

Definition: pi_termco_size
pi_termco_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "zero" then 0
                   if lbl =a "comm" then let pre,z in (size z)
                   if lbl =a "option" then let left,right in (1 (size left)) (size right)
                   if lbl =a "par" then let left,right in (1 (size left)) (size right)
                   if lbl =a "rep" then (size x)
                   if lbl =a "new" then let name,z in (size z)
                   else 0
                   fi )) 
  p

Lemma: pi_termco_size_wf
[p:pi_termco()]. (pi_termco_size(p) ∈ partial(ℕ))

Definition: pi_term
pi_term() ==  {p:pi_termco()| (pi_termco_size(p))↓

Lemma: pi_term_wf
pi_term() ∈ Type

Definition: pi_term_size
pi_term_size(p) ==
  fix((λsize,p. let lbl,x 
                in if lbl =a "zero" then 0
                   if lbl =a "comm" then let pre,z in (size z)
                   if lbl =a "option" then let left,right in (1 (size left)) (size right)
                   if lbl =a "par" then let left,right in (1 (size left)) (size right)
                   if lbl =a "rep" then (size x)
                   if lbl =a "new" then let name,z in (size z)
                   else 0
                   fi )) 
  p

Lemma: pi_term_size_wf
[p:pi_term()]. (pi_term_size(p) ∈ ℕ)

Lemma: pi_term-ext
pi_term() ≡ lbl:Atom × if lbl =a "zero" then Unit
                       if lbl =a "comm" then pre:pi_prefix() × pi_term()
                       if lbl =a "option" then left:pi_term() × pi_term()
                       if lbl =a "par" then left:pi_term() × pi_term()
                       if lbl =a "rep" then pi_term()
                       if lbl =a "new" then name:Name × pi_term()
                       else Void
                       fi 

Definition: pizero
pizero() ==  <"zero", ⋅>

Lemma: pizero_wf
pizero() ∈ pi_term()

Definition: picomm
picomm(pre;body) ==  <"comm", pre, body>

Lemma: picomm_wf
[pre:pi_prefix()]. ∀[body:pi_term()].  (picomm(pre;body) ∈ pi_term())

Definition: pioption
pioption(left;right) ==  <"option", left, right>

Lemma: pioption_wf
[left,right:pi_term()].  (pioption(left;right) ∈ pi_term())

Definition: pipar
pipar(left;right) ==  <"par", left, right>

Lemma: pipar_wf
[left,right:pi_term()].  (pipar(left;right) ∈ pi_term())

Definition: pirep
pirep(body) ==  <"rep", body>

Lemma: pirep_wf
[body:pi_term()]. (pirep(body) ∈ pi_term())

Definition: pinew
pinew(name;body) ==  <"new", name, body>

Lemma: pinew_wf
[name:Name]. ∀[body:pi_term()].  (pinew(name;body) ∈ pi_term())

Definition: pizero?
pizero?(v) ==  fst(v) =a "zero"

Lemma: pizero?_wf
[v:pi_term()]. (pizero?(v) ∈ 𝔹)

Definition: picomm?
picomm?(v) ==  fst(v) =a "comm"

Lemma: picomm?_wf
[v:pi_term()]. (picomm?(v) ∈ 𝔹)

Definition: picomm-pre
picomm-pre(v) ==  fst(snd(v))

Lemma: picomm-pre_wf
[v:pi_term()]. picomm-pre(v) ∈ pi_prefix() supposing ↑picomm?(v)

Definition: picomm-body
picomm-body(v) ==  snd(snd(v))

Lemma: picomm-body_wf
[v:pi_term()]. picomm-body(v) ∈ pi_term() supposing ↑picomm?(v)

Definition: pioption?
pioption?(v) ==  fst(v) =a "option"

Lemma: pioption?_wf
[v:pi_term()]. (pioption?(v) ∈ 𝔹)

Definition: pioption-left
pioption-left(v) ==  fst(snd(v))

Lemma: pioption-left_wf
[v:pi_term()]. pioption-left(v) ∈ pi_term() supposing ↑pioption?(v)

Definition: pioption-right
pioption-right(v) ==  snd(snd(v))

Lemma: pioption-right_wf
[v:pi_term()]. pioption-right(v) ∈ pi_term() supposing ↑pioption?(v)

Definition: pipar?
pipar?(v) ==  fst(v) =a "par"

Lemma: pipar?_wf
[v:pi_term()]. (pipar?(v) ∈ 𝔹)

Definition: pipar-left
pipar-left(v) ==  fst(snd(v))

Lemma: pipar-left_wf
[v:pi_term()]. pipar-left(v) ∈ pi_term() supposing ↑pipar?(v)

Definition: pipar-right
pipar-right(v) ==  snd(snd(v))

Lemma: pipar-right_wf
[v:pi_term()]. pipar-right(v) ∈ pi_term() supposing ↑pipar?(v)

Definition: pirep?
pirep?(v) ==  fst(v) =a "rep"

Lemma: pirep?_wf
[v:pi_term()]. (pirep?(v) ∈ 𝔹)

Definition: pirep-body
pirep-body(v) ==  snd(v)

Lemma: pirep-body_wf
[v:pi_term()]. pirep-body(v) ∈ pi_term() supposing ↑pirep?(v)

Definition: pinew?
pinew?(v) ==  fst(v) =a "new"

Lemma: pinew?_wf
[v:pi_term()]. (pinew?(v) ∈ 𝔹)

Definition: pinew-name
pinew-name(v) ==  fst(snd(v))

Lemma: pinew-name_wf
[v:pi_term()]. pinew-name(v) ∈ Name supposing ↑pinew?(v)

Definition: pinew-body
pinew-body(v) ==  snd(snd(v))

Lemma: pinew-body_wf
[v:pi_term()]. pinew-body(v) ∈ pi_term() supposing ↑pinew?(v)

Lemma: pi_term-induction
[P:pi_term() ⟶ ℙ]
  (P[pizero()]
   (∀pre:pi_prefix(). ∀body:pi_term().  (P[body]  P[picomm(pre;body)]))
   (∀left,right:pi_term().  (P[left]  P[right]  P[pioption(left;right)]))
   (∀left,right:pi_term().  (P[left]  P[right]  P[pipar(left;right)]))
   (∀body:pi_term(). (P[body]  P[pirep(body)]))
   (∀name:Name. ∀body:pi_term().  (P[body]  P[pinew(name;body)]))
   {∀v:pi_term(). P[v]})

Lemma: pi_term-definition
[A:Type]. ∀[R:A ⟶ pi_term() ⟶ ℙ].
  ({x:A| R[x;pizero()]} 
   (∀pre:pi_prefix(). ∀body:pi_term().  ({x:A| R[x;body]}   {x:A| R[x;picomm(pre;body)]} ))
   (∀left,right:pi_term().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;pioption(left;right)]} ))
   (∀left,right:pi_term().  ({x:A| R[x;left]}   {x:A| R[x;right]}   {x:A| R[x;pipar(left;right)]} ))
   (∀body:pi_term(). ({x:A| R[x;body]}   {x:A| R[x;pirep(body)]} ))
   (∀name:Name. ∀body:pi_term().  ({x:A| R[x;body]}   {x:A| R[x;pinew(name;body)]} ))
   {∀v:pi_term(). {x:A| R[x;v]} })

Definition: pi_term_ind
... ==
  fix((λpi_term_ind,v. let lbl,v1 
                       in if lbl="zero" then zero
                          if lbl="comm" then let pre,v2 v1 in comm[pre; v2; pi_term_ind v2]
                          if lbl="option" then let left,v2 v1 in option[left; v2; pi_term_ind left; pi_term_ind v2]
                          if lbl="par" then let left,v2 v1 in par[left; v2; pi_term_ind left; pi_term_ind v2]
                          if lbl="rep" then rep[v1; pi_term_ind v1]
                          if lbl="new" then let name,v2 v1 in new[name; v2; pi_term_ind v2]
                          else Ax
                          fi )) 
  v

Lemma: pi_term_ind_wf
[A:Type]. ∀[R:A ⟶ pi_term() ⟶ ℙ]. ∀[v:pi_term()]. ∀[zero:{x:A| R[x;pizero()]} ]. ∀[comm:pre:pi_prefix()
                                                                                          ⟶ body:pi_term()
                                                                                          ⟶ {x:A| R[x;body]} 
                                                                                          ⟶ {x:A| 
                                                                                              R[x;picomm(pre;body)]} ].
[option:left:pi_term()
         ⟶ right:pi_term()
         ⟶ {x:A| R[x;left]} 
         ⟶ {x:A| R[x;right]} 
         ⟶ {x:A| R[x;pioption(left;right)]} ]. ∀[par:left:pi_term()
                                                     ⟶ right:pi_term()
                                                     ⟶ {x:A| R[x;left]} 
                                                     ⟶ {x:A| R[x;right]} 
                                                     ⟶ {x:A| R[x;pipar(left;right)]} ]. ∀[rep:body:pi_term()
                                                                                              ⟶ {x:A| R[x;body]} 
                                                                                              ⟶ {x:A| 
                                                                                                  R[x;pirep(body)]} ].
[new:name:Name ⟶ body:pi_term() ⟶ {x:A| R[x;body]}  ⟶ {x:A| R[x;pinew(name;body)]} ].
  (... ∈ {x:A| R[x;v]} )

Lemma: pi_term_ind_wf_simple
[A:Type]. ∀[v:pi_term()]. ∀[zero:A]. ∀[comm:pre:pi_prefix() ⟶ body:pi_term() ⟶ A ⟶ A].
[option,par:left:pi_term() ⟶ right:pi_term() ⟶ A ⟶ A ⟶ A]. ∀[rep:body:pi_term() ⟶ A ⟶ A]. ∀[new:name:Name
                                                                                                       ⟶ body:pi_term()
                                                                                                       ⟶ A
                                                                                                       ⟶ A].
  (... ∈ A)

Definition: prefix-match
prefix-match(p1;p2) ==
  if pisend?(p1) ∧b pircv?(p2) then name_eq(pisend-chan(p1);pircv-chan(p2))
  if pisend?(p1) ∧b pircv?(p2) then name_eq(pisend-chan(p1);pircv-chan(p2))
  else ff
  fi 

Lemma: prefix-match_wf
[p1,p2:pi_prefix()].  (prefix-match(p1;p2) ∈ 𝔹)

Lemma: pi-comm-decompose
[P:pi_term()]. picomm(picomm-pre(P);picomm-body(P)) ∈ pi_term() supposing ↑picomm?(P)

Lemma: pi-option-decompose
[P:pi_term()]. pioption(pioption-left(P);pioption-right(P)) ∈ pi_term() supposing ↑pioption?(P)

Lemma: pi-par-decompose
[P:pi_term()]. pipar(pipar-left(P);pipar-right(P)) ∈ pi_term() supposing ↑pipar?(P)

Lemma: pi-new-decompose
[P:pi_term()]. pinew(pinew-name(P);pinew-body(P)) ∈ pi_term() supposing ↑pinew?(P)

Lemma: pi-rep-decompose
[P:pi_term()]. pirep(pirep-body(P)) ∈ pi_term() supposing ↑pirep?(P)

Definition: pi-example
pi-example(x;y;z;v) ==
  pipar(pinew(x;pipar(picomm(pisend(x;z);pizero());picomm(pircv(x;y);picomm(pisend(y;x);picomm(pircv(x;y);...)))));...)

Lemma: pi-example_wf
[x,y,z,v:Name].  (pi-example(x;y;z;v) ∈ pi_term())

Definition: pi-example-inst
pi-example-inst() ==  pi-example([x];[y];[z];[v])

Lemma: pi-example-inst_wf
pi-example-inst() ∈ pi_term()

Definition: ten-locs
ten-locs{$l_server,$l_choose,$l_comm,$l_pi,$a1,$a2,$a3,$a4,$a5,$a6,$a7}() ==  [<0, "$l_server">; <1, "$l_choose">; <2, "\000C$l_comm">; <3, "$l_pi">; <4, "$a1">; <5, "$a2">; <6, "$a3">; <7, "$a4">; <8, "$a5">; <9, "$a6">; <10, "$a7">]

Definition: pi-rank
pi-rank(p) ==
  F(p) where 
  F(0) 0  
  F(pre.P) P1 where  
              P1 F(P)   
  F(P Q) P1 Q1 where  
              P1 F(P)  
              Q1 F(Q)  
  F(P Q) P1 Q1 where  
              P1 F(P)  
              Q1 F(Q)  
  F(!P) P1 where  
           P1 F(P)  
  F(new v.R) R1 where  
                R1 F(R)

Lemma: pi-rank_wf
[p:pi_term()]. (pi-rank(p) ∈ ℕ)

Lemma: rank-zero
pi-rank(pizero()) 0 ∈ ℕ

Lemma: rank-comm
[P:pi_term()]. ∀[pre:pi_prefix()].  (pi-rank(picomm(pre;P)) (pi-rank(P) 1) ∈ ℕ)

Lemma: rank-comm-decompose
[P:pi_term()]. pi-rank(P) (pi-rank(picomm-body(P)) 1) ∈ ℕ supposing ↑picomm?(P)

Lemma: rank-option
[P,Q:pi_term()].  (pi-rank(pioption(P;Q)) ((pi-rank(P) pi-rank(Q)) 1) ∈ ℕ)

Lemma: rank-option-decompose
[P:pi_term()]. pi-rank(P) ((pi-rank(pioption-left(P)) pi-rank(pioption-right(P))) 1) ∈ ℕ supposing ↑pioption?(P)

Lemma: rank-par
[P,Q:pi_term()].  (pi-rank(pipar(P;Q)) ((pi-rank(P) pi-rank(Q)) 1) ∈ ℕ)

Lemma: rank-par-decompose
[P:pi_term()]. pi-rank(P) ((pi-rank(pipar-left(P)) pi-rank(pipar-right(P))) 1) ∈ ℕ supposing ↑pipar?(P)

Lemma: rank-rep
[P:pi_term()]. (pi-rank(pirep(P)) (pi-rank(P) 1) ∈ ℕ)

Lemma: rank-rep-decompose
[P:pi_term()]. pi-rank(P) (pi-rank(pirep-body(P)) 1) ∈ ℕ supposing ↑pirep?(P)

Lemma: rank-new
[v:Name]. ∀[P:pi_term()].  (pi-rank(pinew(v;P)) (pi-rank(P) 1) ∈ ℕ)

Lemma: rank-new-decompose
[P:pi_term()]. pi-rank(P) (pi-rank(pinew-body(P)) 1) ∈ ℕ supposing ↑pinew?(P)

Definition: pi-rank-choices
pi-rank-choices(compList) ==  if (0 =z ||compList||) then else imax-list(map(λcmp.pi-rank(snd(cmp));compList)) fi 

Lemma: pi-rank-choices_wf
[compList:(pi_prefix() × pi_term()) List]. (pi-rank-choices(compList) ∈ ℕ)

Definition: pi-level
pi-level(p) ==
  F(p) where 
  F(0) 0  
  F(pre.P) lvl where  
              lvl F(P)   
  F(P Q) imax(lvl_P;lvl_Q) where  
              lvl_P F(P)  
              lvl_Q F(Q)  
  F(P Q) imax(lvl_P;lvl_Q) where  
              lvl_P F(P)  
              lvl_Q F(Q)  
  F(!P) lvl_P where  
           lvl_P F(P)  
  F(new v.R) lvl_R where  
                lvl_R F(R)

Lemma: pi-level_wf
[p:pi_term()]. (pi-level(p) ∈ ℕ)

Definition: prefix-replace
prefix-replace(t;x;pre) ==
  G(pre) where   
  G(d<v>let d' if name_eq(d;x) then else fi  in
             let v' if name_eq(v;x) then else fi  in
             pisend(d';v')  
  G(c?(w)) let c' if name_eq(c;x) then else fi  in
              let w' if name_eq(w;x) then else fi  in
              pisend(c';w')

Lemma: prefix-replace_wf
[t,x:Name]. ∀[pre:pi_prefix()].  (prefix-replace(t;x;pre) ∈ pi_prefix())

Definition: pi-replace
pi-replace(t;x;P) ==
  F(P) where 
  F(0) pizero()  
  F(pre.P) picomm(prefix-replace(t;x;pre);P1) where  
              P1 F(P)   
  F(P Q) pioption(P1;Q1) where  
              P1 F(P)  
              Q1 F(Q)  
  F(P Q) pipar(P1;Q1) where  
              P1 F(P)  
              Q1 F(Q)  
  F(!P) pirep(P1) where  
           P1 F(P)  
  F(new v.R) let v1 if name_eq(v;x) then else fi  in
                   pinew(v1;P1) where  
                P1 F(R)

Lemma: pi-replace_wf
[t,x:Name]. ∀[P:pi_term()].  (pi-replace(t;x;P) ∈ pi_term())

Lemma: pi-rank-pi-replace
[t,x:Name]. ∀[P:pi_term()].  (pi-rank(pi-replace(t;x;P)) pi-rank(P) ∈ ℤ)

Lemma: pi-level-pi-replace
[t,x:Name]. ∀[P:pi_term()].  (pi-level(pi-replace(t;x;P)) pi-level(P) ∈ ℤ)

Definition: pi-prefix-names
pi-prefix-names(p) ==  pi_prefix_ind(p; pisend(x,y) [x; y]; pircv(x,y) [x; y]) 

Lemma: pi-prefix-names_wf
[p:pi_prefix()]. (pi-prefix-names(p) ∈ Name List)

Definition: pi-names
pi-names(p) ==
  F(p) where 
  F(0) []  
  F(pre.P) pi-prefix-names(pre) ⋃ P1 where  
              P1 F(P)   
  F(P Q) P1 ⋃ Q1 where  
              P1 F(P)  
              Q1 F(Q)  
  F(P Q) P1 ⋃ Q1 where  
              P1 F(P)  
              Q1 F(Q)  
  F(!P) P1 where  
           P1 F(P)  
  F(new v.R) [v R1] where  
                R1 F(R)

Lemma: pi-names_wf
[p:pi_term()]. (pi-names(p) ∈ Name List)

Definition: pi-subst-aux
pi-subst-aux(p) ==
  F(p) where 
  F(0) = λY,s. pizero()  
  F(pre.P) = λY,s. pi_prefix_ind(pre;
                                 pisend(c,v) picomm(pisend(name-subst(s;c);name-subst(s;v));P1 s);
                                 pircv(c,v) if v ∈b Y
                                 then let maybe-new(v;Y pi-names(P)) in
                                       let s' filter(λp.(¬bname_eq(w;fst(p)));s) in
                                       picomm(pircv(name-subst(s;c);w);P1 [w Y] [<v, w> s'])
                                 else let s' filter(λp.(¬bname_eq(v;fst(p)));s) in
                                          picomm(pircv(name-subst(s;c);v);P1 s')
                                 fi )  where  
              P1 F(P)   
  F(P Q) = λY,s. pioption(P1 s;Q1 s) where  
              P1 F(P)  
              Q1 F(Q)  
  F(P Q) = λY,s. pipar(P1 s;Q1 s) where  
              P1 F(P)  
              Q1 F(Q)  
  F(!P) = λY,s. pirep(P1 s) where  
           P1 F(P)  
  F(new v.R) = λY,s. if v ∈b Y
                    then let maybe-new(v;Y pi-names(R)) in
                          let s' filter(λp.(¬bname_eq(w;fst(p)));s) in
                          pinew(w;R1 [w Y] [<v, w> s'])
                    else let s' filter(λp.(¬bname_eq(v;fst(p)));s) in
                             pinew(v;R1 s')
                    fi  where  
                R1 F(R)

Lemma: pi-subst-aux_wf
[p:pi_term()]. (pi-subst-aux(p) ∈ (Name List) ⟶ ((Name × Name) List) ⟶ pi_term())

Definition: pi-simple-subst-aux
pi-simple-subst-aux(t;x;P;avoid) ==
  
  pi-simple-subst-aux,P,avoid. F(P) where 
                                 F(0) pizero()  
                                 F(pre.R) pi_prefix_ind(pre;
                                                          pisend(c,v) let c1 if name_eq(c;x) then else fi  in
                                                                         let v1 if name_eq(v;x) then else fi  in
                                                                         picomm(pisend(c1;v1);pi-simple-subst-aux 
                                                                                              avoid);
                                                          pircv(c,v) let maybe-new(v;avoid) in
                                                                        let c1 if name_eq(c;x) then else fi  in
                                                                        let R1 pi-replace(w;v;R) in
                                                                        picomm(pircv(c1;w);pi-simple-subst-aux R1 
                                                                                           [w avoid]))  where  
                                             F(R)   
                                 F(P Q) pioption(pi-simple-subst-aux avoid;pi-simple-subst-aux avoid) where  
                                             F(P)  
                                             F(Q)  
                                 F(P Q) pipar(pi-simple-subst-aux avoid;pi-simple-subst-aux avoid) where  
                                             F(P)  
                                             F(Q)  
                                 F(!P) pirep(pi-simple-subst-aux avoid) where  
                                          F(P)  
                                 F(new v.P) let maybe-new(v;avoid) in
                                               let P1 pi-replace(w;v;P) in
                                               pinew(w;pi-simple-subst-aux P1 [w avoid]) where  
                                               F(P)) 
  
  avoid

Lemma: pi-simple-subst-aux_wf
[P:pi_term()]. ∀[t,x:Name]. ∀[avoid:Name List].  (pi-simple-subst-aux(t;x;P;avoid) ∈ pi_term())

Lemma: pi-rank-pi-simple-subst-aux
[P:pi_term()]. ∀[t,x:Name]. ∀[avoid:Name List].  (pi-rank(pi-simple-subst-aux(t;x;P;avoid)) pi-rank(P) ∈ ℤ)

Definition: pi-subst
pi-subst(t;x;p) ==  pi-subst-aux(p) [t pi-names(p)] [<x, t>]

Lemma: pi-subst_wf
[x,t:Name]. ∀[p:pi_term()].  (pi-subst(t;x;p) ∈ pi_term())

Definition: pi-simple-subst
pi-simple-subst(t;x;P) ==  pi-simple-subst-aux(t;x;P;[t pi-names(P)])

Lemma: pi-simple-subst_wf
[t,x:Name]. ∀[P:pi_term()].  (pi-simple-subst(t;x;P) ∈ pi_term())

Lemma: pi-rank-pi-simple-subst
[P:pi_term()]. ∀[t,x:Name].  (pi-rank(pi-simple-subst(t;x;P)) pi-rank(P) ∈ ℤ)

Definition: pi-normal
pi-normal(t) ==
  pi_term_ind(t;tt;pr,A,nm.nm;A,B,na,nb.na
  ∧b nb
  ∧b (picomm?(A) ∨bpioption?(A))
  ∧b (picomm?(B) ∨bpioption?(B));A,B,na,nb.na ∧b nb;A,na.na;x,A,na.na)

Lemma: pi-normal_wf
[P:pi_term()]. (pi-normal(P) ∈ 𝔹)

Definition: pi_level
pi_level(pi-term) ==
  pi_term_ind(pi-term;0;pr,P,l.l;P,Q,l_P,l_Q.imax(l_P;l_Q);P,Q,l_P,l_Q.imax(l_P;l_Q) 1;P,l.l 1;x,P,l.l)

Lemma: pi_level_wf
[t:pi_term()]. (pi_level(t) ∈ ℕ)

Definition: pi-choices
pi-choices(t) ==  pi_term_ind(t;[];pr,P,r.[<pr, P>];P,Q,cp,cq.cp cq;P,Q,a,b.[];P,a.[];x,P,a.[])

Lemma: pi-choices_wf
[t:pi_term()]. (pi-choices(t) ∈ (pi_prefix() × pi_term()) List)

Lemma: rank-pi-choices
[P:pi_term()]. ∀[choice:{c:pi_prefix() × pi_term()| (c ∈ pi-choices(P))} ].  pi-rank(snd(choice)) < pi-rank(P)

Definition: bound_in_prefix
bound_in_prefix(pre) ==  G(pre) where   G(chan<c>[]  G(chan?(x)) [x]

Lemma: bound_in_prefix_wf
[pre:pi_prefix()]. (bound_in_prefix(pre) ∈ Name List)

Definition: free_in_prefix
free_in_prefix(pre) ==  G(pre) where   G(chan<c>[chan; c]  G(chan?(v)) [chan]

Lemma: free_in_prefix_wf
[pre:pi_prefix()]. (free_in_prefix(pre) ∈ Name List)

Definition: PiDataVal
PiDataVal() ==
  Id id:Id × Name from:Id × (pi_prefix() List) val:Name × ℕ Unit Unit + ℕ × Id × ℕ × Name (rndv2:ℕ
                                                                                                      × Id
                                                                                                      × ℕ
                                                                                                      × Name × ℕ)

Lemma: PiDataVal_wf
PiDataVal() ∈ Type

Definition: pDVloc
pDVloc(id) ==  inl id

Lemma: pDVloc_wf
[id:Id]. (pDVloc(id) ∈ PiDataVal())

Definition: pDVloc_tag
pDVloc_tag(id;name) ==  inr (inl <id, name>

Lemma: pDVloc_tag_wf
[id:Id]. ∀[name:Name].  (pDVloc_tag(id;name) ∈ PiDataVal())

Definition: pDVguards
pDVguards(from;preList) ==  inr inr (inl <from, preList>)  

Lemma: pDVguards_wf
[from:Id]. ∀[preList:pi_prefix() List].  (pDVguards(from;preList) ∈ PiDataVal())

Definition: pDVmsg
pDVmsg(val;index) ==  inr inr inr (inl <val, index>)   

Lemma: pDVmsg_wf
[val:Name]. ∀[index:ℕ].  (pDVmsg(val;index) ∈ PiDataVal())

Definition: pDVfire
pDVfire() ==  inr inr inr inr (inl ⋅)    

Lemma: pDVfire_wf
pDVfire() ∈ PiDataVal()

Definition: pDVcontinue
pDVcontinue() ==  inr inr inr inr inr (inl ⋅)     

Lemma: pDVcontinue_wf
pDVcontinue() ∈ PiDataVal()

Definition: pDVselex
pDVselex(rndv1) ==  inr inr inr inr inr inr (inl rndv1)      

Lemma: pDVselex_wf
[rndv1:ℕ × Id × ℕ × Name]. (pDVselex(rndv1) ∈ PiDataVal())

Definition: pDVrequest
pDVrequest(rndv2;counter) ==  inr inr inr inr inr inr inr <rndv2, counter>       

Lemma: pDVrequest_wf
[rndv2:ℕ × Id × ℕ × Name]. ∀[counter:ℕ].  (pDVrequest(rndv2;counter) ∈ PiDataVal())

Definition: PiDataVal_ind
F((x) where  
F(id) loc[id]  
F(id,name) loc_tag[id; name]  
F(from,preList) guards[from; preList]  
F(val,index) msg[val; index]  
F(fire) fire  
F(continue) continue  
F(rndv1) selex[rndv1]  
F(rndv2,counter) request[rndv2; counter] ==
  case x
   of inl(x) =>
   loc[x]
   inr(x) =>
   case x
    of inl(x) =>
    loc_tag[fst(x); snd(x)]
    inr(x) =>
    case x
     of inl(x) =>
     guards[fst(x); snd(x)]
     inr(x) =>
     case x
      of inl(x) =>
      msg[fst(x); snd(x)]
      inr(x) =>
      case x
       of inl(x) =>
       fire
       inr(x) =>
       case of inl(x) => continue inr(x) => case of inl(x) => selex[x] inr(x) => request[fst(x); snd(x)]

Lemma: PiDataVal_ind_wf
[X:𝕌{j}]. ∀[x:PiDataVal()]. ∀[loc:id:Id ⟶ X]. ∀[loc_tag:id:Id ⟶ name:Name ⟶ X].
[guards:from:Id ⟶ preList:(pi_prefix() List) ⟶ X]. ∀[msg:val:Name ⟶ index:ℕ ⟶ X]. ∀[fire,continue:X].
[selex:rndv1:(ℕ × Id × ℕ × Name) ⟶ X]. ∀[request:rndv2:(ℕ × Id × ℕ × Name) ⟶ counter:ℕ ⟶ X].
  (F((x) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] ∈ X)

Lemma: PiDataVal-induction
[P:PiDataVal() ⟶ ℙ]
  ((∀id:Id. P[pDVloc(id)])
   (∀id:Id. ∀name:Name.  P[pDVloc_tag(id;name)])
   (∀from:Id. ∀preList:pi_prefix() List.  P[pDVguards(from;preList)])
   (∀val:Name. ∀index:ℕ.  P[pDVmsg(val;index)])
   P[pDVfire()]
   P[pDVcontinue()]
   (∀rndv1:ℕ × Id × ℕ × Name. P[pDVselex(rndv1)])
   (∀rndv2:ℕ × Id × ℕ × Name. ∀counter:ℕ.  P[pDVrequest(rndv2;counter)])
   {∀x:PiDataVal(). P[x]})

Lemma: PiDataVal_ind_pDVloc_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,id:Top.
  (F((pDVloc(id)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] loc[id])

Lemma: PiDataVal_ind_pDVloc_tag_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,name,id:Top.
  (F((pDVloc_tag(id;name)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] loc_tag[id;name])

Lemma: PiDataVal_ind_pDVguards_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,preList,from:Top.
  (F((pDVguards(from;preList)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] guards[from;preList])

Lemma: PiDataVal_ind_pDVmsg_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,index,val:Top.
  (F((pDVmsg(val;index)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] msg[val;index])

Lemma: PiDataVal_ind_pDVfire_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc:Top.
  (F((pDVfire()) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] fire)

Lemma: PiDataVal_ind_pDVcontinue_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc:Top.
  (F((pDVcontinue()) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] continue)

Lemma: PiDataVal_ind_pDVselex_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,rndv1:Top.
  (F((pDVselex(rndv1)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] selex[rndv1])

Lemma: PiDataVal_ind_pDVrequest_lemma
request,selex,continue,fire,msg,guards,loc_tag,loc,counter,rndv2:Top.
  (F((pDVrequest(rndv2;counter)) where  
   F(id) loc[id]  
   F(id,name) loc_tag[id;name]  
   F(from,preList) guards[from;preList]  
   F(val,index) msg[val;index]  
   F(fire) fire  
   F(continue) continue  
   F(rndv1) selex[rndv1]  
   F(rndv2,counter) request[rndv2;counter] request[rndv2;counter])

Definition: pDVloc?
pDVloc?(x) ==
  F((x) where  
  F(id) tt  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVloc?_wf
[x:PiDataVal()]. (pDVloc?(x) ∈ 𝔹)

Definition: pDVloc-id
pDVloc-id(x) ==
  F((x) where  
  F(id) id  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVloc-id_wf
[x:PiDataVal()]. pDVloc-id(x) ∈ Id supposing ↑pDVloc?(x)

Definition: pDVloc_tag?
pDVloc_tag?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) tt  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVloc_tag?_wf
[x:PiDataVal()]. (pDVloc_tag?(x) ∈ 𝔹)

Definition: pDVloc_tag-id
pDVloc_tag-id(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) id  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVloc_tag-id_wf
[x:PiDataVal()]. pDVloc_tag-id(x) ∈ Id supposing ↑pDVloc_tag?(x)

Definition: pDVloc_tag-name
pDVloc_tag-name(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) name  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVloc_tag-name_wf
[x:PiDataVal()]. pDVloc_tag-name(x) ∈ Name supposing ↑pDVloc_tag?(x)

Definition: pDVguards?
pDVguards?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) tt  
  F(val,index) ff  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVguards?_wf
[x:PiDataVal()]. (pDVguards?(x) ∈ 𝔹)

Definition: pDVguards-from
pDVguards-from(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) from  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVguards-from_wf
[x:PiDataVal()]. pDVguards-from(x) ∈ Id supposing ↑pDVguards?(x)

Definition: pDVguards-preList
pDVguards-preList(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) preList  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVguards-preList_wf
[x:PiDataVal()]. pDVguards-preList(x) ∈ pi_prefix() List supposing ↑pDVguards?(x)

Definition: pDVmsg?
pDVmsg?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) tt  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVmsg?_wf
[x:PiDataVal()]. (pDVmsg?(x) ∈ 𝔹)

Definition: pDVmsg-val
pDVmsg-val(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) val  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVmsg-val_wf
[x:PiDataVal()]. pDVmsg-val(x) ∈ Name supposing ↑pDVmsg?(x)

Definition: pDVmsg-index
pDVmsg-index(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) index  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) = ⋅

Lemma: pDVmsg-index_wf
[x:PiDataVal()]. pDVmsg-index(x) ∈ ℕ supposing ↑pDVmsg?(x)

Definition: pDVfire?
pDVfire?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) tt  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVfire?_wf
[x:PiDataVal()]. (pDVfire?(x) ∈ 𝔹)

Definition: pDVcontinue?
pDVcontinue?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) ff  
  F(continue) tt  
  F(rndv1) ff  
  F(rndv2,counter) ff

Lemma: pDVcontinue?_wf
[x:PiDataVal()]. (pDVcontinue?(x) ∈ 𝔹)

Definition: pDVselex?
pDVselex?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) tt  
  F(rndv2,counter) ff

Lemma: pDVselex?_wf
[x:PiDataVal()]. (pDVselex?(x) ∈ 𝔹)

Definition: pDVselex-rndv1
pDVselex-rndv1(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) rndv1  
  F(rndv2,counter) = ⋅

Lemma: pDVselex-rndv1_wf
[x:PiDataVal()]. pDVselex-rndv1(x) ∈ ℕ × Id × ℕ × Name supposing ↑pDVselex?(x)

Definition: pDVrequest?
pDVrequest?(x) ==
  F((x) where  
  F(id) ff  
  F(id,name) ff  
  F(from,preList) ff  
  F(val,index) ff  
  F(fire) ff  
  F(continue) ff  
  F(rndv1) ff  
  F(rndv2,counter) tt

Lemma: pDVrequest?_wf
[x:PiDataVal()]. (pDVrequest?(x) ∈ 𝔹)

Definition: pDVrequest-rndv2
pDVrequest-rndv2(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) rndv2

Lemma: pDVrequest-rndv2_wf
[x:PiDataVal()]. pDVrequest-rndv2(x) ∈ ℕ × Id × ℕ × Name supposing ↑pDVrequest?(x)

Definition: pDVrequest-counter
pDVrequest-counter(x) ==
  F((x) where  
  F(id) = ⋅  
  F(id,name) = ⋅  
  F(from,preList) = ⋅  
  F(val,index) = ⋅  
  F(fire) = ⋅  
  F(continue) = ⋅  
  F(rndv1) = ⋅  
  F(rndv2,counter) counter

Lemma: pDVrequest-counter_wf
[x:PiDataVal()]. pDVrequest-counter(x) ∈ ℕ supposing ↑pDVrequest?(x)

Comment: PiData-comment
PiData represents all possible "data messages":

  loc(id1)  --  request for fresh location and serial number, sent from 
                target process to loc-Server; *id1* is the return
                   address;
                 

  loc_tag(id2,name) -- response from loc-Server; *id2* is the fresh
                       location; *name* is its (globally unique)
                       serial number
                       (Note that, officially, Name == Atom List)

  guards(from,preList) -- request from target process to Comm for a
                          communication; *from* is location of target
                          process; *preList* is list of possible 
                          sends and receives

  msg(val,index) -- reply from Comm to target process; *index* identifies
                    which communication was chosen and *val* is the value
                    transmitted (or, in the case of send, the ack)

  fire     -- used by the boot process to launch newly installed component

  continue -- sent by Comm to itself

  selex(rndv1: Nat Id Nat Name)    
           -- sent by Choose to Comm; the choice that Comm is blocked
              and waiting for' "triple" is an unfortunate name;
              the point: in <i,loc,j, nm>, <i,loc,j> identifies the
              rendezvous and nm is the value sent              
  request(rndv2: Nat Id Nat Name, counter: Nat)
           -- sent by Comm to Choose; counte; counter *indices* is one possible
              choice on round counter

The terminology has gotten slightly out of sync with the paper of Jan 28, 2010.  
The type called Reqs there corresponds to the "guards" messages here -- "guards"
messages are sent to Comm, which then stores them in either queue (before
handling) or an array (after, if they're still "alive").  
              ⋅

Definition: piM
piM(T) ==  PiDataVal()

Lemma: piM_wf
[T:Type]. (piM(T) ∈ Type)

Lemma: piM-continuous
Continuous+(T.piM(T))

Definition: pi-process
pi-process() ==  Process(T.piM(T))

Lemma: pi-process_wf
pi-process() ∈ Type

Lemma: rec-process_wf_pi
[S:Type ⟶ Type]
  ∀[s0:S[pi-process()]]. ∀[next:⋂T:{T:Type| pi-process() ⊆T} 
                                  (S[piM(T) ⟶ (T × LabeledDAG(Id × (Com(T.piM(T)) T)))]
                                  ⟶ piM(T)
                                  ⟶ (S[T] × LabeledDAG(Id × (Com(T.piM(T)) T))))].
    (RecProcess(s0;s,m.next[s;m]) ∈ pi-process()) 
  supposing Continuous+(T.S[T])

Lemma: rec-process_wf_pi_simple_state
[S:Type]. ∀[s0:S]. ∀[next:⋂T:{T:Type| pi-process() ⊆T} (S ⟶ piM(T) ⟶ (S × LabeledDAG(Id × (Com(T.piM(T)) T))))].
  (RecProcess(s0;s,m.next[s;m]) ∈ pi-process())

Comment: loc-server-comment
The intention: When loc-server receives message consisting of location
it sends to message consisting of pair:
   globally fresh location and 
   unique serial number (the name corresponding to an integer counter, 
        followed by "/")

Note that the only property really required of is this:

  Let x_0 f([])
      x_1 f([x_0])
      x_2 f([x_0,x_1])
      x_3 f([x_0,x_1,x_2]), etc.

Then all the x_i must be different.  So, e.g., if 

  Nat -> Id is 1-1

we can define 

  f(idlist) f(|idList|)

and that will do.

If we do that, the initial seeding of loc-server must always take the form

   [x_0, x_1, ..., x_n]



 ⋅

Definition: loc-Server
loc-Server(f;L0;n0) ==
  RecProcess(<L0, n0>;s,i.if pDVloc?(i)
  then let sndr pDVloc-id(i) in
           let L,n 
           in <<[f L], 1>make-lg([<sndr, mk-tagged("msg";pDVloc_tag(f L;nat-to-incomparable(n 1)))>])>
  else <s, make-lg([])>
  fi )

Lemma: loc-Server_wf2
[f:(Id List) ⟶ Id]. ∀[L:Id List]. ∀[n:ℕ].  (loc-Server(f;L;n) ∈ pi-process())

Comment: maybe-new-local-comment
When it's created, each target process will be assigned unique "serial number."
Serial numbers will have the property that if two of them are distinct, then
neither is prefix of the other. Thus, if serial1 \= serial2 then, for any
name1 and name2 (whether they differ or not)

   serial1@name1 \= serial2@name2

In the intended use of

   maybe-new-local(me, s, avoid)

me is the process serial number, the name for which replacement is sought,
and avoid is the list of already used (local) names to be avoided.  (The 
definition will be slightly wasteful -- generating longer names than strictly
needed, because we only need to avoid the "postfix" 
part of the names that follow "me"; if desired, that could be fixed.)⋅

Definition: maybe-new-local
maybe-new-local(me;s;avoid) ==  me maybe-new(s;avoid)

Lemma: maybe-new-local_wf
[me,s:Name]. ∀[avoid:Name List].  (maybe-new-local(me;s;avoid) ∈ Name)

Comment: processChoose-comment
The Choose process communicates only with the Comm process.  The parameter
to processChoose is the location of Comm.  Its definition is straight out of
the note: each request carries with it an identifying counter; the state of 
Choose is the list of counters it's seen.  When it sees new counter value
it echoes the other component of the request back to Comm; when it sees an 
old counter value or any other message (it's not supposed to receive any
others) it's no-op.  (We could cut down on the size of the state if we
wanted to, but who cares?)⋅

Definition: processChoose
processChoose(l_comm) ==
  RecProcess([];seen,i.if pDVrequest?(i)
  then let trip pDVrequest-rndv2(i) in
        let cntr pDVrequest-counter(i) in
        if cntr ∈b seen then <seen, make-lg([])> else <[cntr seen], make-lg([<l_comm, mk-tagged("msg";pDVselex(trip))>\000C])> fi 
  else <seen, make-lg([])>
  fi )

Lemma: processChoose_wf2
[l_comm:Id]. (processChoose(l_comm) ∈ pi-process())

Definition: select-tagged-indices-aux
select-tagged-indices-aux(P;tg;L) ==
  rec-case(L) of
  [] => λn.[]
  fst::rest =>
   r.if fst then λn.[<n, tg fst> (r (n 1))] else λn.(r (n 1)) fi 

Lemma: select-tagged-indices-aux_wf
[A,B:Type]. ∀[tg:A ⟶ B]. ∀[P:A ⟶ 𝔹]. ∀[L:A List].  (select-tagged-indices-aux(P;tg;L) ∈ ℕ ⟶ ((ℕ × B) List))

Definition: select-tagged-indices
select-tagged-indices(P;tg;L) ==  select-tagged-indices-aux(P;tg;L) 0

Lemma: select-tagged-indices_wf
[A,B:Type]. ∀[tg:A ⟶ B]. ∀[P:A ⟶ 𝔹]. ∀[L:A List].  (select-tagged-indices(P;tg;L) ∈ (ℕ × B) List)

Comment: accum-matching-indices-comment
NEED TO CHANGE THIS (SEE BELOW)

The main state component of Comm (call it st) is finite partial function
Id -fp-> pi_prefix List.  For any prefix pre,

   accum-matching-indices(pre,st) \in (Id Nat List) List

consists of those pairs (loc,ns) such that ns lists the 
indices in st(loc) of those prefixes matching pre.

It's computed from fpf-accum, where (in the notation of the wf lemma)

   =  (Id Nat List) List
   =  Id
   =  \a. pi_prefix List  

We want

   fpf-accum(f; []; st)

where, as usual, [] is the initial value and is the accumulating function:

   f(result_so_far, loc, preList) cons("the new pair", result_so_far) where

preList will be computed by fpf-accum as st(loc) and 

   "the new pair" (loc, "the indices matching preList") 

where

   "the indices ..." select-indices(P; preList) 

and, finally

    \p. prefix-match(p;pre)

CHANGES:

    Instead of just returning indices we have to return list of (index,name) pairs,
    the name being the value communicated if the rendezvous succeeds.  If pre is 
    send, then it determines the value, which is the same for all matches; if pre is
    receive, then the value depends on the matching send chosen.

  accum-matching-tagged-indices(pre,st) \in (Id (Nat pi_prefix) List) List

We use fpf-accum, where now

   (Id (Nat name) List) List
   Id
   la. pi_prefix List

So we want

   fpf-accum(f; []; st) 

where now

    f(result_so_far, loc, preList) cons("the new pair with tags", result_so_far) where

Revising, from the bottom up, we need

     is as before

    "the tagged-indices ...." select-tagged-indices(P; tg; preList)
          where tg \x. if "x is c(v)" then v
                         else if "pre is c(v)" then v
                         else []  -- can't happen   

    "the new pair with tags" (loc, "the tagged-indices matching preList")



Definition: accum-matching-tagged-indices
accum-matching-tagged-indices(pre;st) ==
  let = λp.prefix-match(p;pre) in
   let tg = λx.if pisend?(x) then pisend-var(x)
               if pisend?(pre) then pisend-var(pre)
               else []
               fi  in
   let indices = λpreList.select-tagged-indices(P;tg;preList) in
   let new_pair = λloc,preList. <loc, indices preList> in
   let = λrslt,loc,preList. [new_pair loc preList rslt] in
   fpf-accum(c,loc,preList.f loc preList;[];st)

Lemma: accum-matching-tagged-indices_wf
[pre:pi_prefix()]. ∀[st:a:Id fp-> pi_prefix() List].
  (accum-matching-tagged-indices(pre;st) ∈ (Id × ((ℕ × Name) List)) List)

Comment: interleaves-and-gets-comment
Note: It used to return triples; now returns quadruples -- not worth
changing the name.

CHANGES NEEDED.  HERE'S THE ORIGINAL:

We're interested in the case

  Nat
  Id
  Nat

    interleave1(i; (loc, [m;n;...]) i,loc,m; i,loc,n; ... ]

    interleave2(i; loc1, [m1;n1;...]; loc2, [m2;n2;...]; ... ])
        concatenation of application of interleave1 to each element
        i,loc1,m1; i,loc1,n1; ...; i,loc2,m2; i,loc2,n2; ... ]

    get-triples([pre0;pre1;...], st)
        list of all the matching triple
        let \n,p. <n, accum-matching-indices(p;st)>
          let tmp map-index(f; [pre1; pre2; ... ]
                 0,v1; 1,v2; ... ]
                      where each \in (Id Nat List) List
          in concat(map(\x. interleave2(pi1 x; p12 x; tmp)

Note: if pre_j has no match in st, then the component of tmp is (j,[]),
and when you apply \x.interelave2(pi x, pi2 x) to this you get [].  So
this contributes nothing to the concat.  We could get slight boost to
efficiency by filtering tmp to eliminate all the (j,[]) before submitting
it to the "concat(map(...))".

CHANGES MADE:

Now

 Nat
 Id
 Nat Name

    interleave1(i; (loc, [m, name_m; n,name_n; ,,,]
        i, loc, <m,name_m]; i, loc, <n,name_n>,,, ]

    interleave2(i; [loc1, [m1, name_m1; n1, name_n1; ...] ;
                    loc2, [m2, name_m2; n2, name_n2; ...]
                    ... =
        i, loc1, <m1, name_m1>i, loc1, <n1, name_n1> ...
            i, loc2, <m2, name_m2>i, loc2, <m2, name_m2> ... ]

     get-tagged-triples(pre0; pre1; ... 
        let \n,p. <n, accum-matching-tagged-indices(p,st)
          ... the rest is the same




Definition: interleave1
interleave1(a;u) ==  map(λc.<a, fst(u), c>;snd(u))

Lemma: interleave1_wf
[C,B,A:Type]. ∀[a:A]. ∀[u:B × (C List)].  (interleave1(a;u) ∈ (A × B × C) List)

Definition: interleave2
interleave2(a;v) ==  concat(map(λu.interleave1(a;u);v))

Lemma: interleave2_wf
[C,B,A:Type]. ∀[a:A]. ∀[v:(B × (C List)) List].  (interleave2(a;v) ∈ (A × B × C) List)

Definition: get-triples
get-triples(preList;st) ==
  let = λn,p. <n, accum-matching-tagged-indices(p;st)> in
   let tmp map-index(f;preList) in
   concat(map(λx.interleave2(fst(x);snd(x));tmp))

Lemma: get-triples_wf
[preList:pi_prefix() List]. ∀[st:x:Id fp-> pi_prefix() List].  (get-triples(preList;st) ∈ (ℕ × Id × ℕ × Name) List)

Comment: pi-add-comment
NO LONGER NEEDED: This operation is rolled into Comm-process-q:

pi-add is the operation called "add" in the note on modeling pi-calculus.
Given

    q  (Id (pi_prefix List)) List
          -- though of as queue, whose first element is head(q)

    st Id -fp-> (pi_prefix List)

  pi-add(q,st) result of doing the update st[loc := preList]
                 for each (loc, preList) \in q

Order will be unimportant because this will only be applied when, for
all (loc, preList) \in q, loc \not\in fpf-domain(st).

This definition may suggest that using finite partial functions, though
conceptually reasonable, is overkill.  (Because there's never any possibility
of conflict we could get by making st list of pairs.) 


Definition: pi-add
pi-add(eq;q;st) ==
  accumulate (with value and list item req):
   fst(req) snd(req) ⊕ s
  over list:
    q
  with starting value:
   st)

Lemma: pi-add_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[eq:EqDecider(A)]. ∀[q:(a:A × B[a]) List]. ∀[st:a:A fp-> B[a]].
  (pi-add(eq;q;st) ∈ a:A fp-> B[a])

Comment: Comm-next-comment
Split the definition of the next-state function for process Comm
into separate functions for each kind of message it can receive.
It has two parameters, l_comm and l_choose, which will denote the
locations of Comm and Choose:

Comm-state Id -fp-> (pi_prefix) List x
             (Id pi_prefix List) List x
              Id x                 -- when it's waiting for selection
                                      it remembers the party whose
                                      request triggered this wait.
                                      The value is meaningful only
                                      when the Bool component is true.
              Bool x               -- whether it's waiting or not
              Nat                  -- counter identifying the current
                                      round of choices

Write typical element of Comm-state as <st,q,req_from,waiting,count>  or just __

Comm-next-guards( guards(from,preList), __ 
        <<st, q@<from,preList>,req_from,waiting,count>[l_comm, continue]>
   -- enqueue the request and continue

Comm-next-selex( selex<i, loc, j, val>__ =
           < "st with loc removed from domain", req_from, q, false,count>
              [l_comm, continue; req_from, <val,i>loc,<val,j>] >
   -- Recall: <i, loc, j, val> means that the i^th choice
      at location req_from matches the j^th choice at location
      loc, and rendezvous would communicate value val.
      So the receiver learns which choice was selected and the
      value it receives; the sender learns which choice was selected
      and gets back the value sent, which can be regarded as an ack.
      It's safe to leave req_from unchanged, because waiting has 
      become false.


To define the remaining case, Comm-next-continue, we first need to define 

    Comm-process-q(q, id, st) 
        []  =>  < [], st, id, [] >
        = <loc,preList>::tl => 
            let mList get-triples(preList,st) in
            if mList [] then Comm-process-q(tail, st + <loc,preList>
            else
               <tail, st, loc, mList>


In the output, tail is the remaining unprocessed queue, st is the state
of the request array after this preliminary processing of the front part
of the queue, loc is the location for which rendezvous will now
be chosen, mList encodes the possible rendezvous outcomes, from which
Choose will select one.  The value of the "id" parameter doesn't matter
because no rendezvous will take place when it's used.  Casting this 
definition in the form required to use list_ind we first need an 
auxiliary function:

   Comm-process-q_aux(q) =
        list_ind(q;
                 \id,st. <q,st,id,[]>;
                 hd,rl,r. \id,st. let <loc,preList> hd in
                                  let mList get-triples(preList,st) in
                                  if mList [] then (st hd)
                                  else 
                                     <tail, st, loc, mList> 


   Comm-process-q(st,q,dont_care) Comm-process-q_aux(q,dont_care) st

 
Now, this function has the additional parameters, l_comm, l_choose,
and the counter.

   Comm-next-continue( continue, <st,q,req_from,waiting,count> =
      if waiting then (__, [])
      else let <q', st', loc, mList> Comm-process(st,q) in
          if q' [] then <<st', [], _, false>[]>
          else <<st',q',loc,true,count+1>
                 [<l_comm, continue>;
                  <l_choose, request(trip1,count)>;
                  <l_choose, request(trip2,count) ...> ... ]


where mList [trip1, trip2, ... ]

processComm(l_comm, l_choose) =
   RecProcess(<fpf-empty, [], l_comm, bfalse,0>;
               cSt,piD. if pDVguards?(piD) then Comm-next-guards(l_comm; piD; cSt)
                        else if pDVselex?(piD) then Comm-next-selex(l_comm; piD; cSt)
                        else if pDVcontinue?(piD) then Comm-next-continue(l_comm;l_choose;piD;cSt)
                        else <cSt,[]>  

 .  


Definition: Comm-state
Comm-state() ==  st:Id fp-> pi_prefix() List × (Id × (pi_prefix() List)) List × Id × 𝔹 × ℕ

Lemma: Comm-state_wf
Comm-state() ∈ Type

Definition: Comm-output
Comm-output() ==  Comm-state() × LabeledDAG(Id × "msg"×PiDataVal())

Lemma: Comm-output_wf
Comm-output() ∈ Type

Definition: Comm-st
Comm-st(cSt) ==  fst(cSt)

Lemma: Comm-st_wf
[cSt:Comm-state()]. (Comm-st(cSt) ∈ st:Id fp-> pi_prefix() List)

Definition: Comm-q
Comm-q(cSt) ==  fst(snd(cSt))

Lemma: Comm-q_wf
[cSt:Comm-state()]. (Comm-q(cSt) ∈ (Id × (pi_prefix() List)) List)

Definition: Comm-req_from
Comm-req_from(cSt) ==  fst(snd(snd(cSt)))

Lemma: Comm-req_from_wf
[cSt:Comm-state()]. (Comm-req_from(cSt) ∈ Id)

Definition: Comm-waiting
Comm-waiting(cSt) ==  fst(snd(snd(snd(cSt))))

Lemma: Comm-waiting_wf
[cSt:Comm-state()]. (Comm-waiting(cSt) ∈ 𝔹)

Definition: Comm-count
Comm-count(cSt) ==  snd(snd(snd(snd(cSt))))

Lemma: Comm-count_wf
[cSt:Comm-state()]. (Comm-count(cSt) ∈ ℕ)

Definition: Comm-next-guards
Comm-next-guards(l_comm;piD;cSt) ==
  <<Comm-st(cSt)
   Comm-q(cSt) [<pDVguards-from(piD), pDVguards-preList(piD)>]
   Comm-req_from(cSt)
   Comm-waiting(cSt)
   Comm-count(cSt)>
  make-lg([<l_comm, mk-tagged("msg";pDVcontinue())>])
  >

Lemma: Comm-next-guards_wf
[l_comm:Id]. ∀[piD:PiDataVal()]. ∀[cSt:Comm-state()].
  Comm-next-guards(l_comm;piD;cSt) ∈ Comm-output() supposing ↑pDVguards?(piD)

Definition: Comm-next-selex
Comm-next-selex(l_comm;piD;cSt) ==
  let rndv pDVselex-rndv1(piD) in
   let val snd(snd(snd(rndv))) in
   let fst(rndv) in
   let loc fst(snd(rndv)) in
   let fst(snd(snd(rndv))) in
   let cSt' fpf-restrict(Comm-st(cSt);λx.(¬bloc)) in
   <<cSt', Comm-q(cSt), Comm-req_from(cSt), ff, Comm-count(cSt)>
   make-lg([<l_comm, mk-tagged("msg";pDVcontinue())>; <Comm-req_from(cSt), mk-tagged("msg";pDVmsg(val;i))>; <loc, mk-t\000Cagged("msg";pDVmsg(val;j))>])
   >

Lemma: Comm-next-selex_wf
[l_comm:Id]. ∀[piD:PiDataVal()]. ∀[cSt:Comm-state()].
  Comm-next-selex(l_comm;piD;cSt) ∈ Comm-output() supposing ↑pDVselex?(piD)

Definition: Comm-process-q_aux
Comm-process-q_aux(q) ==
  rec-case(q) of
  [] => λid,st. <[], st, id, []>
  hd::tl =>
   aux.λid,st. let loc fst(hd) in
                let preList snd(hd) in
                let mList get-triples(preList;st) in
                if null(mList) then aux id loc preList ⊕ st else <tl, st, loc, mList> fi 

Lemma: Comm-process-q_aux_wf
[q:(Id × (pi_prefix() List)) List]
  (Comm-process-q_aux(q) ∈ Id
   ⟶ st:Id fp-> pi_prefix() List
   ⟶ ((Id × (pi_prefix() List)) List × st:Id fp-> pi_prefix() List × Id × ((ℕ × Id × ℕ × Name) List)))

Definition: Comm-process-q
Comm-process-q(q;id;st) ==  Comm-process-q_aux(q) id st

Lemma: Comm-process-q_wf
[q:(Id × (pi_prefix() List)) List]. ∀[id:Id]. ∀[st:st:Id fp-> pi_prefix() List].
  (Comm-process-q(q;id;st) ∈ (Id × (pi_prefix() List)) List
   × st:Id fp-> pi_prefix() List
   × Id
   × ((ℕ × Id × ℕ × Name) List))

Definition: Comm-next-continue
Comm-next-continue(l_comm;l_choose;piD;cSt) ==
  let st Comm-st(cSt) in
   let Comm-q(cSt) in
   let waiting Comm-waiting(cSt) in
   let count Comm-count(cSt) in
   if waiting
   then <cSt, make-lg([])>
   else let Comm-process-q(q;l_comm;st) in
         let q' fst(p) in
         let st' fst(snd(p)) in
         let loc fst(snd(snd(p))) in
         let mList snd(snd(snd(p))) in
         if null(q')
         then <<st', [], l_comm, ff, count>make-lg([])>
         else let msgs2choose map(λx.<l_choose, mk-tagged("msg";pDVrequest(x;count))>;mList) in
                  <<st', q', loc, tt, count 1>make-lg([<l_comm, mk-tagged("msg";pDVcontinue())>msgs2choose)>
         fi 
   fi 

Lemma: Comm-next-continue_wf
[l_comm,l_choose:Id]. ∀[piD:PiDataVal()]. ∀[cSt:Comm-state()].
  Comm-next-continue(l_comm;l_choose;piD;cSt) ∈ Comm-output() supposing ↑pDVcontinue?(piD)

Definition: processComm
processComm(l_comm;l_choose) ==
  RecProcess(<⊗[], l_comm, ff, 0>;cSt,piD.if pDVguards?(piD) then Comm-next-guards(l_comm;piD;cSt)
  if pDVselex?(piD) then Comm-next-selex(l_comm;piD;cSt)
  if pDVcontinue?(piD) then Comm-next-continue(l_comm;l_choose;piD;cSt)
  else <cSt, make-lg([])>
  fi )

Lemma: processComm_wf2
[l_comm,l_choose:Id].  (processComm(l_comm;l_choose) ∈ pi-process())

Comment: pi-bar-trans-comment
If 

    l_loc: Id               -- location of loc-Server
    P, pi_term
    pi_term -> Id -> Name -> Name List -> pi-process

then 

    pi-bar-trans(l_loc; P; Q; g) Id -> Name -> Name List -> pi-processK

defines the translation of (P|Q) -- assuming that is
recursively instantiated with function that translates
and Q.  The meaning of the parameters to (and to the output of pi-bar-trans) are

    Id        -- location at which process is installed
    Name      -- serial number of the process (globally unique)
    Name List -- list of local names to be avoided in choosing fresh name 

(P|Q) will have four states, which we'll just code as

   waiting for the "fire" message
   waiting for the first new location and serial number (from loc-Server)
   waiting for the second new location and serial number
   the null proczess

   pi-bar-trans(l_loc; P; Q; g) \l,serial,avoid. rec-process(0;next)

where

    nex(s;m) if fire then                 -- pDVfire?(m)
                    <1, make-lg([l_loc, "msg":l])>      -- code "l" as pDVloc(l)

                  -- change to state 1, request new location and tag
                  -- external part is the labeled graph containing one node,
                  -- the request (that's what make-lg is making)

               elsif = <l1,n1> then  
                        -- test is pDVloc_tag?(m)
                        -- "l1" is pDVloc_tag-id(m),
                        -- "n1" is pDVloc_tag-name(m)
                    <2, graph containing the nodes
                            <l1, "create": l1 n1 avoid 
                            <l1, "msg": fire>
                            <l_loc, "msg": l       -- coded as pDVloc(l) 
                            and an edge pointing from the first of these to the second

                  -- wait for new location and tag,
                     change to state 2, 
                     create process modeling P, then send its fire message
                     note that the "avoid" set needs to contain the
                       free names in P; rather than compute the free names,
                       this simply passes on the avoid set for P|Q 
                     request another location and tag

               elseif = <l2,n2> then
                    <3, graph containing the modes
                            <l2, "create": l2 n2 avoid>
                            <l2, "msg": fire>
                         and an edge pointing from the first to the second
                       
                  -- wait for new location and tag.
                     change to state (essentially, the null process),
                     create process modeling Q

               else <s, []>                     -- no-op
                  -- in particular, once it reaches state it's the null process

To apply the wf lemma for rec-process, the relevant `S' function is simply

    \T. {0,1,2,3}        
    




Definition: pi-bar-trans
pi-bar-trans(l_loc;P;Q;g) ==
  λl,serial,avoid. RecProcess(0;s,m.if (s =z 0) ∧b pDVfire?(m) then <1, make-lg([<l_loc, mk-tagged("msg";pDVloc(l))>])>
                  if (s =z 1) ∧b pDVloc_tag?(m)
                    then let l1 pDVloc_tag-id(m) in
                          let n1 pDVloc_tag-name(m) in
                          <2
                          lg-add(make-lg([<l1, mk-tagged("create";g l1 n1 avoid)>; <l1, mk-tagged("msg";pDVfire())>;\000C <l_loc, mk-tagged("msg";pDVloc(l))>]);0;1)
                          >
                  if (s =z 2) ∧b pDVloc_tag?(m)
                    then let l2 pDVloc_tag-id(m) in
                          let n2 pDVloc_tag-name(m) in
                          <3
                          lg-add(make-lg([<l2, mk-tagged("create";g l2 n2 avoid)>; <l2, mk-tagged("msg";pDVfire())>]\000C);0;1)
                          >
                  else <s, make-lg([])>
                  fi )

Lemma: pi-bar-trans_wf
[l_loc:Id]. ∀[P,Q:pi_term()]. ∀[g:{R:pi_term()| pi-rank(R) < pi-rank(pipar(P;Q))} 
                                   ⟶ Id
                                   ⟶ Name
                                   ⟶ (Name List)
                                   ⟶ pi-process()].
  (pi-bar-trans(l_loc;P;Q;g) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Comment: pi-rep-trans-comment
Translating !P

If 

    l_loc: Id               -- location of loc-Server
    P: pi_term
    pi_term -> Id -> Name -> Name List -> pi-process

then 

    pi-bar-trans(l_loc; P; g) Id -> Name -> Name List -> pi-process

defines the translation of (!P) -- assuming that is
recursively instantiated with function that translates
and Q.

(P|Q) will have two states, which we'll just code as

   waiting for the "fire" message
   waiting for the new location (from loc-Server) 

   pi-bar-trans(l_loc; P; g) \l, serial, avoid. rec-process(0;next)

where

    nex(s;m) if fire then        -- pDVfire?(m)
                    <1, make-lg([l_loc, "msg":l])>   -- code "l" as pDVloc(l)

                  -- change to state 1, request new location and tag

               elsif = <l1,n1> then  
                        -- test ipDVloc_tag?(m))
                        -- "l1" is pDVloc_tag-id(m),
                        -- "n1" is pDVloc_tag-name(m)

                    <0, graph consisting of three nodes
                        <l1, "create": l1 n1 avoid>
                        <l1, "msg":fire>
                        <l,  "msg":fire> 
                    with an edge from the first to the second
                        -- "l" is inl pDVloc(l)
                        -- "fire" is pDVfire

                  -- wait for new location and tag,
                     change back to state 0, 
                     create process modeling P,
                     send fire message to yourself
                  -- the only names that the new process needs to
                     avoid are the free names in P, but instead of
                     computing those we simply pass long superset,
                     the avoid set of !P

               else <s,make-lg([])>


Definition: pi-rep-trans
pi-rep-trans(l_loc;P;g) ==
  λl,serial,avoid. RecProcess(0;s,m.if (s =z 0) ∧b pDVfire?(m) then <1, make-lg([<l_loc, mk-tagged("msg";pDVloc(l))>])>
                  if (s =z 1) ∧b pDVloc_tag?(m)
                    then let l1 pDVloc_tag-id(m) in
                          let n1 pDVloc_tag-name(m) in
                          <0
                          lg-add(make-lg([<l1, mk-tagged("create";g l1 n1 avoid)>; <l1, mk-tagged("msg";pDVfire())>;\000C <l, mk-tagged("msg";pDVfire())>]);0;1)
                          >
                  else <s, make-lg([])>
                  fi )

Lemma: pi-rep-trans_wf
[l_loc:Id]. ∀[P:pi_term()]. ∀[g:{Q:pi_term()| pi-rank(Q) < pi-rank(pirep(P))} 
                                 ⟶ Id
                                 ⟶ Name
                                 ⟶ (Name List)
                                 ⟶ pi-process()].
  (pi-rep-trans(l_loc;P;g) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Comment: pi-new-trans-comment
Translating (new x.P)

If

   x: name
   P: pi_term
   g: pi_term -> Id -> Name -> Name List -> pi-process

then

   pi-new-trans(x;P;g) Id -> Name -> Name List -> pi-process


where

   pi-new-trans(x;P;g) =  
      \l, serial, avoid. 
             \m. if fire then 
                 let maybe-new-local(serial;x;avoid) in
                    < g(P[x := n]) serial avoid+n, make-lg([<l, fire>]) >
              else pi-new-trans(x;P;g)

and

     P[x := n] pi-simple-subst(n;x;P)

Note: we don't actually have to make this recursive; we could simply
have it transition on the receipt of any message at all.  Still, it
seems more reasonable, conceptually, to do it this way.

Question: We could eliminate the "avoid" parameter altogether by adopting
different strategy -- installing translation of P[x := n] at new 
location.  Then pi-new-trans could be defined by using rec-process.  It 
seems instructive not to do that (unless get stuck).
 


Definition: pi-new-trans
pi-new-trans(x;P;g)
==r λl,serial,avoid,m. if pDVfire?(m)
                      then let maybe-new-local(serial;x;avoid) in
                            let P' pi-simple-subst(n;x;P) in
                            <P' serial avoid ⋃ pi-names(P), make-lg([<l, mk-tagged("msg";pDVfire())>])>
                      else <pi-new-trans(x;P;g) serial avoid, make-lg([])>
                      fi 

Lemma: pi-new-trans_wf
[x:Name]. ∀[P:pi_term()]. ∀[g:{Q:pi_term()| pi-rank(Q) < pi-rank(pinew(x;P))} 
                               ⟶ Id
                               ⟶ Name
                               ⟶ (Name List)
                               ⟶ pi-process()].
  (pi-new-trans(x;P;g) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Definition: pi-null-trans
pi-null-trans() ==  RecProcess(⋅;s,m.<⋅make-lg([])>)

Lemma: pi-null-trans_wf
pi-null-trans() ∈ pi-process()

Comment: pi-guarded-trans-comment
Translating (pre1.P1 pre2.P2 ... pre_n.P_n

Note:  The type of pi terms wasn't set up with primitive operation to 
produce such guarded choices.  Everything has to be built from binary "+"
and unary prefix.  So we assume normal terms (or whatever they got called)
in which occurs only between prefixed terms; and the first step in 
translating "+" is to turn it into list

   [<pre1,P1>; <pre2,P2>... ; <pre_n,P_n>]

We first define an auxiliary function pi-guarded-aux.  This defines a
process that responds to selection received from the Comm process.
Then pi-gurded-trans will be the process that sends request to
Comm and then becomes the process defined by pi-select-aux.

** The aux function

If

  compList (pi_prefix pi_term) List
  g        pi_term -> Id -> Name -> Name List -> pi-process

then

  pi-guarded-aux(compList;g) Id -> Name -> Name List -> pi-process 

where

  pi-guarded-aux(compList;g) \l,serial,avoid.
     \m. if pDVmsg?(m)        -- is reply from Comm
            become process P' -- see below
          else
            <pi-guarded-aux(compList;g) serial avoid,
             make-lg([])>     -- ignore other messages

Process P' is 

    let <val, i> unpack m
         -- val pDVmsg-val(m)
         -- pDVmsg-index(m)
    let <pre_i,P_i> compList[i]    -- select(compList;i)
    let P' if pircv?(pre_i) then   -- i.e., pre_i c(x)
                P_i[ x:= val]        -- pircv-var(pre_i) 
             else 
                P_i
          -- P' is the appropriate process to become 
      <P' serial avoid@pi-names(P'), make-lg[l, "msg":pDVfire]>
          -- become that process and tell yourself to fire

          -- all we really need to append to avoid is the collection
             of new news introduced by the substitution operator;
             for simplicity's sake, and with much redundancy,
             we append all the names in P'

Note the implicit invariant -- that the in the message returned
by Comm is less than the length of compList.  (The official definition
will say that in that (impossible) case, the return message is ignored.)                          
 
        
** The translation function

  pi-trans-guarded(compList;g) \l,serial,avoid.
     \m. if pDVfire?(m)       
            <pi-guarded-aux(compList;g) serial avoid, 
             make-lg([<l, "msg":fire>])>
          else
            <pi-guarded-aux(compList;g) serial avoid,
             make-lg([])>




Definition: pi-guarded-aux
pi-guarded-aux(compList;g)
==r λl,serial,avoid,m. if pDVmsg?(m)
                      then let val pDVmsg-val(m) in
                            let pDVmsg-index(m) in
                            if i <||compList||
                            then let pre_i,P_i compList[i] 
                                 in let P' ... in
                                        <P' serial avoid ⋃ pi-names(P'), make-lg([<l, mk-tagged("msg";pDVfire())>])>
                            else <pi-guarded-aux(compList;g) serial avoid, make-lg([])>
                            fi 
                      else <pi-guarded-aux(compList;g) serial avoid, make-lg([])>
                      fi 

Lemma: pi-guarded-aux_wf
[compList:(pi_prefix() × pi_term()) List]. ∀[g:{Q:pi_term()| (∃C∈compList. pi-rank(Q) ≤ pi-rank(snd(C)))} 
                                                ⟶ Id
                                                ⟶ Name
                                                ⟶ (Name List)
                                                ⟶ pi-process()].
  (pi-guarded-aux(compList;g) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Definition: pi-guarded-trans
pi-guarded-trans(compList;g)
==r λl,serial,avoid,m. if pDVfire?(m)
                      then <pi-guarded-aux(compList;g) serial avoid, make-lg([<l, mk-tagged("msg";pDVfire())>])>
                      else <pi-guarded-trans(compList;g) serial avoid, make-lg([])>
                      fi 

Lemma: pi-guarded-trans_wf
[compList:(pi_prefix() × pi_term()) List]. ∀[g:{Q:pi_term()| (∃C∈compList. pi-rank(Q) ≤ pi-rank(snd(C)))} 
                                                ⟶ Id
                                                ⟶ Name
                                                ⟶ (Name List)
                                                ⟶ pi-process()].
  (pi-guarded-trans(compList;g) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Comment: pi-trans-comment
Perhaps some poor planning: The way the pi-*-trans operations are
defined it's not possible to formualte this definition using 
pi_term_ind.  It has to be made into recursive definition.

Also ... not clear whether it makes more sense to include an
impossible case in the definition (the last one) or not.


Definition: pi-trans
pi-trans(l_loc;P)
==r if pizero?(P) then λl,serial,avoid. pi-null-trans()
    if pipar?(P) then let pipar-left(P) in let pipar-right(P) in pi-bar-trans(l_loc;Q;R;λp.pi-trans(l_loc;p))
    if pirep?(P) then let pirep-body(P) in pi-rep-trans(l_loc;Q;λp.pi-trans(l_loc;p))
    if pinew?(P) then let pinew-name(P) in let pinew-body(P) in pi-new-trans(x;Q;λp.pi-trans(l_loc;p))
    if pioption?(P) then pi-guarded-trans(pi-choices(P);λp.pi-trans(l_loc;p))
    if picomm?(P) then pi-guarded-trans(pi-choices(P);λp.pi-trans(l_loc;p))
    else λl,serial,avoid. impossible
    fi 

Lemma: pi-trans_wf
[l_loc:Id]. ∀[P:pi_term()].  (pi-trans(l_loc;P) ∈ Id ⟶ Name ⟶ (Name List) ⟶ pi-process())

Definition: pi-system
pi-system(f;l_server;l_choose;l_comm;l_pi) ==
  λP.<[<l_server, loc-Server(f;[l_server; l_choose; l_comm; l_pi];1)>; <l_choose, processChoose(l_comm)>; <l_comm, proce\000CssComm(l_comm;l_choose)>; <l_pi, pi-trans(l_server;P) l_pi nat-to-incomparable(0) pi-names(P)>]
     make-lg([<<-1, l_pi>l_pi, "msg", pDVfire()>])
     >

Lemma: pi-system_wf
[f:(Id List) ⟶ Id]. ∀[l_server,l_choose,l_comm,l_pi:Id].
  (pi-system(f;l_server;l_choose;l_comm;l_pi) ∈ pi_term() ⟶ System(P.piM(P)))

Definition: pi-system-example
pi-system-example{$l_server,$l_choose,$l_comm,$l_pi,$a,$b,$c,$d,$e,$f,$g}(P) ==
  pi-system(λl.case apply-alist(NatDeq;ten-locs{$l_server,$l_choose,$l_comm,$l_pi,$a,$b,$c,$d,$e,$f,$g}();||l|| 4)
                of inl(x) =>
                x
                inr(x) =>
                "$g";"$l_server";"$l_choose";"$l_comm";"$l_pi") 
  P

Lemma: pi-system-example_wf
[P:pi_term()]. (pi-system-example{$a,$b,$c,$d,$e,$f,$g,$h,$i,$j,$k}(P) ∈ System(P.piM(P)))

Definition: pi-run-example
pi-run-example{$l_server,$l_choose,$l_comm,$l_pi,$a,$b,$c,$d,$e,$f,$g}(P;t) ==
  pRun2(pi-system-example{$l_server,$l_choose,$l_comm,$l_pi,$a,$b,$c,
  $d,$e,$f,$g}(P);std-env("$l_server");λn.pDVfire();λl.pDVfire();t)

Lemma: pi-run-example_wf
[P:pi_term()]. ∀[t:ℕ].
  (pi-run-example{$a,$b,$c,$d,$e,$f,$g,$h,$i,$j,$k}(P;t) ∈ {L:(ℤ × Id × Id × pMsg(P.piM(P))? × System(P.piM(P))) List| 
                                                            ||L|| (t 1) ∈ ℤ)

Definition: run-to-n
run-to-n(r;n) ==  map(λt.(fst((r t)));upto(n))

Lemma: run-to-n_wf
[B1,a1:Type]. ∀[r:ℤ ⟶ (a1 × B1)]. ∀[n:ℤ].  (run-to-n(r;n) ∈ a1 List)

Comment: things-that-can-be-run
let pi-run-example{l_server,l_choose,l_comm,l_pi,a,b,c,d,e,f,g}(pi-example([x];[y];[z];[v]);20) in
     map(λp.(fst(p));r)⌝

let pi-run-example{l_server,l_choose,l_comm,l_pi,a,b,c,d,e,f,g}(pi-example([x];[y];[z];[v]);1) in
     norm-list(λp.let a,b in eval a' in eval b' in   <a', b'>r)⌝

let pi-run-example{l_server,l_choose,l_comm,l_pi,a,b,c,d,e,f,g}(pi-example([x];[y];[z];[v]);2) in
     norm-list(norm-pair(λx.x;norm-lg(λx.norm-intransit(x)))) map(λx.let info,comps,intrans in <info, intrans>;r)⌝

Comment: more-things-that-can-be-run
let pi-run-example{l_server,l_choose,l_comm,l_pi,a,b,c,d,e,f,g}(pi-example([x];[y];[z];[v]);4) in
     norm-list(λx.let a,b,c in eval a' in eval c' in   <a', c'>r)⌝⋅

Comment: NOTES_ON_JUNK
What's in here is incoherent -- originally done for definition

  piM(T) PiDataValue (Id T)

(I think) and corresponding definition of pi-process.

The function for that, think, was

   E(T) 

something now forget.⋅

Lemma: loc-Server_wf
[f:(Id List) ⟶ Id]. ∀[L:Id List]. ∀[n:ℕ].  (loc-Server(f;L;n) ∈ Process(P.piM(P)))

Lemma: processChoose_wf
[l_comm:Id]. (processChoose(l_comm) ∈ Process(P.piM(P)))

Definition: select-indices-aux
select-indices-aux(f;L) ==
  rec-case(L) of
  [] => λn.[]
  fst::rest =>
   r.if fst then λn.[n (r (n 1))] else λn.(r (n 1)) fi 

Lemma: select-indices-aux_wf
[A:Type]. ∀[f:A ⟶ 𝔹]. ∀[L:A List].  (select-indices-aux(f;L) ∈ ℕ ⟶ (ℕ List))

Definition: select-indices
select-indices(f;L) ==  select-indices-aux(f;L) 0

Lemma: select-indices_wf
[A:Type]. ∀[f:A ⟶ 𝔹]. ∀[L:A List].  (select-indices(f;L) ∈ ℕ List)

Definition: accum-matching-indices
accum-matching-indices(pre;st) ==
  let = λp.prefix-match(p;pre) in
   let indices = λpreList.select-indices(P;preList) in
   let new_pair = λloc,preList. <loc, indices preList> in
   let = λrslt,loc,preList. [new_pair loc preList rslt] in
   fpf-accum(c,loc,preList.f loc preList;[];st)

Lemma: accum-matching-indices_wf
[pre:pi_prefix()]. ∀[st:a:Id fp-> pi_prefix() List].  (accum-matching-indices(pre;st) ∈ (Id × (ℕ List)) List)

Lemma: processComm_wf
[l_comm,l_choose:Id].  (processComm(l_comm;l_choose) ∈ Process(P.piM(P)))

Definition: class-value-has
X(e) has ==  (↑e ∈b X) ∧ a#X(e):T)

Lemma: class-value-has_wf
[Info:Type]. ∀[a:Atom1]. ∀[T:Type]. ∀[X:EClass(T)]. ∀[es:EO+(Info)]. ∀[e:E].  (X(e) has a ∈ ℙ)

Definition: sdata
SecurityData ==  tree(Id Atom1)

Lemma: sdata_wf
SecurityData ∈ Type

Definition: id-sdata
data(x) ==  tree_leaf(inl x)

Lemma: id-sdata_wf
[x:Id]. (data(x) ∈ SecurityData)

Definition: atom-sdata
data(a) ==  tree_leaf(inr )

Lemma: atom-sdata_wf
[a:Atom1]. (data(a) ∈ SecurityData)

Definition: sdata-pair
<d1, d2> ==  tree_node(d1;d2)

Lemma: sdata-pair_wf
[d1,d2:SecurityData].  (<d1, d2> ∈ SecurityData)

Definition: sdata-pair?
sdata-pair?(d) ==  tree_node?(d)

Lemma: sdata-pair?_wf
[p:SecurityData]. (sdata-pair?(p) ∈ 𝔹)

Lemma: sdata_pair_atom_lemma
a:Top. (sdata-pair?(data(a)) ff)

Lemma: sdata_pair_id_lemma
x:Top. (sdata-pair?(data(x)) ff)

Lemma: sdata_pair_rec_lemma
d2,d1:Top.  (sdata-pair?(<d1, d2>tt)

Definition: sdata-left
sdata-left(p) ==  tree_node-left(p)

Lemma: sdata-left_wf
[p:SecurityData]. sdata-left(p) ∈ SecurityData supposing ↑sdata-pair?(p)

Definition: sdata-right
sdata-right(p) ==  tree_node-right(p)

Lemma: sdata-right_wf
[p:SecurityData]. sdata-right(p) ∈ SecurityData supposing ↑sdata-pair?(p)

Lemma: sdata_left_pair_lemma
d2,d1:Top.  (sdata-left(<d1, d2>d1)

Lemma: sdata_right_pair_lemma
d2,d1:Top.  (sdata-right(<d1, d2>d2)

Lemma: sdata-pair-one-one
[x1,x2,y1,y2:SecurityData].
  uiff(<x1, x2> = <y1, y2> ∈ SecurityData;{(x1 y1 ∈ SecurityData) ∧ (x2 y2 ∈ SecurityData)})

Lemma: id-sdata-one-one
[a,b:Id].  uiff(data(a) data(b) ∈ SecurityData;a b ∈ Id)

Lemma: atom-sdata-one-one
[a,b:Atom1].  uiff(data(a) data(b) ∈ SecurityData;a b ∈ Atom1)

Lemma: atom-sdata-not-pair
[a:Atom1]. ∀[d1,d2:SecurityData].  uiff(data(a) = <d1, d2> ∈ SecurityData;False)

Lemma: id-sdata-not-pair
[a:Id]. ∀[d1,d2:SecurityData].  uiff(data(a) = <d1, d2> ∈ SecurityData;False)

Lemma: atom-sdata-not-pair2
[a:Atom1]. ∀[d1,d2:SecurityData].  uiff(<d1, d2> data(a) ∈ SecurityData;False)

Lemma: id-sdata-not-pair2
[a:Id]. ∀[d1,d2:SecurityData].  uiff(<d1, d2> data(a) ∈ SecurityData;False)

Lemma: sdata-test
[a:Atom1]. ∀[x,y:Id].  (<<data(a), data(y)>data(x)> ∈ SecurityData)

Lemma: sdata_subtype_base
SecurityData ⊆Base

Lemma: sdata_sq
SQType(SecurityData)

Lemma: atom-sdata-has-atom
[a,b:Atom1].  uiff(b#data(a):SecurityData;¬(a b ∈ Atom1))

Lemma: sdata-pair-has-atom
[d1,d2:SecurityData]. ∀[b:Atom1].  uiff(b#<d1, d2>:SecurityData;b#d1:SecurityData ∧ b#d2:SecurityData)

Lemma: id-sdata-has-atom
[x:Id]. ∀[b:Atom1].  uiff(b#data(x):SecurityData;True)

Definition: sdata-atoms
sdata-atoms(d) ==  tree_ind(d; tree_leaf(x) if ¬bisl(x) then [outr(x)] else [] fi tree_node(x,y) as,bs.as bs) 

Lemma: sdata-atoms_wf
[d:SecurityData]. (sdata-atoms(d) ∈ Atom1 List)

Lemma: sdata_atoms_pair_lemma
d2,d1:Top.  (sdata-atoms(<d1, d2>sdata-atoms(d1) sdata-atoms(d2))

Lemma: sdata_atoms_atom_lemma
a:Top. (sdata-atoms(data(a)) [a])

Lemma: sdata_atoms_id_lemma
x:Top. (sdata-atoms(data(x)) [])

Lemma: sdata-free-from-atom
[d:SecurityData]. ∀[a:Atom1].  uiff(a#d:SecurityData;¬(a ∈ sdata-atoms(d)))

Lemma: sdata-has-atom
d:SecurityData. ∀a:Atom1.  a#d:SecurityData ⇐⇒ (a ∈ sdata-atoms(d)))

Lemma: sdata-pair-free-from-atom
[b:Atom1]. ∀[d1,d2:SecurityData].  uiff(b#<d1, d2>:SecurityData;b#d1:SecurityData ∧ b#d2:SecurityData)

Definition: encryption-key
Key ==  Id Atom1 Atom1

Lemma: encryption-key_wf
Key ∈ Type

Definition: symmetric-key
symmetric-key(a) ==  inr inr a  

Lemma: symmetric-key_wf
[a:Atom1]. (symmetric-key(a) ∈ Key)

Lemma: encryption-key_sq
SQType(Key)

Definition: encryption-key-atoms
encryption-key-atoms(k) ==  case of inl(x) => [] inr(x) => case of inl(a) => [a] inr(a) => [a]

Lemma: encryption-key-atoms_wf
[k:Key]. (encryption-key-atoms(k) ∈ Atom1 List)

Lemma: encryption-key-free-from-atom
[k:Key]. ∀[a:Atom1].  uiff(¬(a ∈ encryption-key-atoms(k));a#k:Key)

Definition: security-event-structure
SES ==
  Info:Type
  × New:EClass(Atom1)
  × Send:EClass(SecurityData)
  × Rcv:EClass(SecurityData)
  × Encrypt:EClass(SecurityData × Key × Atom1)
  × Decrypt:EClass(SecurityData × Key × Atom1)
  × Sign:EClass(SecurityData × Id × Atom1)
  × Verify:EClass(SecurityData × Id × Atom1)
  × Honest:Id ⟶ 𝔹
  × KeyRel:Key ⟶ Key ⟶ 𝔹
  × PrivKey:Id ⟶ Atom1
  × Top

Lemma: security-event-structure_wf
SES ∈ 𝕌'

Definition: ses-info
Info ==  fst(s)

Lemma: ses-info_wf
[s:SES]. (Info ∈ Type)

Definition: ses-new
New ==  fst(snd(s))

Lemma: ses-new_wf
[s:SES]. (New ∈ EClass(Atom1))

Definition: ses-send
Send ==  fst(snd(snd(s)))

Lemma: ses-send_wf
[s:SES]. (Send ∈ EClass(SecurityData))

Definition: ses-rcv
Rcv ==  fst(snd(snd(snd(s))))

Lemma: ses-rcv_wf
[s:SES]. (Rcv ∈ EClass(SecurityData))

Definition: ses-sign
Sign ==  fst(snd(snd(snd(snd(snd(snd(s)))))))

Lemma: ses-sign_wf
[s:SES]. (Sign ∈ EClass(SecurityData × Id × Atom1))

Definition: ses-sig
signature(e) ==  snd(snd(Sign(e)))

Lemma: ses-sig_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Sign)].  (signature(e) ∈ Atom1)

Definition: ses-signer
signer(e) ==  fst(snd(Sign(e)))

Lemma: ses-signer_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Sign)].  (signer(e) ∈ Id)

Definition: ses-signed
signed(e) ==  fst(Sign(e))

Lemma: ses-signed_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Sign)].  (signed(e) ∈ SecurityData)

Definition: ses-verify
Verify ==  fst(snd(snd(snd(snd(snd(snd(snd(s))))))))

Lemma: ses-verify_wf
[s:SES]. (Verify ∈ EClass(SecurityData × Id × Atom1))

Definition: ses-verify-signed
signed(e) ==  fst(Verify(e))

Lemma: ses-verify-signed_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Verify)].  (signed(e) ∈ SecurityData)

Definition: ses-verify-signer
signer(e) ==  fst(snd(Verify(e)))

Lemma: ses-verify-signer_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Verify)].  (signer(e) ∈ Id)

Definition: ses-verify-sig
signature(e) ==  snd(snd(Verify(e)))

Lemma: ses-verify-sig_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Verify)].  (signature(e) ∈ Atom1)

Definition: ses-encrypt
Encrypt ==  fst(snd(snd(snd(snd(s)))))

Lemma: ses-encrypt_wf
[s:SES]. (Encrypt ∈ EClass(SecurityData × Key × Atom1))

Definition: ses-encrypted
plainText(e) ==  fst(Encrypt(e))

Definition: ses-encryption-key
key(e) ==  fst(snd(Encrypt(e)))

Lemma: ses-encryption-key_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Encrypt)].  (key(e) ∈ Key)

Lemma: ses-encrypted_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Encrypt)].  (plainText(e) ∈ SecurityData)

Definition: ses-crypt
cipherText(e) ==  snd(snd(Encrypt(e)))

Lemma: ses-crypt_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Encrypt)].  (cipherText(e) ∈ Atom1)

Definition: ses-decrypt
Decrypt ==  fst(snd(snd(snd(snd(snd(s))))))

Lemma: ses-decrypt_wf
[s:SES]. (Decrypt ∈ EClass(SecurityData × Key × Atom1))

Definition: ses-decrypted
plainText(e) ==  fst(Decrypt(e))

Lemma: ses-decrypted_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Decrypt)].  (plainText(e) ∈ SecurityData)

Definition: ses-decryption-key
key(e) ==  fst(snd(Decrypt(e)))

Lemma: ses-decryption-key_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Decrypt)].  (key(e) ∈ Key)

Definition: ses-cipher
cipherText(e) ==  snd(snd(Decrypt(e)))

Lemma: ses-cipher_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E(Decrypt)].  (cipherText(e) ∈ Atom1)

Definition: ses-honest
Honest(A) ==  ↑((fst(snd(snd(snd(snd(snd(snd(snd(snd(s)))))))))) A)

Lemma: ses-honest_wf
[s:SES]. ∀[A:Id].  (Honest(A) ∈ ℙ)

Lemma: ses-honest_witness
[s:SES]. ∀[A:Id].  (Honest(A)  (Ax ∈ Honest(A)))

Definition: ses-key-rel
MatchingKeys(k1;k2) ==  ↑((fst(snd(snd(snd(snd(snd(snd(snd(snd(snd(s))))))))))) k1 k2)

Lemma: ses-key-rel_wf
[s:SES]. ∀[k1,k2:Key].  (MatchingKeys(k1;k2) ∈ ℙ)

Lemma: ses-key-rel_witness
[s:SES]. ∀[k1,k2:Key].  (MatchingKeys(k1;k2)  (Ax ∈ MatchingKeys(k1;k2)))

Definition: ses-private
Private(A) ==  (fst(snd(snd(snd(snd(snd(snd(snd(snd(snd(snd(s)))))))))))) A

Lemma: ses-private_wf
[s:SES]. ∀[A:Id].  (Private(A) ∈ Atom1)

Definition: ses-private-key
PrivateKey(A) ==  inr (inl Private(A)) 

Lemma: ses-private-key_wf
[s:SES]. ∀[A:Id].  (PrivateKey(A) ∈ Key)

Definition: ses-public-key
PublicKey(A) ==  inl A

Lemma: ses-public-key_wf
[A:Id]. (PublicKey(A) ∈ Key)

Lemma: ses-public-not-private
[s:SES]. ∀[A,B:Id].  uiff(PublicKey(A) PrivateKey(B) ∈ Key;False)

Lemma: ses-private-not-public
[s:SES]. ∀[A,B:Id].  uiff(PrivateKey(B) PublicKey(A) ∈ Key;False)

Lemma: ses-public-one-one
[A,B:Id].  uiff(PublicKey(A) PublicKey(B) ∈ Key;A B ∈ Id)

Lemma: ses-public-key-atoms
[A:Id]. (encryption-key-atoms(PublicKey(A)) [])

Lemma: ses-private-key-atoms
[s:SES]. ∀[A:Id].  (encryption-key-atoms(PrivateKey(A)) [Private(A)])

Lemma: symmetric-key-atoms
[s:SES]. ∀[a:Atom1].  (encryption-key-atoms(symmetric-key(a)) [a])

Lemma: ses-sign-has-atom
s:SES. ∀es:EO+(Info). ∀e:E. ∀a:Atom1.
  (Sign(e) has ⇐⇒ (↑e ∈b Sign) ∧ ((a signature(e) ∈ Atom1) ∨ (a ∈ sdata-atoms(signed(e)))))

Definition: ses-action
Action(e) ==
  (↑e ∈b New) ∨ (↑e ∈b Send) ∨ (↑e ∈b Rcv) ∨ (↑e ∈b Encrypt) ∨ (↑e ∈b Decrypt) ∨ (↑e ∈b Sign) ∨ (↑e ∈b Verify)

Lemma: ses-action_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E].  (Action(e) ∈ ℙ)

Lemma: decidable__ses-action
s:SES. ∀es:EO+(Info). ∀e:E.  Dec(Action(e))

Definition: ses-act
Act ==  {e:E| Action(e)} 

Lemma: ses-act_wf
[s:SES]. ∀[es:EO+(Info)].  (Act ∈ Type)

Definition: event-has
(e has a) ==
  New(e) has a ∨ Send(e) has a ∨ Rcv(e) has a ∨ Encrypt(e) has a ∨ Decrypt(e) has a ∨ Sign(e) has a ∨ Verify(e) has a

Lemma: event-has_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E]. ∀[a:Atom1].  ((e has a) ∈ ℙ)

Definition: same-action
same-action(x;y) ==
  ((↑x ∈b New) ∧ (↑y ∈b New) ∧ (New(x) New(y) ∈ Atom1))
  ∨ ((↑x ∈b Send) ∧ (↑y ∈b Send) ∧ (Send(x) Send(y) ∈ SecurityData))
  ∨ ((↑x ∈b Rcv) ∧ (↑y ∈b Rcv) ∧ (Rcv(x) Rcv(y) ∈ SecurityData))
  ∨ ((↑x ∈b Encrypt) ∧ (↑y ∈b Encrypt) ∧ (Encrypt(x) Encrypt(y) ∈ (SecurityData × Key × Atom1)))
  ∨ ((↑x ∈b Decrypt) ∧ (↑y ∈b Decrypt) ∧ (Decrypt(x) Decrypt(y) ∈ (SecurityData × Key × Atom1)))
  ∨ ((↑x ∈b Sign) ∧ (↑y ∈b Sign) ∧ (Sign(x) Sign(y) ∈ (SecurityData × Id × Atom1)))
  ∨ ((↑x ∈b Verify) ∧ (↑y ∈b Verify) ∧ (Verify(x) Verify(y) ∈ (SecurityData × Id × Atom1)))

Lemma: same-action_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[x,y:E].  (same-action(x;y) ∈ ℙ)

Definition: ses-flow
ses-flow(s;es;a;e1;e2) ==
  
  ses-flow,a,e1,e2. (((e1 has a) ∧ (e2 has a) ∧ e1 ≤loc e2 )
                     ∨ (∃snd:E(Send)
                         ∃rcv:E(Rcv)
                          (e1 c≤ snd
                          ∧ rcv c≤ e2
                          ∧ (snd < rcv)
                          ∧ (Send(snd) Rcv(rcv) ∈ SecurityData)
                          ∧ (ses-flow e1 snd)
                          ∧ (ses-flow rcv e2)))
                     ∨ (∃encr:E(Encrypt)
                         ∃decr:E(Decrypt)
                          ((e1 <loc encr)
                          ∧ decr c≤ e2
                          ∧ (encr < decr)
                          ∧ ((plainText(decr) plainText(encr) ∈ SecurityData)
                            ∧ (cipherText(decr) cipherText(encr) ∈ Atom1)
                            ∧ MatchingKeys(key(decr);key(encr)))
                          ∧ (key(decr) symmetric-key(a) ∈ Key))
                          ∧ (ses-flow e1 encr)
                          ∧ (ses-flow cipherText(encr) encr decr)
                          ∧ (ses-flow decr e2))))) 
  
  e1 
  e2

Lemma: ses-flow_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e1,e2:E]. ∀[a:Atom1].  (ses-flow(s;es;a;e1;e2) ∈ ℙ)

Lemma: ses-flow-induction
s:SES. ∀es:EO+(Info).
  ∀[P:E ⟶ E ⟶ Atom1 ⟶ ℙ]
    ((∀e1,e2:E.
        ((∀x,y:E.
            (e1 c≤  c≤  c≤ e2  ((e1 < x) ∨ (y < e2))  (∀a:Atom1. (ses-flow(s;es;a;x;y)  P[x;y;a]))))
         (∀a:Atom1. (ses-flow(s;es;a;e1;e2)  P[e1;e2;a]))))
     {∀e1,e2:E. ∀a:Atom1.  (ses-flow(s;es;a;e1;e2)  P[e1;e2;a])})

Lemma: ses-flow-implies'
s:SES. ∀es:EO+(Info).
  ∀[P:E ⟶ E ⟶ Atom1 ⟶ ℙ]
    ((∀e1,e2:E. ∀a:Atom1.  (((e1 has a) ∧ (e2 has a) ∧ e1 ≤loc e2  P[e1;e2;a]))
     (∀e1,e2:E. ∀snd:E(Send). ∀rcv:E(Rcv). ∀a:Atom1.
          ((e1 c≤ snd
          ∧ rcv c≤ e2
          ∧ (snd < rcv)
          ∧ (Send(snd) Rcv(rcv) ∈ SecurityData)
          ∧ (P[e1;snd;a] ∧ ses-flow(s;es;a;e1;snd))
          ∧ P[rcv;e2;a]
          ∧ ses-flow(s;es;a;rcv;e2))
           P[e1;e2;a]))
     (∀e1,e2:E. ∀encr:E(Encrypt). ∀decr:E(Decrypt). ∀a:Atom1.
          (((e1 <loc encr)
          ∧ decr c≤ e2
          ∧ (encr < decr)
          ∧ ((plainText(decr) plainText(encr) ∈ SecurityData)
            ∧ (cipherText(decr) cipherText(encr) ∈ Atom1)
            ∧ MatchingKeys(key(decr);key(encr)))
          ∧ (key(decr) symmetric-key(a) ∈ Key))
          ∧ (P[e1;encr;a] ∧ ses-flow(s;es;a;e1;encr))
          ∧ (P[encr;decr;cipherText(encr)] ∧ ses-flow(s;es;cipherText(encr);encr;decr))
          ∧ P[decr;e2;a]
          ∧ ses-flow(s;es;a;decr;e2))
           P[e1;e2;a]))
     {∀e1,e2:E. ∀a:Atom1.  (ses-flow(s;es;a;e1;e2)  P[e1;e2;a])})

Lemma: ses-flow-implies
s:SES. ∀es:EO+(Info).
  ∀[P:E ⟶ E ⟶ Atom1 ⟶ ℙ]
    ((∀e1,e2:E. ∀a:Atom1.  (((e1 has a) ∧ (e2 has a) ∧ e1 ≤loc e2  P[e1;e2;a]))
     (∀e1,e2:E. ∀snd:E(Send). ∀rcv:E(Rcv). ∀a:Atom1.
          ((e1 c≤ snd ∧ rcv c≤ e2 ∧ (snd < rcv) ∧ (Send(snd) Rcv(rcv) ∈ SecurityData) ∧ P[e1;snd;a] ∧ P[rcv;e2;a])
           P[e1;e2;a]))
     (∀e1,e2:E. ∀encr:E(Encrypt). ∀decr:E(Decrypt). ∀a:Atom1.
          (((e1 <loc encr)
          ∧ decr c≤ e2
          ∧ (encr < decr)
          ∧ ((plainText(decr) plainText(encr) ∈ SecurityData)
            ∧ (cipherText(decr) cipherText(encr) ∈ Atom1)
            ∧ MatchingKeys(key(decr);key(encr)))
          ∧ (key(decr) symmetric-key(a) ∈ Key))
          ∧ P[e1;encr;a]
          ∧ P[encr;decr;cipherText(encr)]
          ∧ P[decr;e2;a])
           P[e1;e2;a]))
     {∀e1,e2:E. ∀a:Atom1.  (ses-flow(s;es;a;e1;e2)  P[e1;e2;a])})

Lemma: ses-flow-has
s:SES. ∀es:EO+(Info). ∀e1,e2:E. ∀a:Atom1.  (ses-flow(s;es;a;e1;e2)  (e2 has a))

Definition: ses-info-flow
->> ==  λe',e. ((↑e' ∈b Encrypt) ∧ (e has cipherText(e')))

Lemma: ses-info-flow_wf
[s:SES]. ∀[es:EO+(Info)].  (->> ∈ E ⟶ E ⟶ ℙ)

Definition: event-has*
has* ==  ∃e':E. ((e' has a) ∧ (e' (->>^*) e))

Lemma: event-has*_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E]. ∀[a:Atom1].  (e has* a ∈ ℙ)

Lemma: event-has*-iff
s:SES. ∀es:EO+(Info). ∀e:E. ∀a:Atom1.  (e has* ⇐⇒ (e has a) ∨ (∃z:E. ((z ->> e) ∧ has* a)))

Lemma: ses-flow-causle
s:SES. ∀es:EO+(Info). ∀e1,e2:E. ∀a:Atom1.  (ses-flow(s;es;a;e1;e2)  e1 c≤ e2)

Lemma: event-has*-transitive-encrypt
s:SES. ∀es:EO+(Info). ∀a:Atom1. ∀e1,e2:E.  e1 has*  e2 has* cipherText(e1)  e2 has* supposing ↑e1 ∈b Encrypt

Lemma: ses-flow-has*
s:SES. ∀es:EO+(Info). ∀e1,e2:E. ∀a:Atom1.
  (ses-flow(s;es;a;e1;e2)  ((e1 ≤loc e2  ∧ (e2 has a)) ∨ (∃snd:E(Send). (e1 ≤loc snd  ∧ (snd < e2) ∧ snd has* a))))

Definition: ses-secrecy
Only agents in As have atoms in secrets ==  ∀e:E. ((∃a∈secrets. (e has a))  (loc(e) ∈ As))

Lemma: ses-secrecy_wf
[s:SES]. ∀[agents:Id List]. ∀[secrets:Atom1 List]. ∀[es:EO+(Info)].  (Only agents in agents have atoms in secrets ∈ ℙ)

Definition: ses-cabal
cabal is cabal for atms ==
  (∀a∈atms.(∃i∈cabal. ∃e:E
                       ((loc(e) i ∈ Id)
                       ∧ (((↑e ∈b New) ∧ (a New(e) ∈ Atom1))
                         ∨ ((↑e ∈b Encrypt)
                           ∧ (a cipherText(e) ∈ Atom1)
                           ∧ (∃b∈atms. (b a ∈ Atom1)) ∧ (e has b))))))
     ∧ (∀e:E(Send). ((loc(e) ∈ cabal)  (e has a))))
     ∧ (∀e:E(Encrypt)
          ((loc(e) ∈ cabal)
           (e has a)
           ((∃A∈cabal. key(e) PublicKey(A) ∈ Key)
             ∨ (∃k∈atms. key(e) symmetric-key(k) ∈ Key)
             ∨ (cipherText(e) ∈ atms)))))

Lemma: ses-cabal_wf
[s:SES]. ∀[C:Id List]. ∀[as:Atom1 List]. ∀[es:EO+(Info)].  (C is cabal for as ∈ ℙ)

Definition: release-before
(a released before e) ==  ∃e':E(Send). ((e' <loc e) ∧ e' has* a)

Lemma: release-before_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[a:Atom1]. ∀[e:E].  ((a released before e) ∈ ℙ)

Definition: ses-ordering
PropertyO ==
  ∀es:EO+(Info)
    ((∀n:E(New). ∀e:E.  ((e has New(n))  (n ≤loc e  ∨ (∃snd:E(Send). ((n <loc snd) ∧ (snd < e) ∧ snd has* New(n))))))
    ∧ (∀e1:E(Sign). ∀b:E.
         ((b has signature(e1))
          (∃e2:E(Sign)
              ((Sign(e2) Sign(e1) ∈ (SecurityData × Id × Atom1))
              ∧ (e2 ≤loc b  ∨ (∃snd:E(Send). ((e2 <loc snd) ∧ (snd < b) ∧ snd has* signature(e1))))))))
    ∧ (∀e1:E(Encrypt). ∀b:E.
         ((b has cipherText(e1))
          (∃e2:E(Encrypt)
              ((Encrypt(e2) Encrypt(e1) ∈ (SecurityData × Key × Atom1))
              ∧ (e2 ≤loc b  ∨ (∃snd:E(Send). ((e2 <loc snd) ∧ (snd < b) ∧ snd has* cipherText(e1)))))))))

Lemma: ses-ordering_wf
[s:SES]. (PropertyO ∈ ℙ')

Definition: ses-ordering'
ses-ordering'(s) ==
  ∀es:EO+(Info). ∀b:E.
    ((∀n:E(New). (b has* New(n)  (n ≤loc b  ∨ (∃snd:E(Send). ((n <loc snd) ∧ (snd < b) ∧ snd has* New(n))))))
    ∧ (∀e1:E(Sign)
         (b has* signature(e1)
          (∃e2:E(Sign)
              ((Sign(e2) Sign(e1) ∈ (SecurityData × Id × Atom1))
              ∧ (e2 ≤loc b  ∨ (∃snd:E(Send). ((e2 <loc snd) ∧ (snd < b) ∧ snd has* signature(e1))))))))
    ∧ (∀e1:E(Encrypt)
         (b has* cipherText(e1)
          (∃e2:E(Encrypt)
              ((Encrypt(e2) Encrypt(e1) ∈ (SecurityData × Key × Atom1))
              ∧ (e2 ≤loc b  ∨ (∃snd:E(Send). ((e2 <loc snd) ∧ (snd < b) ∧ snd has* cipherText(e1)))))))))

Lemma: ses-ordering'_wf
[s:SES]. (ses-ordering'(s) ∈ ℙ')

Definition: ses-flow-axiom
PropertyF ==
  ∀es:EO+(Info)
    ((∀n:E(New). ∀e:E.  ((e has New(n))  ses-flow(s;es;New(n);n;e)))
    ∧ (∀e1:E(Sign). ∀b:E.
         ((b has signature(e1))
          (∃e2:E(Sign). ((Sign(e2) Sign(e1) ∈ (SecurityData × Id × Atom1)) ∧ ses-flow(s;es;signature(e2);e2;b)))))
    ∧ (∀e1:E(Encrypt). ∀b:E.
         ((b has cipherText(e1))
          (∃e2:E(Encrypt)
              ((Encrypt(e2) Encrypt(e1) ∈ (SecurityData × Key × Atom1)) ∧ ses-flow(s;es;cipherText(e2);e2;b))))))

Lemma: ses-flow-axiom_wf
[s:SES]. (PropertyF ∈ ℙ')

Definition: ses-nonce
PropertyN ==
  ∀es:EO+(Info). ∀n:E(New). ∀e:E.
    (e has* New(n)
     (n c≤ e ∧ ((¬(loc(e) loc(n) ∈ Id))  (∃e':E(Send). ((n <loc e') ∧ e' has* New(n) ∧ (e' < e))))))

Lemma: ses-nonce_wf
[s:SES]. (PropertyN ∈ ℙ')

Definition: ses-NU
PropertyNU ==  ∀es:EO+(Info). ∀n,e:E(New).  ((New(e) New(n) ∈ Atom1)  (e n ∈ E))

Lemma: ses-NU_wf
[s:SES]. (PropertyNU ∈ ℙ')

Lemma: ses-nonce-unique
[s:SES]. PropertyNU supposing PropertyN

Lemma: ses-signature-unique
[s:SES]
  ∀[es:EO+(Info)]. ∀[a,b:E(Sign)].
    Sign(a) Sign(b) ∈ (SecurityData × Id × Atom1) supposing signature(a) signature(b) ∈ Atom1 
  supposing PropertyO

Lemma: ses-signature-unique2
[s:SES]
  ∀[es:EO+(Info)]. ∀[a,b:E(Sign)].
    {(signed(a) signed(b) ∈ SecurityData) ∧ (signer(a) signer(b) ∈ Id)} 
    supposing signature(a) signature(b) ∈ Atom1 
  supposing PropertyO

Lemma: ses-cipher-unique
[s:SES]
  ∀[es:EO+(Info)]. ∀[a,b:E(Encrypt)].
    Encrypt(a) Encrypt(b) ∈ (SecurityData × Key × Atom1) supposing cipherText(a) cipherText(b) ∈ Atom1 
  supposing PropertyO

Lemma: ses-cipher-unique2
[s:SES]
  ∀[es:EO+(Info)]. ∀[a,b:E(Encrypt)].
    {(plainText(a) plainText(b) ∈ SecurityData) ∧ (key(a) key(b) ∈ Key)} 
    supposing cipherText(a) cipherText(b) ∈ Atom1 
  supposing PropertyO

Definition: ses-NR
PropertyNR ==
  ∀es:EO+(Info). ∀n:E(New). ∀e:E.
    ((loc(e) loc(n) ∈ Id)
     (New(n) released before e))
     (∀e':E. ((¬(loc(e') loc(e) ∈ Id))  e' has* New(n)  (e < e'))))

Lemma: ses-NR_wf
[s:SES]. (PropertyNR ∈ ℙ')

Lemma: ses-nonce-release
s:SES. (PropertyN  PropertyNR)

Definition: ses-V
PropertyV ==
  ∀es:EO+(Info). ∀e:E(Verify).  ∃e':E(Sign). ((e' < e) ∧ (Verify(e) Sign(e') ∈ (SecurityData × Id × Atom1)))

Lemma: ses-V_wf
[s:SES]. (PropertyV ∈ ℙ')

Definition: ses-R
PropertyR ==  ∀es:EO+(Info). ∀e:E(Rcv).  ∃e':E(Send). ((e' < e) ∧ (Rcv(e) Send(e') ∈ SecurityData))

Lemma: ses-R_wf
[s:SES]. (PropertyR ∈ ℙ')

Definition: ses-D
PropertyD ==
  ∀es:EO+(Info). ∀e:E(Decrypt).
    ∃e':E(Encrypt)
     ((e' < e)
     ∧ (plainText(e) plainText(e') ∈ SecurityData)
     ∧ (cipherText(e) cipherText(e') ∈ Atom1)
     ∧ MatchingKeys(key(e);key(e')))

Lemma: ses-D_wf
[s:SES]. (PropertyD ∈ ℙ')

Definition: ses-S
PropertyS ==
  ∀es:EO+(Info)
    ((∀A:Id. (Honest(A)  (∀e:E(Sign). ((signer(e) A ∈ Id)  (loc(e) A ∈ Id)))))
    ∧ (∀A:Id
         (Honest(A)
          ((∀e:E(Encrypt). ((key(e) PrivateKey(A) ∈ Key)  (loc(e) A ∈ Id)))
            ∧ (∀e:E(Decrypt). ((key(e) PrivateKey(A) ∈ Key)  (loc(e) A ∈ Id)))))))

Lemma: ses-S_wf
[s:SES]. (PropertyS ∈ ℙ')

Definition: ses-K
PropertyK ==
  Sym(Key;k1,k2.MatchingKeys(k1;k2))
  ∧ (∀A:Id. ∀k:Key.  (MatchingKeys(PrivateKey(A);k) ⇐⇒ PublicKey(A) ∈ Key))
  ∧ (∀A:Id. ∀k:Key.  (MatchingKeys(PublicKey(A);k) ⇐⇒ PrivateKey(A) ∈ Key))
  ∧ (∀A,B:Id.  (Private(A) Private(B) ∈ Atom1 ⇐⇒ B ∈ Id))
  ∧ (∀a:Atom1. ∀k:Key.  (MatchingKeys(symmetric-key(a);k) ⇐⇒ symmetric-key(a) ∈ Key))

Lemma: ses-K_wf
[s:SES]. (PropertyK ∈ ℙ)

Definition: ses-disjoint
ActionsDisjoint ==
  ∀es:EO+(Info)
    ∃f:E ⟶ ℤ
     ∀e:E
       (((↑e ∈b New)  ((f e) 1 ∈ ℤ))
       ∧ ((↑e ∈b Send)  ((f e) 2 ∈ ℤ))
       ∧ ((↑e ∈b Rcv)  ((f e) 3 ∈ ℤ))
       ∧ ((↑e ∈b Sign)  ((f e) 4 ∈ ℤ))
       ∧ ((↑e ∈b Verify)  ((f e) 5 ∈ ℤ))
       ∧ ((↑e ∈b Encrypt)  ((f e) 6 ∈ ℤ))
       ∧ ((↑e ∈b Decrypt)  ((f e) 7 ∈ ℤ)))

Lemma: ses-disjoint_wf
[ses:SES]. (ActionsDisjoint ∈ ℙ')

Definition: ses-disjoint-old
ses-disjoint-old{i:l}(ses) ==
  ∀es:EO+(Info). ∀e:E.
    (((↑e ∈b New)
      ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Send)
       ((((((¬↑e ∈b New) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Rcv)
       ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b New)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Sign)
       ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b New)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Verify)
       ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b New)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Encrypt)
       ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b New)) ∧ (¬↑e ∈b Decrypt)))
    ∧ ((↑e ∈b Decrypt)
       ((((((¬↑e ∈b Send) ∧ (¬↑e ∈b Rcv)) ∧ (¬↑e ∈b Sign)) ∧ (¬↑e ∈b Verify)) ∧ (¬↑e ∈b Encrypt)) ∧ (¬↑e ∈b New))))

Lemma: ses-flow-axiom-ordering
s:SES. (ActionsDisjoint  PropertyF  PropertyO)

Definition: ses-nonce-disjoint
NoncesCiphersAndKeysDisjoint ==
  ∀es:EO+(Info)
    ((∀n:E(New). ∀e:E(Sign).  (New(n) signature(e) ∈ Atom1)))
    ∧ (∀n:E(New). ∀e:E(Encrypt).  (New(n) cipherText(e) ∈ Atom1)))
    ∧ (∀n:E(New). ∀A:Id.  (New(n) Private(A) ∈ Atom1)))
    ∧ (∀A:Id. ∀e:E(Encrypt).  (cipherText(e) Private(A) ∈ Atom1)))
    ∧ (∀A:Id. ∀e:E(Sign).  (signature(e) Private(A) ∈ Atom1)))
    ∧ (∀e1:E(Sign). ∀e2:E(Encrypt).  (signature(e1) cipherText(e2) ∈ Atom1))))

Lemma: ses-nonce-disjoint_wf
[ses:SES]. (NoncesCiphersAndKeysDisjoint ∈ ℙ')

Lemma: ses-private-one-one
[s:SES]. ∀[A,B:Id].  uiff(PrivateKey(A) PrivateKey(B) ∈ Key;A B ∈ Id) supposing PropertyK

Definition: ses-axioms
SecurityAxioms ==
  ActionsDisjoint ∧ NoncesCiphersAndKeysDisjoint ∧ PropertyF ∧ PropertyV ∧ PropertyR ∧ PropertyD ∧ PropertyS ∧ PropertyK

Lemma: ses-axioms_wf
[s:SES]. (SecurityAxioms ∈ ℙ')

Definition: ses-D-public
ses-D-public(s) ==
  ∀es:EO+(Info). ∀e:E(Decrypt). ∀A:Id. ∀d:SecurityData. ∀c:Atom1.
    ((Decrypt(e) = <d, PublicKey(A), c> ∈ (SecurityData × Key × Atom1))
     (∃e':E(Encrypt). ((e' < e) ∧ (Encrypt(e') = <d, PrivateKey(A), c> ∈ (SecurityData × Key × Atom1)))))

Lemma: ses-D-public_wf
[s:SES]. (ses-D-public(s) ∈ ℙ')

Definition: ses-D-private
ses-D-private(s) ==
  ∀es:EO+(Info). ∀e:E(Decrypt). ∀A:Id. ∀d:SecurityData. ∀c:Atom1.
    ((Decrypt(e) = <d, PrivateKey(A), c> ∈ (SecurityData × Key × Atom1))
     (∃e':E(Encrypt). ((e' < e) ∧ (Encrypt(e') = <d, PublicKey(A), c> ∈ (SecurityData × Key × Atom1)))))

Lemma: ses-D-private_wf
[s:SES]. (ses-D-private(s) ∈ ℙ')

Lemma: ses-D-implies
s:SES. (PropertyD  {ses-D-public(s) ∧ ses-D-private(s)} supposing PropertyK)

Lemma: event-has_functionality
s:SES. (ActionsDisjoint  (∀es:EO+(Info). ∀x,y:E.  (same-action(x;y)  (∀a:Atom1. ((x has a)  (y has a))))))

Lemma: ses-info-flow_functionality
s:SES. (ActionsDisjoint  (∀es:EO+(Info). ∀x,y:E.  (same-action(x;y)  (∀e:E. ((e ->> x)  (e ->> y))))))

Lemma: ses-info-flow-exp_functionality
s:SES
  (ActionsDisjoint  (∀es:EO+(Info). ∀x,y:E.  (same-action(x;y)  (∀e:E. ∀n:ℕ+.  ((e ->>^n x)  (e ->>^n y))))))

Lemma: ses-ordering-ordering'
s:SES. (PropertyO  ActionsDisjoint  ses-ordering'(s))

Lemma: ses-nonce-from-ordering
s:SES. (ActionsDisjoint  PropertyO  PropertyN)

Lemma: ses-ordering-implies
s:SES. (ActionsDisjoint  PropertyO  {PropertyNU ∧ PropertyNR})

Lemma: ses-axioms-imply
s:SES. (SecurityAxioms  {ses-ordering'(s) ∧ PropertyNR ∧ PropertyNU})

Definition: security-theory
SecurityTheory ==  ses:SES × A:SecurityAxioms × Top

Lemma: security-theory_wf
SecurityTheory ∈ ℙ'

Definition: sth-es
sth-es(s) ==  fst(s)

Lemma: sth-es_wf
[sth:SecurityTheory]. (sth-es(sth) ∈ SES)

Lemma: sth-axioms
sth:SecurityTheory. SecurityAxioms

Definition: ses-msg
isMsg(e) ==  e ∈b Send ∨be ∈b Rcv

Lemma: ses-msg_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E].  (isMsg(e) ∈ 𝔹)

Definition: thread-messages
thread-messages(thr) ==  filter(λe.isMsg(e);thr)

Lemma: thread-messages_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[thr:E List].  (thread-messages(thr) ∈ {e:E| (↑e ∈b Send) ∨ (↑e ∈b Rcv)}  List)

Lemma: ses-msg-cases
[ses:SES]
  ∀[es:EO+(Info)]
    ((∀[e:E]. isMsg(e) ff supposing ↑e ∈b Sign)
    ∧ (∀[e:E]. isMsg(e) ff supposing ↑e ∈b Verify)
    ∧ (∀[e:E]. isMsg(e) ff supposing ↑e ∈b Encrypt)
    ∧ (∀[e:E]. isMsg(e) ff supposing ↑e ∈b Decrypt)
    ∧ (∀[e:E]. isMsg(e) ff supposing ↑e ∈b New)
    ∧ (∀[e:E]. isMsg(e) tt supposing ↑e ∈b Send)
    ∧ (∀[e:E]. isMsg(e) tt supposing ↑e ∈b Rcv)) 
  supposing ActionsDisjoint

Definition: send-rcv-match
send-rcv-match(m1;m2) ==  (↑m1 ∈b Send) ∧ (↑m2 ∈b Rcv) ∧ (m1 < m2) ∧ (Send(m1) Rcv(m2) ∈ SecurityData)

Lemma: send-rcv-match_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[m1,m2:E].  (send-rcv-match(m1;m2) ∈ ℙ)

Definition: matching-conversation
MatchingConversation(n;thr1;thr2) ==
  (n ≤ ||thread-messages(thr1)||)
  ∧ (n ≤ ||thread-messages(thr2)||)
  ∧ (∀p∈zip(firstn(n;thread-messages(thr1));firstn(n;thread-messages(thr2))).let m1,m2 
                                                                             in send-rcv-match(m1;m2)
                                                                                ∨ send-rcv-match(m2;m1))

Definition: ses-thread
Thread ==  {thr:Act List| ∀i:ℕ||thr|| 1. (thr[i] <loc thr[i 1])} 

Lemma: ses-thread_wf
[s:SES]. ∀[es:EO+(Info)].  (Thread ∈ Type)

Lemma: ses-thread-subtype
[s:SES]. ∀[es:EO+(Info)].  (Thread ⊆(E List))

Lemma: ses-thread-order
s:SES. ∀es:EO+(Info). ∀thr:Thread. ∀i,j:ℕ||thr||.  (thr[i] <loc thr[j]) supposing i < j

Lemma: ses-thread-no_repeats
[s:SES]. ∀[es:EO+(Info)]. ∀[thr:Thread].  no_repeats(Act;thr)

Lemma: ses-thread-weak-order
s:SES. ∀es:EO+(Info). ∀thr:Thread. ∀i,j:ℕ||thr||.  thr[i] ≤loc thr[j]  supposing i ≤ j

Lemma: matching-conversation_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[thr1,thr2:Thread]. ∀[n:ℕ].  (MatchingConversation(n;thr1;thr2) ∈ ℙ)

Definition: ses-thread-member
e ∈ thr ==  ∃i:ℕ||thr||. (e thr[i] ∈ E)

Lemma: ses-thread-member_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:Act]. ∀[thr:Thread].  (e ∈ thr ∈ ℙ)

Definition: ses-thread-loc
loc(thr)= ==  ∀e:Act. (e ∈ thr  (loc(e) A ∈ Id))

Lemma: ses-thread-loc_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[thr:Thread]. ∀[A:Id].  (loc(thr)= A ∈ ℙ)

Definition: ses-used-atoms
UsedAtoms(e) ==
  if e ∈b Send then sdata-atoms(Send(e))
  if e ∈b Sign then sdata-atoms(signed(e))
  if e ∈b Verify then [signature(e) sdata-atoms(signed(e))]
  if e ∈b Encrypt then sdata-atoms(plainText(e)) encryption-key-atoms(key(e))
  if e ∈b Decrypt then [cipherText(e) encryption-key-atoms(key(e))]
  else []
  fi 

Lemma: ses-used-atoms_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E].  (UsedAtoms(e) ∈ Atom1 List)

Definition: ses-useable-atoms
UseableAtoms(e) ==
  if e ∈b Rcv then sdata-atoms(Rcv(e))
  if e ∈b Sign then [signature(e)]
  if e ∈b Encrypt then [cipherText(e)]
  if e ∈b Decrypt then sdata-atoms(plainText(e))
  if e ∈b New then [New(e)]
  else []
  fi 

Lemma: ses-useable-atoms_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[e:E].  (UseableAtoms(e) ∈ Atom1 List)

Lemma: member-useable-atoms
s:SES
  (ActionsDisjoint
   (∀es:EO+(Info). ∀e:E. ∀a:Atom1.
        ((a ∈ UseableAtoms(e))
        ⇐⇒ ((↑e ∈b Rcv) ∧ (a ∈ sdata-atoms(Rcv(e))))
            ∨ ((↑e ∈b Decrypt) ∧ (a ∈ sdata-atoms(plainText(e))))
            ∨ ((↑e ∈b New) ∧ (a New(e) ∈ Atom1))
            ∨ ((↑e ∈b Sign) ∧ (a signature(e) ∈ Atom1))
            ∨ ((↑e ∈b Encrypt) ∧ (a cipherText(e) ∈ Atom1)))))

Lemma: member-used-atoms
s:SES
  (ActionsDisjoint
   (∀es:EO+(Info). ∀e:E. ∀a:Atom1.
        ((a ∈ UsedAtoms(e))
        ⇐⇒ ((↑e ∈b Send) ∧ (a ∈ sdata-atoms(Send(e))))
            ∨ ((↑e ∈b Decrypt) ∧ ((a cipherText(e) ∈ Atom1) ∨ (a ∈ encryption-key-atoms(key(e)))))
            ∨ ((↑e ∈b Verify) ∧ ((a signature(e) ∈ Atom1) ∨ (a ∈ sdata-atoms(signed(e)))))
            ∨ ((↑e ∈b Sign) ∧ (a ∈ sdata-atoms(signed(e))))
            ∨ ((↑e ∈b Encrypt) ∧ ((a ∈ sdata-atoms(plainText(e))) ∨ (a ∈ encryption-key-atoms(key(e))))))))

Lemma: ses-act-has-atom
s:SES
  (ActionsDisjoint  (∀es:EO+(Info). ∀e:Act. ∀a:Atom1.  ((e has a) ⇐⇒ (a ∈ UseableAtoms(e)) ∨ (a ∈ UsedAtoms(e)))))

Definition: ses-legal-thread
Legal(thr)@A ==  ∀i:ℕ||thr||. UsedAtoms(thr[i]) ⊆ [Private(A) concat(map(λj.UseableAtoms(thr[j]);[0, i)))]

Lemma: ses-legal-thread_wf
[s:SES]. ∀[A:Id]. ∀[es:EO+(Info)]. ∀[thr:Thread].  (Legal(thr)@A ∈ ℙ)

Lemma: ses-legal-thread-has-atom
s:SES
  (ActionsDisjoint
   (∀es:EO+(Info). ∀A:Id. ∀thr:Thread.
        (Legal(thr)@A
         (∀a:Atom1. ∀e:Act.
              (e ∈ thr
               (e has a)
               ((∃e':Act. ((e' <loc e) ∧ (e' has a) ∧ e' ∈ thr))
                 ∨ (a ∈ UseableAtoms(e))
                 ∨ (a Private(A) ∈ Atom1)))))))

Lemma: ses-legal-thread-decrypt
s:SES
  (ActionsDisjoint
   PropertyD
      (∀es:EO+(Info). ∀A:Id. ∀thr:Thread. ∀e:E.
           (Legal(thr)@A
            e ∈ thr  (∃a:Act. ((a <loc e) ∧ a ∈ thr ∧ (a has cipherText(e)))) supposing ↑e ∈b Decrypt)) 
     supposing NoncesCiphersAndKeysDisjoint)

Lemma: ses-legal-thread-has*-nonce
s:SES
  (SecurityAxioms
   (∀es:EO+(Info). ∀A:Id. ∀thr:Thread.
        (Legal(thr)@A
         (∀n:E(New). ∀e:Act.
              (e ∈ thr
               has* New(n)
               (∃e':E. ((e' <loc e) ∧ Action(e') ∧ e' has* New(n) ∧ e' ∈ thr)) ∨ (e n ∈ E) 
                 supposing ∀e':E
                             ((e' < e)  Action(e')  e' has* New(n)  ((loc(e') A ∈ Id) ∧ (¬↑e' ∈b Send))))))))

Definition: noncelike-signatures
noncelike-signatures(s;es;thr) ==
  ∀e1,e2:E(Sign).  ((signature(e1) signature(e2) ∈ Atom1)  e1 ∈ thr  (e1 e2 ∈ E))

Lemma: noncelike-signatures_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[thr:Thread].  (noncelike-signatures(s;es;thr) ∈ ℙ)

Lemma: ses-legal-thread-has*-signature
s:SES
  (SecurityAxioms
   (∀es:EO+(Info). ∀A:Id. ∀thr:Thread.
        ((Legal(thr)@A ∧ noncelike-signatures(s;es;thr))
         (∀sg:E(Sign). ∀e:Act.
              (e ∈ thr
               has* signature(sg)
               (∃e':E. ((e' <loc e) ∧ Action(e') ∧ e' has* signature(sg) ∧ e' ∈ thr)) ∨ (e sg ∈ E) 
                 supposing ∀e':E
                             ((e' < e)
                              Action(e')
                              e' has* signature(sg)
                              ((loc(e') A ∈ Id) ∧ (¬↑e' ∈b Send))))))))

Definition: protocol-action
ProtocolAction ==
  n:Atom × if =a "new" then Atom1
           if =a "send" then SecurityData
           if =a "rcv" then SecurityData
           if =a "encrypt" then SecurityData × Key × Atom1
           if =a "decrypt" then SecurityData × Key × Atom1
           if =a "sign" then SecurityData × Id × Atom1
           if =a "verify" then SecurityData × Id × Atom1
           else Top
           fi 

Lemma: protocol-action_wf
ProtocolAction ∈ Type

Definition: pa-is-sign-implies
pa-is-sign-implies(a;v.P[v]) ==  ((fst(a)) "sign" ∈ Atom)  P[snd(a)]

Lemma: pa-is-sign-implies_wf
a:ProtocolAction. ∀P:(SecurityData × Id × Atom1) ⟶ ℙ.  (pa-is-sign-implies(a;v.P[v]) ∈ ℙ)

Lemma: decidable__pa-is-sign-implies
a:ProtocolAction. ∀P:(SecurityData × Id × Atom1) ⟶ ℙ.
  ((∀v:SecurityData × Id × Atom1. Dec(P[v]))  Dec(pa-is-sign-implies(a;v.P[v])))

Definition: pa-is-new-and
pa-is-new-and(a;v.P[v]) ==  ((fst(a)) "new" ∈ Atom) ∧ P[snd(a)]

Lemma: pa-is-new-and_wf
a:ProtocolAction. ∀P:Atom1 ⟶ ℙ.  (pa-is-new-and(a;v.P[v]) ∈ ℙ)

Lemma: decidable__pa-is-new-and
a:ProtocolAction. ∀P:Atom1 ⟶ ℙ.  ((∀v:Atom1. Dec(P[v]))  Dec(pa-is-new-and(a;v.P[v])))

Definition: pa-used
pa-used(pa) ==
  let n,v pa 
  in if =a "send" then sdata-atoms(v)
     if =a "encrypt" then sdata-atoms(fst(v)) encryption-key-atoms(fst(snd(v)))
     if =a "sign" then sdata-atoms(fst(v))
     if =a "verify" then [snd(snd(v)) sdata-atoms(fst(v))]
     if =a "decrypt" then [snd(snd(v)) encryption-key-atoms(fst(snd(v)))]
     else []
     fi 

Lemma: pa-used_wf
[pa:ProtocolAction]. (pa-used(pa) ∈ Atom1 List)

Definition: pa-useable
pa-useable(pa) ==
  let n,v pa 
  in if =a "new" then [v]
     if =a "rcv" then sdata-atoms(v)
     if =a "encrypt" then [snd(snd(v))]
     if =a "decrypt" then sdata-atoms(fst(v))
     if =a "sign" then [snd(snd(v))]
     else []
     fi 

Lemma: pa-useable_wf
[pa:ProtocolAction]. (pa-useable(pa) ∈ Atom1 List)

Definition: ses-legal-sequence
Legal(pas) given prvt ==  ∀i:ℕ||pas||. pa-used(pas[i]) ⊆ [prvt concat(map(λj.pa-useable(pas[j]);[0, i)))]

Lemma: ses-legal-sequence_wf
[prvt:Atom1]. ∀[pas:ProtocolAction List].  (Legal(pas) given prvt ∈ ℙ)

Lemma: decidable__ses-legal-sequence
prvt:Atom1. ∀pas:ProtocolAction List.  Dec(Legal(pas) given prvt)

Definition: ses-is-legal
ses-is-legal(prvt; pas) ==  isl(TERMOF{decidable__ses-legal-sequence:o, 1:l} prvt pas)

Lemma: ses-is-legal_wf
[prvt:Atom1]. ∀[pas:ProtocolAction List].  (ses-is-legal(prvt; pas) ∈ 𝔹)

Lemma: assert-ses-is-legal
prvt:Atom1. ∀pas:ProtocolAction List.  (↑ses-is-legal(prvt; pas) ⇐⇒ Legal(pas) given prvt)

Definition: ses-fresh-sequence
ses-fresh-sequence(f;A;pas) ==
  ∀i:ℕ||pas||
    (((fst(pas[i])) "sign" ∈ Atom)
     (((fst(snd(snd(pas[i])))) A ∈ Id)
       ∧ (∃j:ℕi
           (((fst(pas[j])) "new" ∈ Atom)
           ∧ ((↑isl(f (fst(snd(pas[i]))))) ∧ (outl(f (fst(snd(pas[i])))) (snd(pas[j])) ∈ Atom1))
           ∧ (∀k:{j 1..i-}
                (((fst(pas[k])) "sign" ∈ Atom)
                 (↑isl(f (fst(snd(pas[k])))))
                 (outl(f (fst(snd(pas[k])))) (snd(pas[j])) ∈ Atom1))))))))

Lemma: ses-fresh-sequence_wf
[f:SecurityData ⟶ (Atom1?)]. ∀[A:Id]. ∀[pas:ProtocolAction List].  (ses-fresh-sequence(f;A;pas) ∈ ℙ)

Lemma: decidable-ses-fresh-sequence
f:SecurityData ⟶ (Atom1?). ∀A:Id. ∀pas:ProtocolAction List.  Dec(ses-fresh-sequence(f;A;pas))

Lemma: debug1
m:ℤDec(∀i:ℕm. ∃j:ℕi. True)

Lemma: debug2
m:ℤDec(∀i:ℕm. 0 < i)

Definition: ses-is-fresh
ses-is-fresh(f;A;pas) ==  isl(TERMOF{decidable-ses-fresh-sequence:o, 1:l} pas)

Lemma: ses-is-fresh_wf
[f:SecurityData ⟶ (Atom1?)]. ∀[A:Id]. ∀[pas:ProtocolAction List].  (ses-is-fresh(f;A;pas) ∈ 𝔹)

Lemma: assert-ses-is-fresh
f:SecurityData ⟶ (Atom1?). ∀A:Id. ∀pas:ProtocolAction List.  (↑ses-is-fresh(f;A;pas) ⇐⇒ ses-fresh-sequence(f;A;pas))

Definition: mk-pa
n(v) ==  <n, v>

Lemma: mk-pa_wf
(∀[a:Atom1]. ("new"(a) ∈ ProtocolAction))
∧ (∀[d:SecurityData]. ("send"(d) ∈ ProtocolAction))
∧ (∀[d:SecurityData]. ("rcv"(d) ∈ ProtocolAction))
∧ (∀[tr:SecurityData × Key × Atom1]. ("encrypt"(tr) ∈ ProtocolAction))
∧ (∀[tr:SecurityData × Key × Atom1]. ("decrypt"(tr) ∈ ProtocolAction))
∧ (∀[tr:SecurityData × Id × Atom1]. ("sign"(tr) ∈ ProtocolAction))
∧ (∀[tr:SecurityData × Id × Atom1]. ("verify"(tr) ∈ ProtocolAction))

Definition: ses-is-protocol-action
pa(e) ==
  let n,v pa 
  in if =a "new" then (↑e ∈b New) ∧ (New(e) v ∈ Atom1)
     if =a "send" then (↑e ∈b Send) ∧ (Send(e) v ∈ SecurityData)
     if =a "rcv" then (↑e ∈b Rcv) ∧ (Rcv(e) v ∈ SecurityData)
     if =a "encrypt" then (↑e ∈b Encrypt) ∧ (Encrypt(e) v ∈ (SecurityData × Key × Atom1))
     if =a "decrypt" then (↑e ∈b Decrypt) ∧ (Decrypt(e) v ∈ (SecurityData × Key × Atom1))
     if =a "sign" then (↑e ∈b Sign) ∧ (Sign(e) v ∈ (SecurityData × Id × Atom1))
     if =a "verify" then (↑e ∈b Verify) ∧ (Verify(e) v ∈ (SecurityData × Id × Atom1))
     else False
     fi 

Lemma: ses-is-protocol-action_wf
[s:SES]. ∀[pa:ProtocolAction]. ∀[es:EO+(Info)]. ∀[e:E].  (pa(e) ∈ ℙ)

Lemma: ses-sign-is-protocol-action
[s:SES]
  ∀[pa:ProtocolAction]. ∀[es:EO+(Info)]. ∀[e:E].
    ({((fst(pa)) "sign" ∈ Atom) ∧ ((snd(pa)) Sign(e) ∈ (SecurityData × Id × Atom1))}) supposing 
       ((↑e ∈b Sign) and 
       pa(e)) 
  supposing ActionsDisjoint

Lemma: ses-is-protocol-action-used
[s:SES]
  ∀[pa:ProtocolAction]. ∀[es:EO+(Info)]. ∀[e:E].  UsedAtoms(e) pa-used(pa) ∈ (Atom1 List) supposing pa(e) 
  supposing ActionsDisjoint

Lemma: ses-is-protocol-action-useable
[s:SES]
  ∀[pa:ProtocolAction]. ∀[es:EO+(Info)]. ∀[e:E].  UseableAtoms(e) pa-useable(pa) ∈ (Atom1 List) supposing pa(e) 
  supposing ActionsDisjoint

Definition: ses-is-protocol-actions
pas(thr) ==  (||thr|| ||pas|| ∈ ℤ) ∧ (∀i:ℕ||pas||. pas[i](thr[i]))

Lemma: ses-is-protocol-actions_wf
[s:SES]. ∀[pas:ProtocolAction List]. ∀[es:EO+(Info)]. ∀[thr:Thread].  (pas(thr) ∈ ℙ)

Lemma: ses-is-protocol-actions-legal
s:SES
  (ActionsDisjoint
   (∀pas:ProtocolAction List. ∀es:EO+(Info). ∀thr:Thread. ∀A:Id.
        (pas(thr)  Legal(pas) given Private(A)  Legal(thr)@A)))

Definition: ses-basic-sequence1
Basic1 ==  Id ⟶ Id ⟶ es:EO+(Info) ⟶ Thread ⟶ ℙ

Lemma: ses-basic-sequence1_wf
[s:SES]. (Basic1 ∈ ℙ')

Definition: is-basic-seq
thr[A;B] |= bs ==  bs es thr

Lemma: is-basic-seq_wf
[s:SES]. ∀[bs:Basic1]. ∀[A,B:Id]. ∀[es:EO+(Info)]. ∀[thr:Thread].  (thr[A;B] |= bs ∈ ℙ)

Definition: authentication
prtcl |= bs authenticates messages  ==
  ∀es:EO+(Info). ∀A,B:Id.
    ((Honest(A) ∧ Honest(B) ∧ (A B ∈ Id)) ∧ (prtcl A) ∧ (prtcl B))
     (∀thr1:Thread
          (loc(thr1)=  thr1[A;B] |= bs  (∃thr2:Thread. (MatchingConversation(n;thr1;thr2) ∧ loc(thr2)= B)))))

Lemma: authentication_wf
[s:SES]. ∀[prtcl:Id ⟶ ℙ']. ∀[bs:Basic1]. ∀[n:ℕ].  (prtcl |= bs authenticates messages  ∈ ℙ')

Definition: basic-seq-1-1
A ⟶ B: pas[A; B; m] ==  λA,B,es,thr. ∃m:Atom1. pas[A; B; m](thr)

Lemma: basic-seq-1-1_wf
[s:SES]. ∀[pas:Id ⟶ Id ⟶ Atom1 ⟶ (ProtocolAction List)].  (A ⟶ B: pas[A;B;m] ∈ Basic1)

Definition: basic-seq-1-2
A ⟶ B: pas[A; B; m; n] ==  λA,B,es,thr. ∃m,n:Atom1. pas[A; B; m; n](thr)

Lemma: basic-seq-1-2_wf
[s:SES]. ∀[pas:Id ⟶ Id ⟶ Atom1 ⟶ Atom1 ⟶ (ProtocolAction List)].  (A ⟶ B: pas[A;B;m;n] ∈ Basic1)

Definition: basic-seq-1-3
A ⟶ B: pas[A; B; m; n; x] ==  λA,B,es,thr. ∃m,n,x:Atom1. pas[A; B; m; n; x](thr)

Lemma: basic-seq-1-3_wf
[s:SES]. ∀[pas:Id ⟶ Id ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ (ProtocolAction List)].  (A ⟶ B: pas[A;B;m;n;x] ∈ Basic1)

Definition: basic-seq-1-4
A ⟶ B: pas[A;B;m;n;x;y] ==  λA,B,es,thr. ∃m,n,x,y:Atom1. pas[A;B;m;n;x;y](thr)

Lemma: basic-seq-1-4_wf
[s:SES]. ∀[pas:Id ⟶ Id ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ (ProtocolAction List)].
  (A ⟶ B: pas[A;B;m;n;x;y] ∈ Basic1)

Definition: basic-seq-1-5
A ⟶ B: pas[A;B;m;n;x;y;z] ==  λA,B,es,thr. ∃m,n,x,y,z:Atom1. pas[A;B;m;n;x;y;z](thr)

Lemma: basic-seq-1-5_wf
[s:SES]. ∀[pas:Id ⟶ Id ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ Atom1 ⟶ (ProtocolAction List)].
  (A ⟶ B: pas[A;B;m;n;x;y;z] ∈ Basic1)

Definition: ses-protocol1-thread
(thr is one of bss at A) ==  ∃B:Id. (∃bs∈bss. thr[A;B] |= bs)

Lemma: ses-protocol1-thread_wf
[s:SES]. ∀[bss:Basic1 List]. ∀[A:Id]. ∀[es:EO+(Info)]. ∀[thr:Thread].  ((thr is one of bss at A) ∈ ℙ')

Definition: ses-protocol1
Protocol1(bss) ==
  λA.∀es:EO+(Info)
       (Honest(A)
        (∀e:Act
             ((loc(e) A ∈ Id)
              ((∃thr:Thread. (e ∈ thr ∧ (thr is one of bss at A) ∧ loc(thr)= A))
                ∧ (∀thr1,thr2:Thread.
                     (e ∈ thr1
                      (thr1 is one of bss at A)
                      e ∈ thr2
                      (thr2 is one of bss at A)
                      (thr1 ≤ thr2 ∨ thr2 ≤ thr1)))))))

Lemma: ses-protocol1_wf
[s:SES]. ∀[bss:Basic1 List].  (Protocol1(bss) ∈ Id ⟶ ℙ')

Definition: ses-protocol1-legal
Legal(bss) ==  ActionsDisjoint  (∀A:Id. ∀es:EO+(Info). ∀thr:Thread.  ((thr is one of bss at A)  Legal(thr)@A))

Lemma: ses-protocol1-legal_wf
[s:SES]. ∀[bss:Basic1 List].  (Legal(bss) ∈ ℙ')

Definition: ses-fresh-thread
ses-fresh-thread(s;es;f;A;thr) ==
  ∀i:ℕ||thr||
    ((↑thr[i] ∈b Sign)
     ((signer(thr[i]) A ∈ Id)
       ∧ (∃j:ℕi
           ((↑thr[j] ∈b New)
           ∧ (↑isl(f signed(thr[i])))
           ∧ (outl(f signed(thr[i])) New(thr[j]) ∈ Atom1)
           ∧ (∀k:{j 1..i-}
                ((↑thr[k] ∈b Sign)
                 (↑isl(f signed(thr[k])))
                 (outl(f signed(thr[k])) New(thr[j]) ∈ Atom1))))))))

Lemma: ses-fresh-thread_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[f:SecurityData ⟶ (Atom1?)]. ∀[A:Id]. ∀[thr:Thread].  (ses-fresh-thread(s;es;f;A;thr) ∈ ℙ)

Lemma: ses-is-protocol-actions-fresh
s:SES
  (ActionsDisjoint
   (∀pas:ProtocolAction List. ∀es:EO+(Info). ∀thr:Thread. ∀A:Id.
        (pas(thr)  (∀f:SecurityData ⟶ (Atom1?). (ses-fresh-sequence(f;A;pas)  ses-fresh-thread(s;es;f;A;thr))))))

Definition: fresh-sig-protocol1
FreshSignatures(bss) ==
  ∃f:SecurityData ⟶ (Atom1?)
   ∀es:EO+(Info). ∀thr:Thread. ∀A:Id.  ((thr is one of bss at A)  ses-fresh-thread(s;es;f;A;thr))

Lemma: fresh-sig-protocol1_wf
[s:SES]. ∀[bss:Basic1 List].  (FreshSignatures(bss) ∈ ℙ')

Definition: unique-sig-protocol
UniqueSignatures(bss) ==
  ∀A:Id
    (Honest(A)
     (Protocol1(bss) A)
     (∀es:EO+(Info). ∀thr:Thread.  ((thr is one of bss at A)  noncelike-signatures(s;es;thr))))

Lemma: unique-sig-protocol_wf
[s:SES]. ∀[bss:Basic1 List].  (UniqueSignatures(bss) ∈ ℙ')

Lemma: fresh-sig-protocol1_property
[s:SES]
  (∀[bss:Basic1 List]. UniqueSignatures(bss) supposing FreshSignatures(bss)) supposing 
     (PropertyS and 
     PropertyO and 
     ActionsDisjoint)

Lemma: nonce-release-lemma
[ses:SES]
  ∀[bss:Basic1 List]
    ∀[A:Id]
      (∀[es:EO+(Info)]. ∀[thr:Thread].
         (∀[i:ℕ||thr||]. ∀[j:ℕi].
            (New(thr[j]) released before thr[i])) supposing 
               ((∀k:{j 1..i-}. (¬↑thr[k] ∈b Send)) and 
               (↑thr[j] ∈b New))) supposing 
            (loc(thr)= and 
            (thr is one of bss at A))) supposing 
         ((Protocol1(bss) A) and 
         Honest(A)) 
    supposing Legal(bss) 
  supposing SecurityAxioms

Lemma: signature-release-lemma
[ses:SES]
  ∀[bss:Basic1 List]
    ∀[A:Id]
      (∀[es:EO+(Info)]. ∀[thr:Thread].
         (∀[i:ℕ||thr||]. ∀[j:ℕi].
            (signature(thr[j]) released before thr[i])) supposing 
               ((∀k:{j 1..i-}. (¬↑thr[k] ∈b Send)) and 
               (↑thr[j] ∈b Sign))) supposing 
            (loc(thr)= and 
            (thr is one of bss at A))) supposing 
         ((Protocol1(bss) A) and 
         Honest(A)) 
    supposing Legal(bss) ∧ UniqueSignatures(bss) 
  supposing SecurityAxioms

Lemma: nonce-release-lemma2
[s:SecurityTheory]. ∀[bss:Basic1 List].
  ∀[A:Id]
    (∀[es:EO+(Info)]. ∀[thr:Thread].
       (∀[i:ℕ||thr||]. ∀[j:ℕi].
          (New(thr[j]) released before thr[i])) supposing 
             ((∀k:{j 1..i-}. (¬↑thr[k] ∈b Send)) and 
             (↑thr[j] ∈b New))) supposing 
          (loc(thr)= and 
          (thr is one of bss at A))) supposing 
       ((Protocol1(bss) A) and 
       Honest(A)) 
  supposing Legal(bss)

Definition: sig-release-thread
sig-release-thread(s;es;A;thr) ==
  ∀i:ℕ||thr||. ∀j:ℕi.
    ((↑thr[j] ∈b Sign)
     (∀k:{j 1..i-}. (¬↑thr[k] ∈b Send))
     (∀e:E. ((¬(loc(e) A ∈ Id))  (e has signature(thr[j]))  (thr[i] < e))))

Lemma: sig-release-thread_wf
[s:SES]. ∀[es:EO+(Info)]. ∀[thr:Thread]. ∀[A:Id].  (sig-release-thread(s;es;A;thr) ∈ ℙ)

Lemma: signature-release-lemma2
s:SecurityTheory. ∀bss:Basic1 List.
  ((Legal(bss) ∧ FreshSignatures(bss))
   (∀A:Id
        (Protocol1(bss) A)
         (∀es:EO+(Info). ∀thr:Thread.
              ((thr is one of bss at A)  sig-release-thread(sth-es(s);es;A;thr) supposing loc(thr)= A)) 
        supposing Honest(A)))

Definition: CR-initiator1
CR-initiator1{i:l}(s) ==  A ⟶ B: ["new"(m); "send"(data(m))]

Definition: CR-initiator2
CR-initiator2{i:l}(s) ==  A ⟶ B: ["new"(m); "send"(data(m))] ["rcv"(<data(n), data(s)>)]

Lemma: CR-initiator2_wf
[s:SES]. (CR-initiator2{i:l}(s) ∈ Basic1)

Lemma: CR-initiator1_wf
[s:SES]. (CR-initiator1{i:l}(s) ∈ Basic1)

Definition: CR-initiator3
CR-initiator3 ==
  A ⟶ B: (["new"(m); "send"(data(m))] ["rcv"(<data(n), data(s)>)])
          ["verify"(<<data(n), <data(A), data(m)>>B, s>); "sign"(<<data(m), <data(B), data(n)>>A, s'>); "send"(dat\000Ca(s'))]

Lemma: CR-initiator3_wf
[s:SES]. (CR-initiator3 ∈ Basic1)

Definition: CR-responder1
CR-responder1{i:l}(s) ==  B ⟶ A: ["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(A), data(m)>>B, s>); "send"(<data\000C(n), data(s)>)]

Lemma: CR-responder1_wf
[s:SES]. (CR-responder1{i:l}(s) ∈ Basic1)

Definition: CR-responder2
CR-responder2{i:l}(s) ==
  B ⟶ A: ["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(A), data(m)>>B, s>); "send"(<data(n), data(s)>)] ["rcv"\000C(data(s'))]

Lemma: CR-responder2_wf
[s:SES]. (CR-responder2{i:l}(s) ∈ Basic1)

Definition: CR-responder3
CR-responder3 ==
  B ⟶ A: (["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(A), data(m)>>B, s>); "send"(<data(n), data(s)>)] ["rcv\000C"(data(s'))])
          ["verify"(<<data(m), <data(B), data(n)>>A, s'>)]

Lemma: CR-responder3_wf
[s:SES]. (CR-responder3 ∈ Basic1)

Definition: CR-protocol
CR-protocol ==
  Protocol1([CR-initiator1{i:l}(s);
             CR-initiator2{i:l}(s);
             CR-initiator3;
             CR-responder1{i:l}(s);
             CR-responder2{i:l}(s);
             CR-responder3])

Lemma: CR-protocol_wf
[ses:SES]. (CR-protocol ∈ Id ⟶ ℙ')

Lemma: CR-protocol-legal
s:SES
  Legal([CR-initiator1{i:l}(s);
         CR-initiator2{i:l}(s);
         CR-initiator3;
         CR-responder1{i:l}(s);
         CR-responder2{i:l}(s);
         CR-responder3])

Lemma: CR-protocol-fresh
s:SES
  (ActionsDisjoint
   FreshSignatures([CR-initiator1{i:l}(s);
                      CR-initiator2{i:l}(s);
                      CR-initiator3;
                      CR-responder1{i:l}(s);
                      CR-responder2{i:l}(s);
                      CR-responder3]))

Lemma: CR-authentication-theorem
sth:SecurityTheory
  (CR-protocol |= CR-initiator3 authenticates messages  ∧ CR-protocol |= CR-responder3 authenticates messages )

Definition: CR-initiator4
CR-initiator4{i:l}(s) ==
  A ⟶ B: (["new"(m); "send"(data(m))] ["rcv"(<data(n), data(s)>)])
          ["verify"(<<data(n), <data(m), data(A)>>B, s>); "new"(k); "sign"(<<data(n), <data(m), <data(k), data(B)>>>\000C, A, s'>); "send"(<data(k), data(s')>)]

Lemma: CR-initiator4_wf
[s:SES]. (CR-initiator4{i:l}(s) ∈ Basic1)

Definition: CR-responder6
CR-responder6{i:l}(s) ==
  B ⟶ A: (["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(m), data(A)>>B, s>); "send"(<data(n), data(s)>)] ["rcv\000C"(<data(k), data(s')>)])
          ["verify"(<<data(n), <data(m), <data(k), data(B)>>>A, s'>)]

Lemma: CR-responder6_wf
[s:SES]. (CR-responder6{i:l}(s) ∈ Basic1)

Definition: CR-responder5
CR-responder5{i:l}(s) ==
  B ⟶ A: ["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(m), data(A)>>B, s>); "send"(<data(n), data(s)>)] ["rcv"\000C(<data(k), data(s')>)]

Lemma: CR-responder5_wf
[s:SES]. (CR-responder5{i:l}(s) ∈ Basic1)

Definition: CR-responder4
CR-responder4{i:l}(s) ==  B ⟶ A: ["rcv"(data(m)); "new"(n); "sign"(<<data(n), <data(m), data(A)>>B, s>); "send"(<data\000C(n), data(s)>)]

Lemma: CR-responder4_wf
[s:SES]. (CR-responder4{i:l}(s) ∈ Basic1)

Definition: CRX-protocol
CRX-protocol{i:l}(s) ==
  Protocol1([CR-initiator1{i:l}(s);
             CR-initiator2{i:l}(s);
             CR-initiator4{i:l}(s);
             CR-responder4{i:l}(s);
             CR-responder5{i:l}(s);
             CR-responder6{i:l}(s)])

Lemma: CRX-protocol_wf
[ses:SES]. (CRX-protocol{i:l}(ses) ∈ Id ⟶ ℙ')

Lemma: CRX-protocol-legal
s:SES
  Legal([CR-initiator1{i:l}(s);
         CR-initiator2{i:l}(s);
         CR-initiator4{i:l}(s);
         CR-responder4{i:l}(s);
         CR-responder5{i:l}(s);
         CR-responder6{i:l}(s)])

Lemma: CRX-authentication-theorem
sth:SecurityTheory
  (CRX-protocol{i:l}(sth-es(sth)) |= CR-initiator4{i:l}(sth-es(sth)) authenticates messages 
  ∧ CRX-protocol{i:l}(sth-es(sth)) |= CR-responder6{i:l}(sth-es(sth)) authenticates messages )

Lemma: testnlrst
es:EO. ∀A:Type.  ((A ⊆E)  (∀a,b,c:A.  ((a < b)  (b < c)  (c < a)  False)))

Lemma: cabal-theorem
s:SecurityTheory. ∀es:EO+(Info). ∀C:Id List.
  ∀ns:Atom1 List. (C is cabal for ns  Only agents in have atoms in ns) supposing (∀A∈C.Honest(A))

Definition: NSL-initiator1
NSL-initiator1{i:l}(s) ==  A ⟶ B: ["new"(m); "encrypt"(<<data(m), data(A)>PublicKey(B), c>); "send"(data(c))]

Lemma: NSL-initiator1_wf
[s:SES]. (NSL-initiator1{i:l}(s) ∈ Basic1)

Definition: NSL-initiator2
NSL-initiator2{i:l}(s) ==  A ⟶ B: ["new"(m); "encrypt"(<<data(m), data(A)>PublicKey(B), c>); "send"(data(c))] ["rcv\000C"(data(c'))]

Lemma: NSL-initiator2_wf
[s:SES]. (NSL-initiator2{i:l}(s) ∈ Basic1)

Definition: NSL-initiator3
NSL-initiator3 ==
  A ⟶ B: (["new"(m); "encrypt"(<<data(m), data(A)>PublicKey(B), c>); "send"(data(c))] ["rcv"(data(c'))])
          ["decrypt"(<<data(m), <data(n), data(B)>>PrivateKey(A), c'>); "encrypt"(<data(n), PublicKey(B), d>); "send\000C"(data(d))]

Definition: NSL-responder3
NSL-responder3 ==
  B ⟶ A: (["rcv"(data(c));
            "decrypt"(<<data(m), data(A)>PrivateKey(B), c>);
            "new"(n);
            "encrypt"(<<data(m), <data(n), data(B)>>PublicKey(A), c'>);
            "send"(data(c'))]
          ["rcv"(data(d))])
          ["decrypt"(<data(n), PrivateKey(B), d>)]

Lemma: NSL-responder3_wf
[s:SES]. (NSL-responder3 ∈ Basic1)

Definition: NSL-responder2
NSL-responder2{i:l}(s) ==
  B ⟶ A: ["rcv"(data(c));
           "decrypt"(<<data(m), data(A)>PrivateKey(B), c>);
           "new"(n);
           "encrypt"(<<data(m), <data(n), data(B)>>PublicKey(A), c'>);
           "send"(data(c'))]
          ["rcv"(data(d))]

Lemma: NSL-responder2_wf
[s:SES]. (NSL-responder2{i:l}(s) ∈ Basic1)

Lemma: NSL-initiator3_wf
[s:SES]. (NSL-initiator3 ∈ Basic1)

Definition: NSL-responder0
NSL-responder0{i:l}(s) ==  B ⟶ A: ["rcv"(data(c))]

Lemma: NSL-responder0_wf
[s:SES]. (NSL-responder0{i:l}(s) ∈ Basic1)

Definition: NSL-responder1
NSL-responder1{i:l}(s) ==
  B ⟶ A: ["rcv"(data(c));
           "decrypt"(<<data(m), data(A)>PrivateKey(B), c>);
           "new"(n);
           "encrypt"(<<data(m), <data(n), data(B)>>PublicKey(A), c'>);
           "send"(data(c'))]

Lemma: NSL-responder1_wf
[s:SES]. (NSL-responder1{i:l}(s) ∈ Basic1)

Definition: NSL-protocol
NSL-protocol ==
  Protocol1([NSL-initiator1{i:l}(s);
             NSL-initiator2{i:l}(s);
             NSL-initiator3;
             NSL-responder0{i:l}(s);
             NSL-responder1{i:l}(s);
             NSL-responder2{i:l}(s);
             NSL-responder3])

Lemma: NSL-protocol_wf
[ses:SES]. (NSL-protocol ∈ Id ⟶ ℙ')

Lemma: NSL-protocol-legal
s:SES
  Legal([NSL-initiator1{i:l}(s);
         NSL-initiator2{i:l}(s);
         NSL-initiator3;
         NSL-responder0{i:l}(s);
         NSL-responder1{i:l}(s);
         NSL-responder2{i:l}(s);
         NSL-responder3])

Definition: NSL-authentication-property
NSL-authentication-property{i:l}(sth) ==
  NSL-protocol |= NSL-initiator3 authenticates messages  ∧ NSL-protocol |= NSL-responder3 authenticates messages 

Lemma: NSL-authentication-property_wf
[sth:SecurityTheory]. (NSL-authentication-property{i:l}(sth) ∈ ℙ')

Definition: oar-deliver
Class OARDeliver outputs (i.e. delivers) the
sequenced message ⌜<sndr, n, msg>⌝ at correct location
only if it has verified that 2*f+1 members of the list
of orderers have signed ⌜<sndr, n, msg>⌝.
Also, it does not deliver ⌜<sndr, n, msg>⌝ unless it has previously
delivered messages ⌜<sndr, n', msg'>⌝ for all n' < n.⋅

oar-deliver(es;M;V;correct;orderers;f;OARDeliver) ==
  ∀n:ℕ. ∀sndr:Id. ∀msg:M. ∀e:E.
    (<sndr, n, msg> ∈ OARDeliver(e)
     (correct loc(e))
     {(0 <  (∃msg':M. ∃e':E. ((e' <loc e) ∧ <sndr, 1, msg'> ∈ OARDeliver(e))))
       ∧ (∃L:Id List
           ((||L|| ((2 f) 1) ∈ ℤ)
           ∧ no_repeats(Id;L)
           ∧ (∀o∈L.o ↓∈ orderers ∧ (∃e':E(V). (e' ≤loc e  ∧ (∃sig:Atom1. <o, sig, sndr, n, msg> ∈ V(e')))))))})

Lemma: oar-deliver_wf
[Info:Type]. ∀[es:EO+(Info)].
  ∀M:Type. ∀V:EClass(Id × Atom1 × Id × ℕ × M). ∀correct:Id ⟶ ℙ. ∀f:ℕ. ∀orderers:bag(Id).
  ∀OARDeliver:EClass(Id × ℕ × M).
    (oar-deliver(es;M;V;correct;orderers;f;OARDeliver) ∈ ℙ)

Definition: oar-order
correct orderer ⌜x ↓∈ orderers⌝ signs ⌜<sndr, n, msg>⌝ only if it has 
previously signed messages ⌜<sndr, n', msg'>⌝ for all n' < n, and
it has never previously signed ⌜<sndr, n, msg>⌝and
it does not simultaneously sign any other msg for sequence number 
and sender sndr.⋅

oar-order(es;M;S;correct;orderers) ==
  ∀n:ℕ. ∀sndr:Id. ∀msg:M. ∀sig:Atom1. ∀x:Id. ∀e:E.
    ((x ↓∈ orderers ∧ <x, sig, sndr, n, msg> ∈ S(e) ∧ (correct x))
     ((0 <  (∃msg':M. ∃sig':Atom1. ∃e':E. ((e' <loc e) ∧ <x, sig', sndr, 1, msg'> ∈ S(e'))))
       ∧ (∀msg':M. ∀sig':Atom1. ∀e':E.  ((e' <loc e)  (¬<x, sig', sndr, n, msg'> ∈ S(e'))))
       ∧ (∀msg':M. ∀sig':Atom1.  (<x, sig', sndr, n, msg'> ∈ S(e)  (msg' msg ∈ M)))))

Lemma: oar-order_wf
[Info:Type]. ∀[es:EO+(Info)].
  ∀M:Type. ∀S:EClass(Id × Atom1 × Id × ℕ × M). ∀correct:Id ⟶ ℙ. ∀orderers:bag(Id).
    (oar-order(es;M;S;correct;orderers) ∈ ℙ)

Definition: oar-crypto
The sign and verify classes and satisfy the assumptions that
for every V-event there was previous matching S-event.
And the location of correct S-event is correct.⋅

oar-crypto(es;M;V;S;correct) ==
  (∀n:ℕ. ∀sndr:Id. ∀msg:M. ∀x:Id. ∀sig:Atom1. ∀e:E.
     (<x, sig, sndr, n, msg> ∈ V(e)  (↓∃e':E. ((e' < e) ∧ <x, sig, sndr, n, msg> ∈ S(e')))))
  ∧ (∀n:ℕ. ∀sndr:Id. ∀msg:M. ∀x:Id. ∀sig:Atom1. ∀e:E.
       (<x, sig, sndr, n, msg> ∈ S(e)  (correct x)  (loc(e) x ∈ Id)))

Lemma: oar-crypto_wf
[Info:Type]. ∀[es:EO+(Info)]. ∀[M:Type]. ∀[V,S:EClass(Id × Atom1 × Id × ℕ × M)]. ∀[correct:Id ⟶ ℙ].
  (oar-crypto(es;M;V;S;correct) ∈ ℙ)

Definition: oar-failure-model
There are 3f+1 orderers and no more than of them are incorrect.
(Without assuming that "correct" is decidable, we can't say that 2f+1 of 
them are correct because we don't know which of them is correct.)⋅

oar-failure-model(orderers;f;correct) ==
  (#(orderers) ((3 f) 1) ∈ ℤ)
  ∧ bag-no-repeats(Id;orderers)
  ∧ (∃b:bag(Id). (f < #(b) ∧ bag-no-repeats(Id;b) ∧ (∀x:Id. (x ↓∈  (x ↓∈ orderers ∧ (correct x))))))))

Lemma: oar-failure-model_wf
[orderers:bag(Id)]. ∀[f:ℤ]. ∀[correct:Id ⟶ ℙ].  (oar-failure-model(orderers;f;correct) ∈ ℙ)

Definition: oar-consistency
Correct agents that both deliver the n-th message from sender
agree on that message (even if it was sent by faulty agent).⋅

oar-consistency(es;M;correct;OARDeliver) ==
  ∀e1,e2:E. ∀n:ℕ. ∀sndr:Id. ∀m1,m2:M.
    (<sndr, n, m1> ∈ OARDeliver(e1)
     <sndr, n, m2> ∈ OARDeliver(e2)
     (correct loc(e1))
     (correct loc(e2))
     (m1 m2 ∈ M))

Lemma: oar-consistency_wf
[Info:Type]
  ∀es:EO+(Info). ∀M:Type. ∀correct:Id ⟶ ℙ. ∀OARDeliver:EClass(Id × ℕ × M).
    (oar-consistency(es;M;correct;OARDeliver) ∈ ℙ)

Lemma: OARcast_lemma1
[Info:Type]
  ∀es:EO+(Info). ∀M:Type. ∀V,S:EClass(Id × Atom1 × Id × ℕ × M). ∀correct:Id ⟶ ℙ. ∀f:ℕ. ∀orderers:bag(Id).
  ∀OARDeliver:EClass(Id × ℕ × M).
    (oar-deliver(es;M;V;correct;orderers;f;OARDeliver)
     oar-order(es;M;S;correct;orderers)
     oar-crypto(es;M;V;S;correct)
     oar-failure-model(orderers;f;correct)
     (∀m1,m2:M.  Dec(m1 m2 ∈ M))
     oar-consistency(es;M;correct;OARDeliver))

Definition: test_xxx
test_xxx(x; y.x[y]) ==  x[2]

Definition: test_yyy
test_yyy(x;x) ==  x

Comment: ordered message buffers
In protocols like Paxos and OARcast, we often need state machine
that maintains state of the form
  ⌜n:ℕ × {L:({n 1...} × Msg) List| sorted-by(λx,y. fst(x) < snd(y);L)} ⌝
pair ⌜<k, m> ∈ ℕ × Msg⌝ is message with sequence number.
The state (n,L) represents the fact that all messages with sequence numbers
less than have been processed, and the messages in the sorted list L
(all with sequence numbers > n) are waiting to be processed  
(but can't be processed until messsage with sequence number arrives).

Given new sequenced message ⌜<k, m>⌝ we compute output, state' by
1) if k < then do nothing (i.e. return [],(n,L) )
2) if k > then return [], (n,L') where L' is 
      insert ⌜<k, m>⌝ in 
      (but if there is already some ⌜(<k, m'> ∈ L)⌝ then do nothing)
3) if then 
    a) if [] or ⌜1 < fst(hd(L))⌝ then return (n+1,L)
    b) otherwise (⌜(fst(hd(L))) (n 1) ∈ ℤ⌝so
       split into L1 L2 where L1 is the maximal prefix of for which
       the sequence numbers are consecutive (i.e. with no gaps).
       return L1, (⌜fst(last(L1))⌝L2)

 
 ⋅

Definition: insert-ordered-message
insert-ordered-message(L;x) ==  insert-combine(int-minus-comparison-inc(λp.(fst(p)));λx,y. y;x;L)

Lemma: insert-ordered-message_wf
[M:Type]. ∀[n:ℕ]. ∀[L:{L:({n 1...} × M) List| sorted-by(λx,y. fst(x) < fst(y);L)} ]. ∀[x:{n 1...} × M].
  (insert-ordered-message(L;x) ∈ {L:({n 1...} × M) List| sorted-by(λx,y. fst(x) < fst(y);L)} )

Definition: find-maximal-consecutive
find-maximal-consecutive(g;L) ==
  fst(accumulate (with value ix and list item y):
       eval in
       let i,x ix 
       in if 1=z  then <1, z>  else ix
      over list:
        tl(L)
      with starting value:
       <1, hd(L)>))

Lemma: find-maximal-consecutive_wf
[T:Type]. ∀[g:T ⟶ ℤ]. ∀[L:T List+].  (find-maximal-consecutive(g;L) ∈ {1..||L|| 1-})

Definition: split-maximal-consecutive
split-maximal-consecutive(g;L) ==
  eval find-maximal-consecutive(g;L) in
  eval L1 eval_list(firstn(a;L)) in
  eval L2 eval_list(nth_tl(a;L)) in
    <L1, L2>

Lemma: split-maximal-consecutive_wf
[T:Type]. ∀[g:T ⟶ ℤ]. ∀[L:T List+].
  (split-maximal-consecutive(g;L) ∈ {p:T List+ × (T List)| ((fst(p)) (snd(p))) ∈ (T List)} )

Definition: process-ordered-message
See this comment: ordered message buffers

process-ordered-message(nL;km) ==
  let n,L nL 
  in let k,m km 
     in if (n) < (k)
           then eval L' eval_list(insert-ordered-message(L;km)) in
                <[], n, L'>
           else if (k) < (n)
                   then <[], n, L>
                   else eval sn in
                        if null(L) ∨bsn <fst(hd(L))
                        then <[km], sn, L>
                        else let L1,L2 split-maximal-consecutive(λp.(fst(p));L) 
                             in eval n' fst(last(L1)) in
                                <[km L1], n', L2>
                        fi 

Lemma: process-ordered-message_wf
[M:Type]. ∀[nL:n:ℕ × {L:({n 1...} × M) List| sorted-by(λx,y. fst(x) < fst(y);L)} ]. ∀[km:ℕ × M].
  (process-ordered-message(nL;km) ∈ {tr:({fst(nL)...} × M) List × n:{fst(nL)...} × (({n 1...} × M) List)| 
                                     let out,n',L' tr in 
                                     sorted-by(λx,y. fst(x) < fst(y);L')
                                     ∧ (0 < ||out||  (((fst(km)) (fst(nL)) ∈ ℤ) ∧ (hd(out) km ∈ (ℕ × M))))
                                     ∧ (((fst(nL)) (fst(km)) ∈ ℤ)
                                        ([km (snd(nL))] (out L') ∈ (({fst(nL)...} × M) List)))
                                     ∧ (fst(nL) < fst(km)
                                        (insert-ordered-message(snd(nL);km)
                                          (out L')
                                          ∈ (({(fst(nL)) 1...} × M) List)))
                                     ∧ (fst(km) < fst(nL)
                                        ((↑null(out)) ∧ (L' (snd(nL)) ∈ (({(fst(nL)) 1...} × M) List))))} )

Lemma: process-ordered-message_wf_simple
[M:Type]. ∀[nL:ℤ × ((ℤ × M) List)]. ∀[km:ℤ × M].  (process-ordered-message(nL;km) ∈ (ℤ × M) List × ℤ × ((ℤ × M) List))

Comment: deliver ordered messages
The OARcast protocol delivers the ordered message ⌜<sndr, n, xxx>⌝ once it
has received 2f+1 signed copies of ⌜<sndr, n, xxx>⌝ from distinct orderers
from known set of 3f+1 orderers.

To implement this, we will spawn state machine to handle the
signed messages for given sndr and n. It will have state that it updates
on receipt of new pair ⌜<orderer, xxx>⌝.

The state machine for this should therefore have state
that includes the locations of the orderers not yet heard from
-- initially it will be the list of all 3f+1 orderers.

The state also contains list of pairs ⌜<xxx, count>⌝ which records how many
copies of xxx have been received. (So we need an equality decider on the type
of messages xxx).

Once some pair with 2f+1 is there, the corresponding xxx can be output.



Definition: update-oarcast-deliver
update-oarcast-deliver(eq;s;p) ==
  let remaining,counts 
  in let orderer,xxx 
     in if bag-deq-member(IdDeq;orderer;remaining)
        then eval rem' eval_bag(remaining orderer) in
             eval counts' eval_list(weak-update-alist(eq;counts;xxx;1;n.n 1)) in
               <rem', counts'>
        else s
        fi 

Lemma: update-oarcast-deliver_wf
[M:Type]. ∀[eq:M ⟶ M ⟶ 𝔹]. ∀[s:bag(Id) × ((M × ℕList)]. ∀[p:Id × M].
  (update-oarcast-deliver(eq;s;p) ∈ bag(Id) × ((M × ℕList))

Definition: oarcast-deliver-output
oarcast-deliver-output(num;s) ==
  let rem,counts 
  in eval eval_list(mapfilter(λp.(fst(p));λp.num ≤snd(p);counts)) in
     if null(L) then inr ⋅  else let xx,_ in inl xx fi 

Lemma: oarcast-deliver-output_wf
[M:Type]. ∀[num:ℤ]. ∀[s:bag(Id) × ((M × ℕList)].  (oarcast-deliver-output(num;s) ∈ M?)

Comment: EventML spec properties
exports:                 []
prefix:                  OARcast:s
name:                    OARcast.esh:s

Comment: ------ OARcast - headers ------


Definition: OARcast_headers
OARcast_headers(deliverhdr;oarcasthdr;orderedhdr;orderhdr) ==  [oarcasthdr; orderhdr; orderedhdr; deliverhdr]

Lemma: OARcast_headers_wf
[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
  (OARcast_headers(deliverhdr;oarcasthdr;orderedhdr;orderhdr) ∈ Name List)

Definition: OARcast_headers_no_rep
OARcast_headers_no_rep(deliverhdr;oarcasthdr;orderedhdr;orderhdr) ==
  no_repeats(Name;OARcast_headers(deliverhdr;oarcasthdr;orderedhdr;orderhdr))

Lemma: OARcast_headers_no_rep_wf
[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
  (OARcast_headers_no_rep(deliverhdr;oarcasthdr;orderedhdr;orderhdr) ∈ ℙ)

Definition: OARcast_headers_fun
OARcast_headers_fun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr) ==
  λhdr.if name_eq(hdr;oarcasthdr) then M
       if name_eq(hdr;orderhdr) then Id × ℤ × M
       if name_eq(hdr;orderedhdr) then Id × Id × ℤ × M
       if name_eq(hdr;deliverhdr) then Id × ℤ × M
       else Top
       fi 

Lemma: OARcast_headers_fun_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
  (OARcast_headers_fun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr) ∈ Name ⟶ Type)

Definition: OARcast_headers_type
OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr) ==
  {mf:Name ⟶ ValueAllType| 
   OARcast_headers_no_rep(deliverhdr;oarcasthdr;orderedhdr;orderhdr)
   ∧ (∀hdr∈OARcast_headers(deliverhdr;oarcasthdr;orderedhdr;orderhdr).(mf hdr)
          (OARcast_headers_fun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr) hdr))} 

Lemma: OARcast_headers_type_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
  (OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr) ∈ ℙ')

Comment: ------ OARcast - specification ------


Definition: OARcast_oarcast'base
OARcast_oarcast'base(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  Base(oarcasthdr)

Lemma: OARcast_oarcast'base_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_oarcast'base(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(M))

Definition: OARcast_oarcast'base-program
OARcast_oarcast'base-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  base-class-program(oarcasthdr)

Lemma: OARcast_oarcast'base-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_oarcast'base-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_oarcast'base(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_order'verify
OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  Verify(orderhdr)

Lemma: OARcast_order'verify_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(Id × ℤ × M))

Definition: OARcast_order'verify-program
OARcast_order'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  verify-class-program(orderhdr)

Lemma: OARcast_order'verify-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_order'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_order'bsign
OARcast_order'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λlocs,z. bag-map(λl.mk-msg-interface(l;make-Authentic-Msg(orderhdr;z));locs)

Lemma: OARcast_order'bsign_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_order'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ bag(Id) ⟶ (Id × ℤ × M) ⟶ bag(Interface))

Definition: OARcast_ordered'verify
OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  Verify(orderedhdr)

Lemma: OARcast_ordered'verify_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(Id × Id × ℤ × M))

Definition: OARcast_ordered'verify-program
OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==  verify-class-program(orderedhdr)

Lemma: OARcast_ordered'verify-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_ordered'bsign
OARcast_ordered'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λlocs,z. bag-map(λl.mk-msg-interface(l;make-Authentic-Msg(orderedhdr;z));locs)

Lemma: OARcast_ordered'bsign_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_ordered'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ bag(Id)
   ⟶ (Id × Id × ℤ × M)
   ⟶ bag(Interface))

Definition: OARcast_deliver'send
OARcast_deliver'send(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λl,z. mk-msg-interface(l;make-Msg(deliverhdr;z))

Lemma: OARcast_deliver'send_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_deliver'send(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ Id ⟶ (Id × ℤ × M) ⟶ Interface)

Definition: OARcast_on_oarcast_update
OARcast_on_oarcast_update(M) ==  λslf,msg,seqnum. (seqnum 1)

Lemma: OARcast_on_oarcast_update_wf
[M:ValueAllType]. (OARcast_on_oarcast_update(M) ∈ Id ⟶ M ⟶ ℤ ⟶ ℤ)

Definition: OARcast_on_oarcast_output
OARcast_on_oarcast_output(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λslf,msg,seqnum. (OARcast_order'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) orderers <slf, seqnum, msg>)

Lemma: OARcast_on_oarcast_output_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_on_oarcast_output(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ Id ⟶ M ⟶ ℤ ⟶ bag(Interface))

Definition: OARcast_SenderState
OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1(λslf.0;OARcast_on_oarcast_update(M);OARcast_oarcast'base(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_SenderState_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(ℤ))

Lemma: OARcast_SenderState-functional
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))].
  OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) is functional

Definition: OARcast_SenderStateFun
OARcast_SenderStateFun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf;es;e) ==
  OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e)

Lemma: OARcast_SenderStateFun_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))]. ∀[e:E].
  (OARcast_SenderStateFun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf;es;e) ∈ ℤ)

Lemma: OARcast_SenderState-classrel
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  ∀es:EO+(Message(mf)). ∀e:E. ∀v:ℤ.
    (v ∈ OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e)
    ⇐⇒ OARcast_SenderStateFun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf;es;e))

Definition: OARcast_SenderState-program
OARcast_SenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1-program(λslf.0;OARcast_on_oarcast_update(M);
                       OARcast_oarcast'base-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_SenderState-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_SenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_Sender
OARcast_Sender(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  ((OARcast_on_oarcast_output(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) o
   OARcast_oarcast'base(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)) o
  OARcast_SenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_Sender_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Sender(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ EClass(Interface))

Definition: OARcast_Sender-program
OARcast_Sender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  eclass1-program(OARcast_on_oarcast_output(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf);
                  OARcast_oarcast'base-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
  OARcast_SenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)

Lemma: OARcast_Sender-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Sender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf)
   ∈ LocalClass(OARcast_Sender(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf)))

Definition: OARcast_oastate_init
OARcast_oastate_init(M) ==  <[], 0, []>

Lemma: OARcast_oastate_init_wf
[M:ValueAllType]. (OARcast_oastate_init(M) ∈ (ℤ × M) List × ℤ × ((ℤ × M) List))

Definition: OARcast_orderer_for_sender
OARcast_orderer_for_sender(M) ==
  λsndr,loc,za,z. let l,seqm za 
                  in let z,state 
                     in if sndr then process-ordered-message(state;seqm) else <[], state> fi 

Lemma: OARcast_orderer_for_sender_wf
[M:ValueAllType]
  (OARcast_orderer_for_sender(M) ∈ Id
   ⟶ Id
   ⟶ (Id × ℤ × M)
   ⟶ ((ℤ × M) List × ℤ × ((ℤ × M) List))
   ⟶ ((ℤ × M) List × ℤ × ((ℤ × M) List)))

Definition: OARcast_OrdererForSenderState
OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λsndr.state-class1(λslf.OARcast_oastate_init(M);OARcast_orderer_for_sender(M) 
                                                  sndr;OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;...;mf))

Lemma: OARcast_OrdererForSenderState_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ Id ⟶ EClass((ℤ × M) List
                                                                                         × ℤ
                                                                                         × ((ℤ × M) List)))

Definition: OARcast_OrdererForSenderState-program
OARcast_OrdererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λsndr.state-class1-program(λslf.OARcast_oastate_init(M);OARcast_orderer_for_sender(M) sndr;
                             OARcast_order'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_OrdererForSenderState-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ ∀x:Id. LocalClass(OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) x))

Definition: OARcast_orderer_for_sender_outputs
OARcast_orderer_for_sender_outputs(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ==
  λsndr,slf,z. let seqms,z 
               in bag-union(map(λseqm.(OARcast_ordered'bsign(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) receivers 
                                       <slf, sndr, seqm>);seqms))

Lemma: OARcast_orderer_for_sender_outputs_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List]. ∀[receivers:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_orderer_for_sender_outputs(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ∈ Id
   ⟶ Id
   ⟶ ((ℤ × M) List × ℤ × ((ℤ × M) List))
   ⟶ bag(Interface))

Definition: OARcast_OrdererForSender
OARcast_OrdererForSender(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ==
  λsndr.(OARcast_orderer_for_sender_outputs(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) 
         sndr OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) sndr)

Lemma: OARcast_OrdererForSender_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List]. ∀[receivers:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererForSender(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ∈ Id ⟶ EClass(Interface))

Definition: OARcast_OrdererForSender-program
OARcast_OrdererForSender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ==
  λsndr.eclass0-program(OARcast_orderer_for_sender_outputs(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) 
                        sndr;OARcast_OrdererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) sndr)

Lemma: OARcast_OrdererForSender-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List]. ∀[receivers:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererForSender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf)
   ∈ ∀x:Id. LocalClass(OARcast_OrdererForSender(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) x))

Definition: OARcast_on_order_update
OARcast_on_order_update(M) ==
  λslf,z,state. let sndr,z in let n,msg in if sndr ∈b state then state else [sndr state] fi 

Lemma: OARcast_on_order_update_wf
[M:ValueAllType]. (OARcast_on_order_update(M) ∈ Id ⟶ (Id × ℤ × M) ⟶ (Id List) ⟶ (Id List))

Definition: OARcast_on_order_output
OARcast_on_order_output(M) ==  λslf,z,state. let sndr,z in let n,msg in if sndr ∈b state then {} else {sndr} fi 

Lemma: OARcast_on_order_output_wf
[M:ValueAllType]. (OARcast_on_order_output(M) ∈ Id ⟶ (Id × ℤ × M) ⟶ (Id List) ⟶ bag(Id))

Definition: OARcast_OrdererState
OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1(λslf.[];OARcast_on_order_update(M);OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_OrdererState_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(Id List))

Definition: OARcast_OrdererState-program
OARcast_OrdererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1-program(λslf.[];OARcast_on_order_update(M);
                       OARcast_order'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_OrdererState-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_OrdererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_Orderer
OARcast_Orderer(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ==
  ((OARcast_on_order_output(M) OARcast_order'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)) o
  OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
   >z> OARcast_OrdererForSender(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) z

Lemma: OARcast_Orderer_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List]. ∀[receivers:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Orderer(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ∈ EClass(Interface))

Definition: OARcast_Orderer-program
OARcast_Orderer-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) ==
  eclass1-program(OARcast_on_order_output(M);
                  OARcast_order'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
  OARcast_OrdererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   >>= λz.(OARcast_OrdererForSender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf) z)

Lemma: OARcast_Orderer-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List]. ∀[receivers:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Orderer-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf)
   ∈ LocalClass(OARcast_Orderer(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf)))

Definition: OARcast_init_odstate
OARcast_init_odstate(M;orderers) ==  <orderers, []>

Lemma: OARcast_init_odstate_wf
[M:ValueAllType]. ∀[orderers:bag(Id)].  (OARcast_init_odstate(M;orderers) ∈ bag(Id) × ((M × ℕList))

Definition: OARcast_deliverer_for_sender_seq_update
OARcast_deliverer_for_sender_seq_update(M;deqM) ==
  λsndr,seq,slf,z,state. let orderer,z 
                         in let l,z 
                            in let s,m 
                               in if sndr l ∧b (seq =z s) then update-oarcast-deliver(deqM;state;<orderer, m>else st\000Cate fi 

Lemma: OARcast_deliverer_for_sender_seq_update_wf
[M:ValueAllType]. ∀[deqM:EqDecider(M)].
  (OARcast_deliverer_for_sender_seq_update(M;deqM) ∈ Id
   ⟶ ℤ
   ⟶ Id
   ⟶ (Id × Id × ℤ × M)
   ⟶ (bag(Id) × ((M × ℕList))
   ⟶ (bag(Id) × ((M × ℕList)))

Definition: OARcast_DelivererForSenderSeqState
OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr,seq.
            state-class1(λslf.OARcast_init_odstate(M;orderers);OARcast_deliverer_for_sender_seq_update(M;deqM) sndr 
                                                               seq;OARcast_ordered'verify(M;deliverhdr;...;...;...;mf))

Lemma: OARcast_DelivererForSenderSeqState_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ Id
   ⟶ ℤ
   ⟶ EClass(bag(Id) × ((M × ℕList)))

Definition: OARcast_DelivererForSenderSeqState-program
OARcast_DelivererForSenderSeqState-program(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr,seq. state-class1-program(λslf.OARcast_init_odstate(M;orderers);OARcast_deliverer_for_sender_seq_update(M;deqM) 
                                                                        sndr 
                                                                        seq;
                                  OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_DelivererForSenderSeqState-program_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderSeqState-program(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf)
   ∈ ∀x:Id. ∀zr:ℤ.
       LocalClass(OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
                  zr))

Definition: OARcast_deliverer_for_sender_seq_output
OARcast_deliverer_for_sender_seq_output(M;deliverhdr;flrs;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λsndr,seq,slf,state. let oarcast-deliver-output((2 flrs) 1;state) in
                           if isl(d)
                           then {OARcast_deliver'send(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) slf 
                                 <sndr, seq, outl(d)>}
                           else {}
                           fi 

Lemma: OARcast_deliverer_for_sender_seq_output_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_deliverer_for_sender_seq_output(M;deliverhdr;flrs;oarcasthdr;orderedhdr;orderhdr;mf) ∈ Id
   ⟶ ℤ
   ⟶ Id
   ⟶ (bag(Id) × ((M × ℕList))
   ⟶ bag(Interface))

Definition: OARcast_DelivererForSenderSeq
OARcast_DelivererForSenderSeq(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr,seq. (OARcast_deliverer_for_sender_seq_output(M;deliverhdr;flrs;oarcasthdr;orderedhdr;orderhdr;mf) sndr 
              seq OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
                    sndr 
                    seq)

Lemma: OARcast_DelivererForSenderSeq_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderSeq(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ Id
   ⟶ ℤ
   ⟶ EClass(Interface))

Definition: OARcast_DelivererForSenderSeq-program
OARcast_DelivererForSenderSeq-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr,seq. eclass0-program(OARcast_deliverer_for_sender_seq_output(M;deliverhdr;flrs;oarcasthdr;orderedhdr;...;mf) 
                             sndr 
                             seq;
             OARcast_DelivererForSenderSeqState-program(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
             sndr 
             seq)

Lemma: OARcast_DelivererForSenderSeq-program_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderSeq-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf)
   ∈ ∀x:Id. ∀zv:ℤ.
       LocalClass(OARcast_DelivererForSenderSeq(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
                  zv))

Definition: OARcast_deliverer_for_sender_update
OARcast_deliverer_for_sender_update(M) ==
  λsndr,slf,z,state. let orderer,z 
                     in let sndr,z 
                        in let seq,m 
                           in if seq ∈b state then state else [seq state] fi 

Lemma: OARcast_deliverer_for_sender_update_wf
[M:ValueAllType]. (OARcast_deliverer_for_sender_update(M) ∈ Id ⟶ Id ⟶ (Id × Id × ℤ × M) ⟶ (ℤ List) ⟶ (ℤ List))

Definition: OARcast_deliverer_for_sender_output
OARcast_deliverer_for_sender_output(M) ==
  λsndr,slf,z,state. let orderer,z in let sndr,z in let seq,m in if seq ∈b state then {} else {seq} fi 

Lemma: OARcast_deliverer_for_sender_output_wf
[M:ValueAllType]. (OARcast_deliverer_for_sender_output(M) ∈ Id ⟶ Id ⟶ (Id × Id × ℤ × M) ⟶ (ℤ List) ⟶ bag(ℤ))

Definition: OARcast_DelivererForSenderState
OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λsndr.state-class1(λslf.[];OARcast_deliverer_for_sender_update(M) 
                             sndr;OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_DelivererForSenderState_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ Id ⟶ EClass(ℤ List))

Definition: OARcast_DelivererForSenderState-program
OARcast_DelivererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  λsndr.state-class1-program(λslf.[];OARcast_deliverer_for_sender_update(M) sndr;
                             OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_DelivererForSenderState-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ ∀x:Id. LocalClass(OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) x))

Definition: OARcast_DelivererForSender
OARcast_DelivererForSender(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr.((OARcast_deliverer_for_sender_output(M) sndr o
         OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)) o
        OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) sndr)
         >z> OARcast_DelivererForSenderSeq(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) sndr z

Lemma: OARcast_DelivererForSender_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSender(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ Id
   ⟶ EClass(Interface))

Definition: OARcast_DelivererForSender-program
OARcast_DelivererForSender-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  λsndr.eclass1-program(OARcast_deliverer_for_sender_output(M) sndr;
                        OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
        OARcast_DelivererForSenderState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) sndr
         >>= λz.(OARcast_DelivererForSenderSeq-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;...;mf) 
                 sndr 
                 z)

Lemma: OARcast_DelivererForSender-program_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererForSender-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf)
   ∈ ∀x:Id. LocalClass(OARcast_DelivererForSender(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) x))

Definition: OARcast_on_ordered_update
OARcast_on_ordered_update(M) ==
  λslf,z,state. let orderer,z 
                in let sndr,z 
                   in let n,msg 
                      in if sndr ∈b state then state else [sndr state] fi 

Lemma: OARcast_on_ordered_update_wf
[M:ValueAllType]. (OARcast_on_ordered_update(M) ∈ Id ⟶ (Id × Id × ℤ × M) ⟶ (Id List) ⟶ (Id List))

Definition: OARcast_on_ordered_output
OARcast_on_ordered_output(M) ==
  λslf,z,state. let orderer,z in let sndr,z in let n,msg in if sndr ∈b state then {} else {sndr} fi 

Lemma: OARcast_on_ordered_output_wf
[M:ValueAllType]. (OARcast_on_ordered_output(M) ∈ Id ⟶ (Id × Id × ℤ × M) ⟶ (Id List) ⟶ bag(Id))

Definition: OARcast_DelivererState
OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1(λslf.[];OARcast_on_ordered_update(M);OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;...;mf))

Lemma: OARcast_DelivererState_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ∈ EClass(Id List))

Definition: OARcast_DelivererState-program
OARcast_DelivererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) ==
  state-class1-program(λslf.[];OARcast_on_ordered_update(M);
                       OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))

Lemma: OARcast_DelivererState-program_wf
[M:ValueAllType]. ∀[deliverhdr,oarcasthdr,orderedhdr,orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_DelivererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   ∈ LocalClass(OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)))

Definition: OARcast_Deliverer
OARcast_Deliverer(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  ((OARcast_on_ordered_output(M) OARcast_ordered'verify(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)) o
  OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
   >z> OARcast_DelivererForSender(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) z

Lemma: OARcast_Deliverer_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Deliverer(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ∈ EClass(Interface))

Definition: OARcast_Deliverer-program
OARcast_Deliverer-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) ==
  eclass1-program(OARcast_on_ordered_output(M);
                  OARcast_ordered'verify-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf))
  OARcast_DelivererState-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)
   >>= λz.(OARcast_DelivererForSender-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf) z)

Lemma: OARcast_Deliverer-program_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_Deliverer-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf)
   ∈ LocalClass(OARcast_Deliverer(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf)))

Definition: OARcast_main
OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf) ==
  OARcast_Orderer(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf)@orderers
  || OARcast_Deliverer(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf)@receivers
  || OARcast_Sender(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf)@senders

Lemma: OARcast_main_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)
   ∈ EClass(Interface))

Definition: OARcast_main-program
OARcast_main-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf) ==
  (OARcast_Orderer-program(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;receivers;mf))@orderers
  || (OARcast_Deliverer-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;mf))@receivers
  || (OARcast_Sender-program(M;deliverhdr;oarcasthdr;orderedhdr;orderers;orderhdr;mf))@senders

Lemma: OARcast_main-program_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_main-program(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)
   ∈ LocalClass(OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)))

Comment: ------ OARcast - extra ------


Definition: OARcast_headers_internal
OARcast_headers_internal(orderedhdr;orderhdr) ==  [orderhdr; orderedhdr]

Lemma: OARcast_headers_internal_wf
[orderedhdr,orderhdr:Atom List].  (OARcast_headers_internal(orderedhdr;orderhdr) ∈ Name List)

Definition: OARcast_headers_no_inputs
OARcast_headers_no_inputs(deliverhdr;orderedhdr;orderhdr) ==  [orderhdr; orderedhdr; deliverhdr]

Lemma: OARcast_headers_no_inputs_wf
[deliverhdr,orderedhdr,orderhdr:Atom List].  (OARcast_headers_no_inputs(deliverhdr;orderedhdr;orderhdr) ∈ Name List)

Definition: OARcast_headers_no_inputs_types
OARcast_headers_no_inputs_types(M;deliverhdr;orderedhdr;orderhdr) ==
  [<orderhdr, Id × ℤ × M>; <orderedhdr, Id × Id × ℤ × M>; <deliverhdr, Id × ℤ × M>]

Lemma: OARcast_headers_no_inputs_types_wf
[M:ValueAllType]. ∀[deliverhdr,orderedhdr,orderhdr:Atom List].
  (OARcast_headers_no_inputs_types(M;deliverhdr;orderedhdr;orderhdr) ∈ (Name × Type) List)

Definition: OARcast_message-constraint
OARcast_message-constraint{i:l}(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf) ==
  msg-interface-constraint{i:l}(OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;...;...;...;...;mf);...;mf)

Lemma: OARcast_message-constraint_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_message-constraint{i:l}(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)
   ∈ ℙ')

Definition: OARcast_messages-delivered
OARcast_messages-delivered{i:l}(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf) ==
  msgs-interface-delivered{i:l}(OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;...;...;...;mf);mf)

Lemma: OARcast_messages-delivered_wf
[M:ValueAllType]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ]. ∀[oarcasthdr,orderedhdr:Atom List].
[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)].
  (OARcast_messages-delivered{i:l}(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)
   ∈ ℙ')

Lemma: OARcast-ilf
[M:{T:Type| valueall-type(T)} ]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ].
[oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))]. ∀[e:E]. ∀[d:ℤ].
[i:Id]. ∀[m:Message(mf)].
  {<d, i, m> ∈ OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhdr;receivers;senders;mf)(e)
  ⇐⇒ ↓((∃e':{e':E| e' ≤loc 
          ∃b:Id List
           ∃b2:(ℤ × M) List
            ∃b4:ℤ
             ∃b5:(ℤ × M) List
              ∃v1:ℤ
               ∃v2:M
                (loc(e) ↓∈ orderers
                ∧ (((↑es-info-auth(es;e')) ∧ (header(e') orderhdr) ∧ has-es-info-type(es;e';mf;Id × ℤ × M))
                  ∧ b ∈ OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e')
                  ∧ (fst(msgval(e')) ∈ b)))
                ∧ <b2, b4, b5> ∈ OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) 
                                 (fst(msgval(e')))(e)
                ∧ <v1, v2> ↓∈ b2
                ∧ (d 0)
                ∧ (m make-Authentic-Msg(orderedhdr;<loc(e), fst(msgval(e')), v1, v2>))
                ∧ i ↓∈ receivers))
       ∨ (∃e':{e':E| e' ≤loc 
           ∃b:Id List
            ∃e1:{e1:E| e1 ≤loc 
             ∃b1:ℤ List
              ∃b2:bag(Id)
               ∃b3:(M × ℕList
                (loc(e) ↓∈ receivers
                ∧ (((↑es-info-auth(es;e')) ∧ (header(e') orderedhdr) ∧ has-es-info-type(es;e';mf;Id × Id × ℤ × M))
                  ∧ b ∈ OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e')
                  ∧ (fst(snd(msgval(e'))) ∈ b)))
                ∧ (((↑es-info-auth(es.e';e1)) ∧ (header(e1) orderedhdr) ∧ has-es-info-type(es;e1;mf;Id × Id × ℤ × M))
                  ∧ b1 ∈ OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) 
                         (fst(snd(msgval(e'))))(e1)
                  ∧ (fst(snd(snd(msgval(e1)))) ∈ b1)))
                ∧ <b2, b3> ∈
                   OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
                   (fst(snd(msgval(e')))) 
                   (fst(snd(snd(msgval(e1)))))(e)
                ∧ (↑isl(oarcast-deliver-output((2 flrs) 1;<b2, b3>)))
                ∧ (d 0)
                ∧ (i loc(e))
                ∧ (m
                  make-Msg(deliverhdr;<fst(snd(msgval(e')))
                                        fst(snd(snd(msgval(e1))))
                                        outl(oarcast-deliver-output((2 flrs) 1;<b2, b3>))>)))))
       ∨ (loc(e) ↓∈ senders
         ∧ ((header(e) oarcasthdr) ∧ has-es-info-type(es;e;mf;M))
         ∧ i ↓∈ orderers
         ∧ (d 0)
         ∧ (m
           make-Authentic-Msg(orderhdr;<loc(e)
                                         OARcast_SenderStateFun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf;es;e)
                                         msgval(e)>)))}

Lemma: OARcast-orderhdr
[M:{T:Type| valueall-type(T)} ]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ].
[oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))]. ∀[e:E]. ∀[d:ℤ].
[i:Id]. ∀[auth:𝔹]. ∀[i1:Id]. ∀[k:ℤ]. ∀[v:M].
  (<d, i, mk-msg(auth;orderhdr;<i1, k, v>)> ∈ OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;orderhd\000Cr;receivers;senders;mf)(e)
  ⇐⇒ loc(e) ↓∈ senders
      ∧ ((header(e) oarcasthdr) ∧ has-es-info-type(es;e;mf;M))
      ∧ i ↓∈ orderers
      ∧ (d 0)
      ∧ auth tt
      ∧ (i1 loc(e))
      ∧ (k OARcast_SenderStateFun(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf;es;e))
      ∧ (v msgval(e)))

Lemma: OARcast-orderedhdr
[M:{T:Type| valueall-type(T)} ]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ].
[oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))]. ∀[e:E]. ∀[d:ℤ].
[i:Id]. ∀[auth:𝔹]. ∀[i1,i2:Id]. ∀[k:ℤ]. ∀[v:M].
  (<d, i, mk-msg(auth;orderedhdr;<i1, i2, k, v>)> ∈ OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;o\000Crderhdr;receivers;senders;mf)(
                      e)
  ⇐⇒ ↓∃b2:(ℤ × M) List
        ∃b4:ℤ
         ∃b5:(ℤ × M) List
          ∃e':{e':E| e' ≤loc 
           ∃b:Id List
            (loc(e) ↓∈ orderers
            ∧ (↑es-info-auth(es;e'))
            ∧ (header(e') orderhdr)
            ∧ has-es-info-type(es;e';mf;Id × ℤ × M)
            ∧ b ∈ OARcast_OrdererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e')
            ∧ (fst(msgval(e')) ∈ b))
            ∧ <b2, b4, b5> ∈ OARcast_OrdererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) 
                             (fst(msgval(e')))(e)
            ∧ (d 0)
            ∧ auth tt
            ∧ (i1 loc(e))
            ∧ (i2 (fst(msgval(e'))))
            ∧ i ↓∈ receivers
            ∧ <k, v> ↓∈ b2))

Lemma: OARcast-deliverhdr
[M:{T:Type| valueall-type(T)} ]. ∀[deliverhdr:Atom List]. ∀[deqM:EqDecider(M)]. ∀[flrs:ℤ].
[oarcasthdr,orderedhdr:Atom List]. ∀[orderers:bag(Id)]. ∀[orderhdr:Atom List]. ∀[receivers,senders:bag(Id)].
[mf:OARcast_headers_type{i:l}(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr)]. ∀[es:EO+(Message(mf))]. ∀[e:E]. ∀[d:ℤ].
[i:Id]. ∀[auth:𝔹]. ∀[i1:Id]. ∀[k:ℤ]. ∀[v:M].
  (<d, i, mk-msg(auth;deliverhdr;<i1, k, v>)> ∈ OARcast_main(M;deliverhdr;deqM;flrs;oarcasthdr;orderedhdr;orderers;order\000Chdr;receivers;senders;mf)(e)
  ⇐⇒ ↓∃e':{e':E| e' ≤loc 
        ∃b:Id List
         ∃e1:{e1:E| e1 ≤loc 
          ∃b1:ℤ List
           ∃b2:bag(Id)
            ∃b3:(M × ℕList
             (loc(e) ↓∈ receivers
             ∧ (((↑es-info-auth(es;e')) ∧ (header(e') orderedhdr) ∧ has-es-info-type(es;e';mf;Id × Id × ℤ × M))
               ∧ b ∈ OARcast_DelivererState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf)(e')
               ∧ (fst(snd(msgval(e'))) ∈ b)))
             ∧ (((↑es-info-auth(es.e';e1)) ∧ (header(e1) orderedhdr) ∧ has-es-info-type(es;e1;mf;Id × Id × ℤ × M))
               ∧ b1 ∈ OARcast_DelivererForSenderState(M;deliverhdr;oarcasthdr;orderedhdr;orderhdr;mf) 
                      (fst(snd(msgval(e'))))(e1)
               ∧ (fst(snd(snd(msgval(e1)))) ∈ b1)))
             ∧ <b2, b3> ∈
                OARcast_DelivererForSenderSeqState(M;deliverhdr;deqM;oarcasthdr;orderedhdr;orderers;orderhdr;mf) 
                (fst(snd(msgval(e')))) 
                (fst(snd(snd(msgval(e1)))))(e)
             ∧ (↑isl(oarcast-deliver-output((2 flrs) 1;<b2, b3>)))
             ∧ (d 0)
             ∧ (i loc(e))
             ∧ auth ff
             ∧ (i1 (fst(snd(msgval(e')))))
             ∧ (k (fst(snd(snd(msgval(e1))))))
             ∧ (v outl(oarcast-deliver-output((2 flrs) 1;<b2, b3>)))))

Comment: EventML spec properties
exports:                 []
prefix:                  nysiad:s
name:                    nysiad.esh:s

Comment: ------ nysiad - headers ------


Definition: nysiad_headers
nysiad_headers(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ==
  [readyhdr; add2baghdr; addwaitinghdr; adeliverhdr; inputmsghdr; tooarcasthdr; kdeliverhdr]

Lemma: nysiad_headers_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
  (nysiad_headers(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ∈ Name List)

Definition: nysiad_headers_no_rep
nysiad_headers_no_rep(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ==
  no_repeats(Name;nysiad_headers(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr))

Lemma: nysiad_headers_no_rep_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
  (nysiad_headers_no_rep(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ∈ ℙ)

Definition: nysiad_headers_fun
nysiad_headers_fun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λhdr.if name_eq(hdr;readyhdr) then Id
       if name_eq(hdr;add2baghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;addwaitinghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;adeliverhdr) then Id × Message(mf)
       if name_eq(hdr;inputmsghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;tooarcasthdr) then Message(mf) × Id? × Id
       if name_eq(hdr;kdeliverhdr) then Message(mf) × Id? × Id
       else Top
       fi 

Lemma: nysiad_headers_fun_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[mf:Name ⟶ Type].
  (nysiad_headers_fun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Name
   ⟶ Type)

Definition: nysiad_headers_type
nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ==
  {mf:Name ⟶ ValueAllType| 
   nysiad_headers_no_rep(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)
   ∧ (∀hdr∈nysiad_headers(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr).(mf hdr)
          (nysiad_headers_fun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) 
             hdr))} 

Lemma: nysiad_headers_type_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
  (nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr) ∈ ℙ')

Comment: ------ nysiad - specification ------


Definition: nysiad_ready'base
nysiad_ready'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  Base(readyhdr)

Lemma: nysiad_ready'base_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_ready'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Id))

Definition: nysiad_ready'base-program
nysiad_ready'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  base-class-program(readyhdr)

Lemma: nysiad_ready'base-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_ready'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_ready'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_ready'send
nysiad_ready'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λl,z. mk-msg-interface(l;make-Msg(readyhdr;z))

Lemma: nysiad_ready'send_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_ready'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ Interface)

Definition: nysiad_add2bag'base
nysiad_add2bag'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  Base(add2baghdr)

Lemma: nysiad_add2bag'base_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_add2bag'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id))

Definition: nysiad_add2bag'base-program
nysiad_add2bag'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  base-class-program(add2baghdr)

Lemma: nysiad_add2bag'base-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_add2bag'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_add2bag'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_add2bag'send
nysiad_add2bag'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λl,z. mk-msg-interface(l;make-Msg(add2baghdr;z))

Lemma: nysiad_add2bag'send_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_add2bag'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ Interface)

Definition: nysiad_addwaiting'base
nysiad_addwaiting'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  Base(addwaitinghdr)

Lemma: nysiad_addwaiting'base_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_addwaiting'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id))

Definition: nysiad_addwaiting'base-program
nysiad_addwaiting'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  base-class-program(addwaitinghdr)

Lemma: nysiad_addwaiting'base-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_addwaiting'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_addwaiting'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_addwaiting'send
nysiad_addwaiting'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λl,z. mk-msg-interface(l;make-Msg(addwaitinghdr;z))

Lemma: nysiad_addwaiting'send_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_addwaiting'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ Interface)

Definition: nysiad_adeliver'send
nysiad_adeliver'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λl,z. mk-msg-interface(l;make-Msg(adeliverhdr;z))

Lemma: nysiad_adeliver'send_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_adeliver'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ (Id × Message(mf))
   ⟶ Interface)

Definition: nysiad_inputmsg'base
nysiad_inputmsg'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  Base(inputmsghdr)

Lemma: nysiad_inputmsg'base_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_inputmsg'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id))

Definition: nysiad_inputmsg'base-program
nysiad_inputmsg'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  base-class-program(inputmsghdr)

Lemma: nysiad_inputmsg'base-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_inputmsg'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_inputmsg'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_tooarcast'broadcast
nysiad_tooarcast'broadcast(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λlocs,z. bag-map(λl.mk-msg-interface(l;make-Msg(tooarcasthdr;z));locs)

Lemma: nysiad_tooarcast'broadcast_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_tooarcast'broadcast(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ bag(Id) ⟶ (Message(mf) × Id? × Id) ⟶ bag(Interface))

Definition: nysiad_kdeliver'base
nysiad_kdeliver'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  Base(kdeliverhdr)

Lemma: nysiad_kdeliver'base_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_kdeliver'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id))

Definition: nysiad_kdeliver'base-program
nysiad_kdeliver'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  base-class-program(kdeliverhdr)

Lemma: nysiad_kdeliver'base-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_kdeliver'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_kdeliver'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_Amessage-deq
nysiad_Amessage-deq(M;deqM) ==  product-deq(M;Id? × Id;deqM;product-deq(Id?;Id;union-deq(Id;Unit;IdDeq;UnitDeq);IdDeq))

Lemma: nysiad_Amessage-deq_wf
[M:ValueAllType]. ∀[deqM:EqDecider(M)].  (nysiad_Amessage-deq(M;deqM) ∈ EqDecider(M × Id? × Id))

Definition: nysiad_add_waiting
nysiad_add_waiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m) ==
  λslf,za,z. let msg,s,d za in 
            let msgs,w 
            in let msg2m msg in
                   if <m, s, d> ∈b msgs
                   then <remove-first(λx.(product-deq(M;Id?
                                          × Id;deqM;product-deq(Id?;Id;union-deq(Id;Unit;IdDeq;UnitDeq);IdDeq)) 
                                          
                                          <m, s, d>);msgs)
                        w
                        >
                   else <msgs, map-sig-add(waitingmap) w>
                   fi 

Lemma: nysiad_add_waiting_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_add_waiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m)
   ∈ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap))
   ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap)))

Definition: nysiad_add_to_bag
nysiad_add_to_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;waitingmap;mf;...) ==
  λslf,zb,z. let msg,s,d zb in 
            let msgs,w 
            in let msg2m msg in
                let map-sig-find(waitingmap) in
                if isl(f)
                then if deqM outl(f) then <msgs, map-sig-remove(waitingmap) w> else <[<m, s, d> msgs], w> fi 
                else <[<m, s, d> msgs], w>
                fi 

Lemma: nysiad_add_to_bag_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_add_to_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;waitingmap;mf;...)
   ∈ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap))
   ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap)))

Definition: nysiad_MessageBagState
nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m) ==
  memory-class2(λslf.<[], map-sig-empty(waitingmap)>;
                nysiad_add_waiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;...);
                nysiad_addwaiting'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf);
                nysiad_add_to_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;...);
                nysiad_add2bag'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf))

Lemma: nysiad_MessageBagState_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m)
   ∈ EClass((M × Id? × Id) List × map-sig-map(waitingmap)))

Lemma: nysiad_MessageBagState-functional
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[es:EO+(Message(mf))].
  nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;...;...;...;mf;msg2m) is functional

Definition: nysiad_MessageBagStateFun
nysiad_MessageBagStateFun(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;msg2m;es;e) ==
  nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;...)(e)

Lemma: nysiad_MessageBagStateFun_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[es:EO+(Message(mf))]. ∀[e:E].
  (nysiad_MessageBagStateFun(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;msg2m;es;e)
   ∈ (M × Id? × Id) List × map-sig-map(waitingmap))

Lemma: nysiad_MessageBagState-classrel
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  ∀es:EO+(Message(mf)). ∀e:E. ∀v:(M × Id? × Id) List × map-sig-map(waitingmap).
    ...

Definition: nysiad_MessageBagState-program
nysiad_MessageBagState-program(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;msg2m) ==
  memory-class2-program(λslf.<[], map-sig-empty(waitingmap)>;...;...;...;...)

Lemma: nysiad_MessageBagState-program_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_MessageBagState-program(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;msg2m)
   ∈ LocalClass(nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;...;...;...;mf;...)))

Definition: nysiad_on_addwaiting
nysiad_on_addwaiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m) ==
  λslf,zc,z. let m,s,d zc in 
            let msgs,w 
            in let map-sig-find(waitingmap) in
                   if isl(f)
                   then if deqM outl(f) (msg2m m)
                        then {nysiad_ready'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;readyhdr;...;mf) 
                              slf 
                              d}
                        else {}
                        fi 
                   else {}
                   fi 

Lemma: nysiad_on_addwaiting_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[A:𝕌'].
  (nysiad_on_addwaiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m)
   ∈ Id ⟶ (Message(mf) × A × Id) ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap)) ⟶ bag(Interface))

Definition: nysiad_on_add2bag
nysiad_on_add2bag(M;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;waitingmap;mf) ==
  λslf,zd,z. let m,s,d zd in 
            let msgs,w 
            in if slf
               then nysiad_tooarcast'broadcast(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;...;...;mf) 
                    senders 
                    <m, s, d>
               else {}
               fi 

Lemma: nysiad_on_add2bag_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr:Atom List].
[senders:bag(Id)]. ∀[tooarcasthdr:Atom List]. ∀[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_on_add2bag(M;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;waitingmap;mf)
   ∈ Id ⟶ (Message(mf) × Id? × Id) ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap)) ⟶ bag(Interface))

Definition: nysiad_on_input_message_bag
nysiad_on_input_message_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;...;mf;...) ==
  λslf,i,s. if isl(i)
           then nysiad_on_addwaiting(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;...;...;...;mf;msg2m) 
                slf 
                outl(i) 
                s
           else nysiad_on_add2bag(M;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;...;...;...;...;mf) 
                slf 
                outr(i) 
                s
           fi 

Lemma: nysiad_on_input_message_bag_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_on_input_message_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;...;mf;...)
   ∈ Id
   ⟶ (Message(mf) × Id? × Id (Message(mf) × Id? × Id))
   ⟶ ((M × Id? × Id) List × map-sig-map(waitingmap))
   ⟶ bag(Interface))

Definition: nysiad_MessageBagInput
nysiad_MessageBagInput(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  slf,x. {inl x} nysiad_addwaiting'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf\000C))
  || slf,x. {inr nysiad_add2bag'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;m\000Cf))

Lemma: nysiad_MessageBagInput_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_MessageBagInput(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id (Message(mf) × Id? × Id)))

Definition: nysiad_MessageBagInput-program
nysiad_MessageBagInput-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  eclass0-program(λslf,x. {inl x};
  nysiad_addwaiting'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf))
  || eclass0-program(λslf,x. {inr };
     nysiad_add2bag'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf))

Lemma: nysiad_MessageBagInput-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_MessageBagInput-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_MessageBagInput(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_MessageBag
nysiad_MessageBag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;...;mf;msg2m) ==
  ((nysiad_on_input_message_bag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;...;...;...;...;mf;msg2m) o
   nysiad_MessageBagInput(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)) o
  nysiad_MessageBagState(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf;msg2m))

Lemma: nysiad_MessageBag_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  (nysiad_MessageBag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;...;mf;msg2m)
   ∈ EClass(Interface))

Definition: nysiad_MessageBag-program
nysiad_MessageBag-program(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;...;...;...;...;mf;...) ==
  eclass1-program(nysiad_on_input_message_bag(M;add2baghdr;addwaitinghdr;...;deqM;...;...;readyhdr;...;...;...;mf;...);
                  nysiad_MessageBagInput-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;readyhdr;...;mf))
  nysiad_MessageBagState-program(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;...;readyhdr;...;...;mf;...)

Lemma: nysiad_MessageBag-program_wf
[M:ValueAllType]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M].
  ...

Definition: nysiad_on_kdeliver
nysiad_on_kdeliver(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,ze,z. let m,s,d ze in 
                let queue,waiting 
                in if loc d
                   then if isl(s) then <queue [<m, s, d>], tt> else <queue, waiting> fi 
                   else <queue, waiting>
                   fi 

Lemma: nysiad_on_kdeliver_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_on_kdeliver(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹)
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹))

Definition: nysiad_on_ready
nysiad_on_ready(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,d,z. let queue,waiting in if loc then <tl(queue), ff> else <queue, waiting> fi 

Lemma: nysiad_on_ready_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_on_ready(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ Id
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹)
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹))

Definition: nysiad_QueueState
nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc.memory-class2(λslf.<[], ff>;
                     nysiad_on_kdeliver(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
                     loc;
                     nysiad_kdeliver'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf);
                     nysiad_on_ready(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
                     loc;
                     nysiad_ready'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf))

Lemma: nysiad_QueueState_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ EClass((Message(mf) × Id? × Id) List × 𝔹))

Lemma: nysiad_QueueState-functional
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[es:EO+(Message(mf))]. ∀[x:Id].
  nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) is functional

Definition: nysiad_QueueStateFun
nysiad_QueueStateFun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;x;es;e) ==
  nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) x(e)

Lemma: nysiad_QueueStateFun_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[es:EO+(Message(mf))]. ∀[e:E]. ∀[x:Id].
  (nysiad_QueueStateFun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;x;es;e)
   ∈ (Message(mf) × Id? × Id) List × 𝔹)

Lemma: nysiad_QueueState-classrel
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  ∀es:EO+(Message(mf)). ∀e:E. ∀x:Id. ∀v:(Message(mf) × Id? × Id) List × 𝔹.
    (v ∈ nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) x(e)
    ⇐⇒ v
        nysiad_QueueStateFun(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;x;es;e))

Definition: nysiad_QueueState-program
nysiad_QueueState-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc.memory-class2-program(λslf.<[]
                                  ff
                                  >;nysiad_on_kdeliver(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;...;...;mf) 
                                    loc;nysiad_kdeliver'base-program(add2baghdr;...;...;...;...;...;...;mf);... 
                                                                                                            loc;...)

Lemma: nysiad_QueueState-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_QueueState-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ ∀x:Id
       LocalClass(nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
                  x))

Definition: nysiad_InputQueue
nysiad_InputQueue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  slf,x. {inl x} nysiad_kdeliver'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf))
  || slf,x. {inr nysiad_ready'base(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)\000C)

Lemma: nysiad_InputQueue_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_InputQueue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ EClass(Message(mf) × Id? × Id Id))

Definition: nysiad_InputQueue-program
nysiad_InputQueue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  eclass0-program(λslf,x. {inl x};
  nysiad_kdeliver'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf))
  || eclass0-program(λslf,x. {inr };
     nysiad_ready'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf))

Lemma: nysiad_InputQueue-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_InputQueue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ LocalClass(nysiad_InputQueue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf)))

Definition: nysiad_deliver_to_replica
nysiad_deliver_to_replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,d,z.
               let queue,waiting 
               in if loc d
                  then if null(queue)
                       then {}
                       else let m',s',d' hd(queue) in 
                            {nysiad_adeliver'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;readyhdr;...;mf) 
                             slf 
                             <d', m'>}
                       fi 
                  else {}
                  fi 

Lemma: nysiad_deliver_to_replica_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_deliver_to_replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ Id
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹)
   ⟶ bag(Interface))

Definition: nysiad_query
nysiad_query(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,zf,z. let m,s,d zf in 
                let queue,waiting 
                in if loc d
                   then if isl(s)
                        then if ¬bwaiting
                             then {nysiad_addwaiting'send(add2baghdr;addwaitinghdr;adeliverhdr;...;...;readyhdr;...;mf) 
                                   slf 
                                   hd(queue [<m, s, d>])}
                             else {}
                             fi 
                        else {nysiad_adeliver'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;...;...;mf) 
                              slf 
                              <d, m>}
                        fi 
                   else {}
                   fi 

Lemma: nysiad_query_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_query(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ (Message(mf) × Id? × Id)
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹)
   ⟶ bag(Interface))

Definition: nysiad_on_input_queue
nysiad_on_input_queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,i,s. if isl(i)
               then nysiad_query(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) 
                    loc 
                    slf 
                    outl(i) 
                    s
               else nysiad_deliver_to_replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;...;...;mf) 
                    loc 
                    slf 
                    outr(i) 
                    s
               fi 

Lemma: nysiad_on_input_queue_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_on_input_queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ (Message(mf) × Id? × Id Id)
   ⟶ ((Message(mf) × Id? × Id) List × 𝔹)
   ⟶ bag(Interface))

Definition: nysiad_Queue
nysiad_Queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc.((nysiad_on_input_queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) 
         loc o
        nysiad_InputQueue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)) o
       nysiad_QueueState(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) loc)

Lemma: nysiad_Queue_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_Queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ EClass(Interface))

Definition: nysiad_Queue-program
nysiad_Queue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc.eclass1-program(nysiad_on_input_queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;...;...;mf) 
                       loc;
                       nysiad_InputQueue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;...;readyhdr;...;mf))
       nysiad_QueueState-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
         loc

Lemma: nysiad_Queue-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_Queue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ ∀x:Id
       LocalClass(nysiad_Queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) 
                  x))

Definition: nysiad_Out
nysiad_Out(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  λloc,slf,z. let delay,dest,m in 
             {nysiad_add2bag'send(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
              slf 
              <m, inl loc, dest>}

Lemma: nysiad_Out_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_Out(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ∈ Id
   ⟶ Id
   ⟶ (ℤ × Id × Message(mf))
   ⟶ bag(Interface))

Definition: nysiad_Replica
nysiad_Replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X) ==
  λloc.nysiad_Queue(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) loc
       || (nysiad_Out(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) 
           loc base-process-class(X;loc;adeliverhdr))

Lemma: nysiad_Replica_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[X:EClass(ℤ × Id × Message(mf))].
  (nysiad_Replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X) ∈ Id
   ⟶ EClass(Interface))

Definition: nysiad_Replica-program
nysiad_Replica-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;prX) ==
  λloc.nysiad_Queue-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) loc
       || eclass0-program(nysiad_Out(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
                          loc;base-process-class-program(prX;loc;adeliverhdr))

Lemma: nysiad_Replica-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[X:EClass(ℤ × Id × Message(mf))]. ∀[prX:LocalClass(X)].
  (nysiad_Replica-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;prX)
   ∈ ∀x:Id
       LocalClass(nysiad_Replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;X) 
                  x))

Definition: nysiad_Replicas
nysiad_Replicas(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X) ==
  return-loc-bag-class(λslf.actors)
   >z> nysiad_Replica(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X) z

Lemma: nysiad_Replicas_wf
[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[X:EClass(ℤ × Id × Message(mf))].
  (nysiad_Replicas(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X)
   ∈ EClass(Interface))

Definition: nysiad_Replicas-program
nysiad_Replicas-program(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;prX) ==
  return-loc-bag-class-program(λslf.actors)
   >>= λz.(nysiad_Replica-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;prX) 
           z)

Lemma: nysiad_Replicas-program_wf
[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[X:EClass(ℤ × Id × Message(mf))]. ∀[prX:LocalClass(X)].
  (nysiad_Replicas-program(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;prX)
   ∈ LocalClass(nysiad_Replicas(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf;X)))

Definition: nysiad_on_receipt_msg
nysiad_on_receipt_msg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;tooarcasthdr;mf) ==
  λloc,z. let m,s,d in 
         if loc d
         then nysiad_tooarcast'broadcast(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;mf) 
              senders 
              <m, s, d>
         else {}
         fi 

Lemma: nysiad_on_receipt_msg_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)].
[tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_on_receipt_msg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;tooarcasthdr;mf)
   ∈ Id ⟶ (Message(mf) × Id? × Id) ⟶ bag(Interface))

Definition: nysiad_OnReceiptMsg
nysiad_OnReceiptMsg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;tooarcasthdr;mf) ==
  (nysiad_on_receipt_msg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;mf) ...)

Lemma: nysiad_OnReceiptMsg_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)].
[tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_OnReceiptMsg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;tooarcasthdr;mf)
   ∈ EClass(Interface))

Definition: nysiad_OnReceiptMsg-program
nysiad_OnReceiptMsg-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;mf) ==
  eclass0-program(nysiad_on_receipt_msg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;...;...;...;mf);
  nysiad_inputmsg'base-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf))

Lemma: nysiad_OnReceiptMsg-program_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)].
[tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_OnReceiptMsg-program(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;mf)
   ∈ LocalClass(nysiad_OnReceiptMsg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;...;...;mf)))

Definition: nysiad_main
nysiad_main(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;...;mf;...;X) ==
  nysiad_Replicas(actors;add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf;X)
  || nysiad_MessageBag(M;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;...;mf;...)
  || nysiad_OnReceiptMsg(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;senders;...;mf)@actors

Lemma: nysiad_main_wf
[M:ValueAllType]. ∀[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[X:EClass(ℤ × Id × Message(mf))].
  (nysiad_main(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;...;...;...;mf;...;X)
   ∈ EClass(Interface))

Definition: nysiad_main-program
nysiad_main-program(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;inputmsghdr;kdeliverhdr;readyhdr;senders;
                    tooarcasthdr;waitingmap;mf;msg2m;prX) ==
  (...)@actors

Lemma: nysiad_main-program_wf
[M:ValueAllType]. ∀[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[X:EClass(ℤ × Id × Message(mf))]. ∀[prX:LocalClass(X)].
  ...

Comment: ------ nysiad - extra ------


Definition: nysiad_headers_internal
nysiad_headers_internal(add2baghdr;addwaitinghdr;adeliverhdr;readyhdr) ==
  [readyhdr; add2baghdr; addwaitinghdr; adeliverhdr]

Lemma: nysiad_headers_internal_wf
[add2baghdr,addwaitinghdr,adeliverhdr,readyhdr:Atom List].
  (nysiad_headers_internal(add2baghdr;addwaitinghdr;adeliverhdr;readyhdr) ∈ Name List)

Definition: nysiad_headers_no_inputs
nysiad_headers_no_inputs(add2baghdr;addwaitinghdr;adeliverhdr;readyhdr;tooarcasthdr) ==
  [readyhdr; add2baghdr; addwaitinghdr; adeliverhdr; tooarcasthdr]

Lemma: nysiad_headers_no_inputs_wf
[add2baghdr,addwaitinghdr,adeliverhdr,readyhdr,tooarcasthdr:Atom List].
  (nysiad_headers_no_inputs(add2baghdr;addwaitinghdr;adeliverhdr;readyhdr;tooarcasthdr) ∈ Name List)

Definition: nysiad_headers_no_inputs_types
nysiad_headers_no_inputs_types(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf) ==
  [<readyhdr, Id>; <add2baghdr, Message(mf) × Id? × Id>; <addwaitinghdr, Message(mf) × Id? × Id>; <adeliverhdr, Id × Mes\000Csage(mf)>; <tooarcasthdr, Message(mf) × Id? × Id>]

Lemma: nysiad_headers_no_inputs_types_wf
[add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,kdeliverhdr,readyhdr,tooarcasthdr:Atom List].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
  (nysiad_headers_no_inputs_types(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr;mf)
   ∈ (Name × Type) List)

Definition: nysiad_message-constraint
nysiad_message-constraint{i:l}(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;...;...;...;...;...;...;mf;msg2m;X) ==
  msg-interface-constraint{i:l}(nysiad_main(M;actors;add2baghdr;...;...;deqM;...;...;...;...;...;...;mf;msg2m;X);...;mf)

Lemma: nysiad_message-constraint_wf
[M:ValueAllType]. ∀[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[X:EClass(ℤ × Id × Message(mf))].
  (nysiad_message-constraint{i:l}(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;...;...;...;...;...;...;mf;msg2m;X)
   ∈ ℙ')

Definition: nysiad_messages-delivered
nysiad_messages-delivered{i:l}(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;...;...;...;...;...;...;mf;msg2m;X) ==
  msgs-interface-delivered{i:l}(nysiad_main(M;actors;add2baghdr;...;...;deqM;...;...;readyhdr;...;...;...;mf;...;X);mf)

Lemma: nysiad_messages-delivered_wf
[M:ValueAllType]. ∀[actors:bag(Id)]. ∀[add2baghdr,addwaitinghdr,adeliverhdr:Atom List]. ∀[deqM:EqDecider(M)].
[inputmsghdr,kdeliverhdr,readyhdr:Atom List]. ∀[senders:bag(Id)]. ∀[tooarcasthdr:Atom List].
[waitingmap:map-sig{i:l}(Id;M)].
[mf:nysiad_headers_type{i:l}(add2baghdr;addwaitinghdr;adeliverhdr;inputmsghdr;kdeliverhdr;readyhdr;tooarcasthdr)].
[msg2m:Message(mf) ⟶ M]. ∀[X:EClass(ℤ × Id × Message(mf))].
  (nysiad_messages-delivered{i:l}(M;actors;add2baghdr;addwaitinghdr;adeliverhdr;deqM;...;...;...;...;...;...;mf;msg2m;X)
   ∈ ℙ')

Definition: nysiad-inst-msg-fun
nysiad-inst-msg-fun(oarcasthdr;orderhdr;orderedhdr;deliverhdr;readyhdr;add2baghdr;addwaitinghdr;...;...;...;...;M;mf) ==
  λhdr.if name_eq(hdr;oarcasthdr) then M
       if name_eq(hdr;orderhdr) then Id × ℤ × M
       if name_eq(hdr;orderedhdr) then Id × Id × ℤ × M
       if name_eq(hdr;deliverhdr) then Id × ℤ × M
       if name_eq(hdr;readyhdr) then Id
       if name_eq(hdr;add2baghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;addwaitinghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;adeliverhdr) then Id × Message(mf)
       if name_eq(hdr;inputmsghdr) then Message(mf) × Id? × Id
       if name_eq(hdr;tooarcasthdr) then Message(mf) × Id? × Id
       if name_eq(hdr;kdeliverhdr) then Message(mf) × Id? × Id
       else Void
       fi 

Lemma: nysiad-inst-msg-fun_wf
[oarcasthdr,orderhdr,orderedhdr,deliverhdr,readyhdr,add2baghdr,addwaitinghdr,adeliverhdr,inputmsghdr,tooarcasthdr,...



Home Index