Built with Alectryon, running Coq+SerAPI v8.18.0+0.18.1. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl. Style: centered; floating; windowed.
[Loading ML file ring_plugin.cmxs (using legacy method) ... done]
[Loading ML file zify_plugin.cmxs (using legacy method) ... done]
[Loading ML file micromega_plugin.cmxs (using legacy method) ... done]
[Loading ML file btauto_plugin.cmxs (using legacy method) ... done]
From VLSM.Lib Require Import Preamble.
[Loading ML file coq-itauto.plugin ... done]

Core: VLSM Projecting Equivocator Traces

Section sec_equivocator_vlsm_projections.
Given an equivocator_vlsm trace ending in a state s, we can obtain a trace in the original vlsm leading to the si, the ith internal state in s, by extracting a path leading to si.
This section is devoting to formalizing this projects studying its properties. In particular, we show that given a valid_trace for the equivocator_vlsm, we can always extract such a trace for any valid index, and, furthermore, that the trace extracted is valid for the original machine.
Context
  {message : Type}
  (X : VLSM message)
  .
Given a transition_item item for the equivocator_vlsm and a MachineDescriptor referring to a position in the destination of item, it returns a transition item for the original machine (if the descriptor matches the copy affected by this transition) and a new machine descriptor referring to a position in the state prior to the transition.
Definition equivocator_vlsm_transition_item_project
  (item : transition_item (equivocator_vlsm X))
  (descriptor : MachineDescriptor X)
  : option (option (transition_item X) * MachineDescriptor X)
  :=
  match descriptor with
  | NewMachine _ => Some (None, descriptor)
  | Existing j =>
    match item with {| l := el; input := im; output := om; destination := s |} =>
      match equivocator_state_project s j with
      | None => None
      | Some sj =>
        match el with
        | Spawn sn =>
          if (decide (j = equivocator_state_last s)) then (* this is the first state *)
            Some (None, NewMachine sn)
          else Some (None, Existing j)
        | ForkWith i lx =>
            if (decide (j = equivocator_state_last s)) then (* this is the copy *)
              Some (Some {| l := lx; input := im; output := om; destination := sj |}, Existing i)
            else Some (None, Existing j)
        | ContinueWith i lx =>
          if decide (i = j) then
              Some (Some {| l := lx; input := im; output := om; destination := sj |}, Existing i)
            else Some (None, Existing j)
        end
      end
    end
  end.
Since equivocators always have machine 0, We can always project a valid equivocator transition item to component 0.
message: Type
X: VLSM message
item: transition_item
s: state (equivocator_vlsm X)
Ht: transition (l item) (s, input item) = (destination item, output item)
Hv: valid (l item) (s, input item)

oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing 0) = Some (oitem, Existing 0)
message: Type
X: VLSM message
item: transition_item
s: state (equivocator_vlsm X)
Ht: transition (l item) (s, input item) = (destination item, output item)
Hv: valid (l item) (s, input item)

oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing 0) = Some (oitem, Existing 0)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |})
Hv: valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |})

oitem : option transition_item, equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing 0) = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: (equivocator_state_extend s s0, None) = (destination, output)
Hv: initial_state_prop s0 ∧ input = None

oitem : option transition_item, match equivocator_state_project destination 0 with | Some _ => if decide (0 = equivocator_state_last destination) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition l (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output)
Hv: match equivocator_state_project s n with | Some si => valid l (si, input) | None => False end
oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (n = 0) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition l (si, input) in (equivocator_state_extend s si', om') | None => (s, input) end = (destination, output)
Hv: match equivocator_state_project s n with | Some si => valid l (si, input) | None => False end
oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (0 = equivocator_state_last destination) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: (equivocator_state_extend s s0, None) = (destination, output)
Hv: initial_state_prop s0 ∧ input = None

oitem : option transition_item, match equivocator_state_project destination 0 with | Some _ => if decide (0 = equivocator_state_last destination) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hv: initial_state_prop s0 ∧ input = None

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s s0) 0 with | Some _ => if decide (0 = equivocator_state_last (equivocator_state_extend s s0)) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hv: initial_state_prop s0
Hinput: input = None

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s s0) 0 with | Some _ => if decide (0 = equivocator_state_last (equivocator_state_extend s s0)) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hv: initial_state_prop s0

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s s0) 0 with | Some _ => if decide (0 = equivocator_state_last (equivocator_state_extend s s0)) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hv: initial_state_prop s0

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s s0) 0 with | Some _ => if decide (0 = equivocator_state_n s) then Some (None, NewMachine s0) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
s0: state X
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hv: initial_state_prop s0

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s s0) 0 with | Some _ => Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
by eexists.
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition l (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output)
Hv: match equivocator_state_project s n with | Some si => valid l (si, input) | None => False end

oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (n = 0) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
by destruct (decide _); subst; eexists.
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition l (si, input) in (equivocator_state_extend s si', om') | None => (s, input) end = (destination, output)
Hv: match equivocator_state_project s n with | Some si => valid l (si, input) | None => False end

oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (0 = equivocator_state_last destination) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
si: state X
Ht: (let (si', om') := transition l (si, input) in (equivocator_state_extend s si', om')) = (destination, output)
Hv: valid l (si, input)

oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (0 = equivocator_state_last destination) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
si, si': state X
om': option message
Ht: (equivocator_state_extend s si', om') = (destination, output)
Hv: valid l (si, input)

oitem : option transition_item, match equivocator_state_project destination 0 with | Some sj => if decide (0 = equivocator_state_last destination) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
si, si': state X
om': option message
Hv: valid l (si, input)

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s si') 0 with | Some sj => if decide (0 = equivocator_state_last (equivocator_state_extend s si')) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
message: Type
X: VLSM message
n: nat
l: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
si, si': state X
om': option message
Hv: valid l (si, input)

oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s si') 0 with | Some sj => if decide (0 = equivocator_state_n s) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing n) else Some (None, Existing 0) | None => None end = Some (oitem, Existing 0)
by eexists. Qed.
An injectivity result for equivocator_vlsm_transition_item_project.
message: Type
X: VLSM message
item, itemX, itemX': transition_item
i, i': nat
idescriptor:= Existing i: MachineDescriptor X
idescriptor':= Existing i': MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
HitemX: equivocator_vlsm_transition_item_project item idescriptor = Some (Some itemX, odescriptor)
HitemX': equivocator_vlsm_transition_item_project item idescriptor' = Some (Some itemX', odescriptor')

i = i' ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
item, itemX, itemX': transition_item
i, i': nat
idescriptor:= Existing i: MachineDescriptor X
idescriptor':= Existing i': MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
HitemX: equivocator_vlsm_transition_item_project item idescriptor = Some (Some itemX, odescriptor)
HitemX': equivocator_vlsm_transition_item_project item idescriptor' = Some (Some itemX', odescriptor')

i = i' ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i, i': nat
idescriptor:= Existing i: MachineDescriptor X
idescriptor':= Existing i': MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
HitemX: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} idescriptor = Some (Some itemX, odescriptor)
HitemX': equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} idescriptor' = Some (Some itemX', odescriptor')

i = i' ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
ls: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i: nat
idescriptor, idescriptor':= Existing i: MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination i = Some si
HitemX: Some (Some {| l := ls; input := input; destination := si; output := output |}, Existing i) = Some (Some itemX, odescriptor)
si': state X
HitemX': Some (Some {| l := ls; input := input; destination := si'; output := output |}, Existing i) = Some (Some itemX', odescriptor')
Hsi': equivocator_state_project destination i = Some si'

i = i ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination (equivocator_state_last destination) = Some si
HitemX: Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX, odescriptor)
si': state X
Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si'
HitemX': Some (Some {| l := l2; input := input; destination := si'; output := output |}, Existing j) = Some (Some itemX', odescriptor')
equivocator_state_last destination = equivocator_state_last destination ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
ls: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i: nat
idescriptor, idescriptor':= Existing i: MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination i = Some si
HitemX: Some (Some {| l := ls; input := input; destination := si; output := output |}, Existing i) = Some (Some itemX, odescriptor)
si': state X
HitemX': Some (Some {| l := ls; input := input; destination := si'; output := output |}, Existing i) = Some (Some itemX', odescriptor')
Hsi': equivocator_state_project destination i = Some si'

i = i ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
ls: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i: nat
idescriptor, idescriptor':= Existing i: MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination i = Some si
HitemX: Some (Some {| l := ls; input := input; destination := si; output := output |}, Existing i) = Some (Some itemX, odescriptor)
Hsi': equivocator_state_project destination i = Some si
HitemX': Some (Some {| l := ls; input := input; destination := si; output := output |}, Existing i) = Some (Some itemX', odescriptor')

i = i ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
ls: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i: nat
idescriptor, idescriptor':= Existing i: MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi, Hsi': equivocator_state_project destination i = Some si
HitemX': Some (Some {| l := ls; input := input; destination := si; output := output |}, Existing i) = Some (Some itemX', odescriptor')

i = i ∧ {| l := ls; input := input; destination := si; output := output |} = itemX' ∧ Existing i = odescriptor'
message: Type
X: VLSM message
ls: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
i: nat
idescriptor, idescriptor':= Existing i: MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi, Hsi': equivocator_state_project destination i = Some si

i = i ∧ {| l := ls; input := input; destination := si; output := output |} = {| l := ls; input := input; destination := si; output := output |} ∧ Existing i = Existing i
by repeat split.
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination (equivocator_state_last destination) = Some si
HitemX: Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX, odescriptor)
si': state X
Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si'
HitemX': Some (Some {| l := l2; input := input; destination := si'; output := output |}, Existing j) = Some (Some itemX', odescriptor')

equivocator_state_last destination = equivocator_state_last destination ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination (equivocator_state_last destination) = Some si
HitemX: Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX, odescriptor)
HitemX': Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX', odescriptor')
Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si

equivocator_state_last destination = equivocator_state_last destination ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination (equivocator_state_last destination) = Some si
HitemX: Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX, odescriptor)
HitemX': Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX', odescriptor')
Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si

equivocator_state_last destination = equivocator_state_last destination ∧ itemX = itemX' ∧ odescriptor = odescriptor'
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi: equivocator_state_project destination (equivocator_state_last destination) = Some si
HitemX': Some (Some {| l := l2; input := input; destination := si; output := output |}, Existing j) = Some (Some itemX', odescriptor')
Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si

equivocator_state_last destination = equivocator_state_last destination ∧ {| l := l2; input := input; destination := si; output := output |} = itemX' ∧ Existing j = odescriptor'
message: Type
X: VLSM message
j: nat
l2: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX, itemX': transition_item
idescriptor, idescriptor':= Existing (equivocator_state_last destination): MachineDescriptor X
odescriptor, odescriptor': MachineDescriptor X
si: state X
Hsi, Hsi': equivocator_state_project destination (equivocator_state_last destination) = Some si

equivocator_state_last destination = equivocator_state_last destination ∧ {| l := l2; input := input; destination := si; output := output |} = {| l := l2; input := input; destination := si; output := output |} ∧ Existing j = Existing j
by repeat split. Qed.
equivocator_vlsm_transition_item_project only fails for an out-of-range descriptor.
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = None

i : nat, descriptor = Existing i ∧ equivocator_state_project (destination item) i = None
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = None

i : nat, descriptor = Existing i ∧ equivocator_state_project (destination item) i = None
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = None

i : nat, descriptor = Existing i ∧ equivocator_state_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) i = None
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
i: nat
Hitem: match equivocator_state_project destination i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = None

i0 : nat, Existing i = Existing i0 ∧ equivocator_state_project destination i0 = None
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
i: nat
Hitem: match equivocator_state_project destination i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = None

Existing i = Existing i ∧ equivocator_state_project destination i = None
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
i: nat
Hitem: match equivocator_state_project destination i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = None

equivocator_state_project destination i = None
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
i: nat
si: state X
Hitem: match l with | Spawn sn => if decide (i = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := si; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := si; output := output |}, Existing i0) else Some (None, Existing i) end = None
Hi: i < equivocator_state_n destination

Some si = None
by destruct l; case_decide. Qed.
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)

is_Some (equivocator_vlsm_transition_item_project item descriptor)
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)

is_Some (equivocator_vlsm_transition_item_project item descriptor)
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
contra: equivocator_vlsm_transition_item_project item descriptor = None

is_Some None
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
contra: i : nat, descriptor = Existing i ∧ equivocator_state_project (destination item) i = None

is_Some None
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
id: nat
Heqd: descriptor = Existing id
Hd: equivocator_state_project (destination item) id = None

is_Some None
message: Type
X: VLSM message
item: transition_item
id: nat
Hproper: proper_descriptor X (Existing id) (destination item)
Hd: equivocator_state_project (destination item) id = None

is_Some None
message: Type
X: VLSM message
item: transition_item
id: nat
Hproper: is_Some (equivocator_state_project (destination item) id)
Hd: equivocator_state_project (destination item) id = None

is_Some None
message: Type
X: VLSM message
item: transition_item
id: nat
x: state X
Hproper: equivocator_state_project (destination item) id = Some x
Hd: equivocator_state_project (destination item) id = None

is_Some None
by congruence. Qed.
If equivocator_vlsm_transition_item_project produces a transition item, then that item has the same input and output as the argument item.
message: Type
X: VLSM message
item, itemX: transition_item
idescriptor, odescriptor: MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item idescriptor = Some (Some itemX, odescriptor)

i : nat, idescriptor = Existing i ∧ proper_descriptor X idescriptor (destination item) ∧ input item = input itemX ∧ output item = output itemX
message: Type
X: VLSM message
item, itemX: transition_item
idescriptor, odescriptor: MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item idescriptor = Some (Some itemX, odescriptor)

i : nat, idescriptor = Existing i ∧ proper_descriptor X idescriptor (destination item) ∧ input item = input itemX ∧ output item = output itemX
message: Type
X: VLSM message
item, itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
Hitem: match item with | {| l := el; input := im; destination := s; output := om |} => match equivocator_state_project s j with | Some sj => match el with | Spawn sn => if decide (j = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last s) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) end | None => None end end = Some (Some itemX, odescriptor)

i : nat, Existing j = Existing i ∧ proper_descriptor X (Existing j) (destination item) ∧ input item = input itemX ∧ output item = output itemX
message: Type
X: VLSM message
item, itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
Hitem: match item with | {| l := el; input := im; destination := s; output := om |} => match equivocator_state_project s j with | Some sj => match el with | Spawn sn => if decide (j = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last s) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) end | None => None end end = Some (Some itemX, odescriptor)

Existing j = Existing j ∧ proper_descriptor X (Existing j) (destination item) ∧ input item = input itemX ∧ output item = output itemX
message: Type
X: VLSM message
item, itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
Hitem: match item with | {| l := el; input := im; destination := s; output := om |} => match equivocator_state_project s j with | Some sj => match el with | Spawn sn => if decide (j = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last s) then Some (Some {| l := lx; input := im; destination := sj; output := om |}, Existing i) else Some (None, Existing j) end | None => None end end = Some (Some itemX, odescriptor)

proper_descriptor X (Existing j) (destination item) ∧ input item = input itemX ∧ output item = output itemX
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
Hitem: match equivocator_state_project destination j with | Some sj => match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) end | None => None end = Some (Some itemX, odescriptor)

proper_descriptor X (Existing j) (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) ∧ VLSM.input {| l := l; input := input; destination := destination; output := output |} = VLSM.input itemX ∧ VLSM.output {| l := l; input := input; destination := destination; output := output |} = VLSM.output itemX
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
Hitem: match equivocator_state_project destination j with | Some sj => match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) end | None => None end = Some (Some itemX, odescriptor)

is_Some (equivocator_state_project destination j) ∧ input = VLSM.input itemX ∧ output = VLSM.output itemX
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
s: state X
Hitem: match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) end = Some (Some itemX, odescriptor)

is_Some (Some s) ∧ input = VLSM.input itemX ∧ output = VLSM.output itemX
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
itemX: transition_item
j: nat
odescriptor: MachineDescriptor X
s: state X
Hitem: match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) end = Some (Some itemX, odescriptor)

input = VLSM.input itemX ∧ output = VLSM.output itemX
by destruct l; case_decide; inversion Hitem. Qed.
If the destination of a valid equivocator transition_item is singleton, then by projecting the item to component 0 we actually obtain a transition_item for the original machine.
message: Type
X: VLSM message
item: transition_item
Hno_equiv_item: is_singleton_state X (destination item)
s: state (equivocator_vlsm X)
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

Hex : existing_equivocator_label X (l item), equivocator_vlsm_transition_item_project item (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X (l item) Hex; input := input item; destination := equivocator_state_descriptor_project (destination item) (Existing 0); output := output item |}, Existing 0)
message: Type
X: VLSM message
item: transition_item
Hno_equiv_item: is_singleton_state X (destination item)
s: state (equivocator_vlsm X)
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

Hex : existing_equivocator_label X (l item), equivocator_vlsm_transition_item_project item (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X (l item) Hex; input := input item; destination := equivocator_state_descriptor_project (destination item) (Existing 0); output := output item |}, Existing 0)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
Hno_equiv_item: is_singleton_state X (VLSM.destination {| l := l; input := input; destination := destination; output := output |})
s: state (equivocator_vlsm X)
Hv: valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |})

Hex : existing_equivocator_label X (VLSM.l {| l := l; input := input; destination := destination; output := output |}), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) Hex; input := VLSM.input {| l := l; input := input; destination := destination; output := output |}; destination := equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) (Existing 0); output := VLSM.output {| l := l; input := input; destination := destination; output := output |} |}, Existing 0)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
Hno_equiv_item: is_singleton_state X destination
s: state (equivocator_vlsm X)
Hv: valid l (s, input)
Ht: transition l (s, input) = (destination, output)

Hex : existing_equivocator_label X l, equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := equivocator_state_descriptor_project destination (Existing 0); output := output |}, Existing 0)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
Hno_equiv_item: is_singleton_state X destination
s: state (equivocator_vlsm X)
Hv: valid l (s, input)
Ht: transition l (s, input) = (destination, output)
li: label X
Heq_eqvi: l = ContinueWith 0 li

Hex : existing_equivocator_label X l, equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := equivocator_state_descriptor_project destination (Existing 0); output := output |}, Existing 0)
message: Type
X: VLSM message
input: option message
destination: state (equivocator_vlsm X)
output: option message
Hno_equiv_item: is_singleton_state X destination
s: state (equivocator_vlsm X)
li: label X
Ht: transition (ContinueWith 0 li) (s, input) = (destination, output)
Hv: valid (ContinueWith 0 li) (s, input)

Hex : existing_equivocator_label X (ContinueWith 0 li), equivocator_vlsm_transition_item_project {| l := ContinueWith 0 li; input := input; destination := destination; output := output |} (Existing 0) = Some (Some {| l := existing_equivocator_label_extract X (ContinueWith 0 li) Hex; input := input; destination := equivocator_state_descriptor_project destination (Existing 0); output := output |}, Existing 0)
message: Type
X: VLSM message
input: option message
destination: state (equivocator_vlsm X)
output: option message
Hno_equiv_item: is_singleton_state X destination
s: state (equivocator_vlsm X)
li: label X
Ht: transition (ContinueWith 0 li) (s, input) = (destination, output)
Hv: valid (ContinueWith 0 li) (s, input)

_ : True, Some (Some {| l := li; input := input; destination := equivocator_state_s destination 0; output := output |}, Existing 0) = Some (Some {| l := li; input := input; destination := equivocator_state_s destination 0; output := output |}, Existing 0)
repeat split. Qed.
For every valid transition there exists a (non-equivocating) MachineDescriptor for its destination such that by projecting the transition item through that descriptor we obtain the transition item corresponding to the input transition.
message: Type
X: VLSM message
item: transition_item
s: state (equivocator_vlsm X)
Hs: proper_existing_equivocator_label X (l item) s
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

proper_equivocator_label X (l item) s ∧ ( dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (destination item) ∧ equivocator_vlsm_transition_item_project item dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (l item) (existing_equivocator_label_forget_proper X Hs); input := input item; destination := equivocator_state_descriptor_project (destination item) dest_eqv; output := output item |}, equivocator_label_descriptor (l item)))
message: Type
X: VLSM message
item: transition_item
s: state (equivocator_vlsm X)
Hs: proper_existing_equivocator_label X (l item) s
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

proper_equivocator_label X (l item) s ∧ ( dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (destination item) ∧ equivocator_vlsm_transition_item_project item dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (l item) (existing_equivocator_label_forget_proper X Hs); input := input item; destination := equivocator_state_descriptor_project (destination item) dest_eqv; output := output item |}, equivocator_label_descriptor (l item)))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
s: state (equivocator_vlsm X)
Hs: proper_existing_equivocator_label X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) s
Hv: valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |})

proper_equivocator_label X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) s ∧ ( dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) ∧ equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (existing_equivocator_label_forget_proper X Hs); input := VLSM.input {| l := l; input := input; destination := destination; output := output |}; destination := equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) dest_eqv; output := VLSM.output {| l := l; input := input; destination := destination; output := output |} |}, equivocator_label_descriptor (VLSM.l {| l := l; input := input; destination := destination; output := output |})))
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X l s
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)

proper_equivocator_label X l s ∧ ( dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv destination ∧ equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X l (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project destination dest_eqv; output := output |}, equivocator_label_descriptor l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ContinueWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (equivocator_state_update s i si') ∧ equivocator_vlsm_transition_item_project {| l := ContinueWith i l; input := input; destination := equivocator_state_update s i si'; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (ContinueWith i l) (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project (equivocator_state_update s i si') dest_eqv; output := output |}, equivocator_label_descriptor (ContinueWith i l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ForkWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message
dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (equivocator_state_extend s si') ∧ equivocator_vlsm_transition_item_project {| l := ForkWith i l; input := input; destination := equivocator_state_extend s si'; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (ForkWith i l) (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project (equivocator_state_extend s si') dest_eqv; output := output |}, equivocator_label_descriptor (ForkWith i l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ContinueWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (equivocator_state_update s i si') ∧ equivocator_vlsm_transition_item_project {| l := ContinueWith i l; input := input; destination := equivocator_state_update s i si'; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (ContinueWith i l) (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project (equivocator_state_update s i si') dest_eqv; output := output |}, equivocator_label_descriptor (ContinueWith i l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ContinueWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

existing_descriptor X (Existing i) (equivocator_state_update s i si') ∧ equivocator_vlsm_transition_item_project {| l := ContinueWith i l; input := input; destination := equivocator_state_update s i si'; output := output |} (Existing i) = Some (Some {| l := existing_equivocator_label_extract X (ContinueWith i l) (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project (equivocator_state_update s i si') (Existing i); output := output |}, equivocator_label_descriptor (ContinueWith i l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ContinueWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

is_Some (equivocator_state_project (equivocator_state_update s i si') i) ∧ match equivocator_state_project (equivocator_state_update s i si') i with | Some sj => if decide (i = i) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing i) | None => None end = Some (Some {| l := l; input := input; destination := default (if decide (i = 0) then si' else equivocator_state_s s 0) (equivocator_state_project (equivocator_state_update s i si') i); output := output |}, Existing i)
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ContinueWith i l) s
si: state X
Hpr: i < equivocator_state_n s
Hv: valid l (si, input)
si': state X
om': option message

is_Some (equivocator_state_project (equivocator_state_update s i si') i) ∧ match equivocator_state_project (equivocator_state_update s i si') i with | Some sj => if decide (i = i) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing i) | None => None end = Some (Some {| l := l; input := input; destination := default (if decide (i = 0) then si' else equivocator_state_s s 0) (equivocator_state_project (equivocator_state_update s i si') i); output := output |}, Existing i)
by rewrite equivocator_state_update_project_eq, decide_True.
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ForkWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

dest_eqv : MachineDescriptor X, existing_descriptor X dest_eqv (equivocator_state_extend s si') ∧ equivocator_vlsm_transition_item_project {| l := ForkWith i l; input := input; destination := equivocator_state_extend s si'; output := output |} dest_eqv = Some (Some {| l := existing_equivocator_label_extract X (ForkWith i l) (existing_equivocator_label_forget_proper X Hs); input := input; destination := equivocator_state_descriptor_project (equivocator_state_extend s si') dest_eqv; output := output |}, equivocator_label_descriptor (ForkWith i l))
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ForkWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message

is_Some (equivocator_state_project (equivocator_state_extend s si') (equivocator_state_n s)) ∧ match equivocator_state_project (equivocator_state_extend s si') (equivocator_state_n s) with | Some sj => if decide (equivocator_state_n s = S (equivocator_state_last s)) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing (equivocator_state_n s)) | None => None end = Some (Some {| l := l; input := input; destination := default (if decide (S (equivocator_state_last s) = 0) then si' else equivocator_state_s s 0) (equivocator_state_project (equivocator_state_extend s si') (equivocator_state_n s)); output := output |}, Existing i)
message: Type
X: VLSM message
i: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
s: Equivocators.bounded_state_copies X
Hs: proper_existing_equivocator_label X (ForkWith i l) s
si: state X
Hpr: equivocator_state_project s i = Some si
Hv: valid l (si, input)
si': state X
om': option message
Hni: ¬ equivocator_state_n s < equivocator_state_n s
Hn: equivocator_state_n s = equivocator_state_n s

is_Some (Some si') ∧ (if decide (equivocator_state_n s = S (equivocator_state_last s)) then Some (Some {| l := l; input := input; destination := si'; output := output |}, Existing i) else Some (None, Existing (equivocator_state_n s))) = Some (Some {| l := l; input := input; destination := si'; output := output |}, Existing i)
by rewrite (equivocator_state_last_n _ s), decide_True. Qed.
This property attempts to characterize the descriptor obtained after applying an equivocator projection (trace, transition_item) function in terms of the input descriptor and the resulting state.
It is assumed that the original_descriptor is a proper descriptor w.r.t. the final state of the trace/transition on which equivocator_vlsm_transition_item_project or equivocator_vlsm_trace_project was applied. In particular this makes s_descriptor a proper descriptor for the state s (see the lemmas above and below).
What this property adds is the fact that it constrains more the output descriptor of a projection operation in terms of the input descriptor (if the input is Newmachine, the output must be Newmachine, if both are Existing, then the output index must be less than the input), while also guaranteeing that the output state of such a projection has a size less than the index of the input descriptor in case that output descriptor becomes NewMachine (signaling that the projection is complete).
This property is crucial for establishing an invariant on known equivocators (see full_node_limited_equivocation_constraint_known_equivocators).
Definition previous_state_descriptor_prop
  (original_descriptor : MachineDescriptor X)
  (s : state (equivocator_vlsm X))
  (s_descriptor : MachineDescriptor X)
  : Prop :=
    match original_descriptor with
    | NewMachine sd => s_descriptor = original_descriptor
    | Existing id =>
      match s_descriptor with
      | NewMachine _ => equivocator_state_n s <= id
      | Existing id' => id' <= id
      end
    end.

message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)

(oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (l itemx) (sx, input itemx) ∧ transition (l itemx) (sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)

(oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (l itemx) (sx, input itemx) ∧ transition (l itemx) (sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (VLSM.destination {| l := l; input := input; destination := destination; output := output |})

(oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := input; destination := destination; output := output |}), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) Hex) ∧ VLSM.input {| l := l; input := input; destination := destination; output := output |} = VLSM.input itemx ∧ VLSM.output {| l := l; input := input; destination := destination; output := output |} = VLSM.output itemx ∧ equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) descriptor = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := input; destination := destination; output := output |}) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) → transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |}) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) descriptor = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (VLSM.destination {| l := l; input := input; destination := destination; output := output |})

(oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination

(oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
s: state X
Heqvi: descriptor = NewMachine s
Hproper: proper_descriptor X (NewMachine s) destination

(oitem : option transition_item) (descriptor' : MachineDescriptor X), Some (None, NewMachine s) = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ s = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s0 : Equivocators.bounded_state_copies X, equivocator_valid X l (s0, input) → equivocator_transition X l (s0, input) = (destination, output) → proper_descriptor X descriptor' s0 ∧ descriptor' = NewMachine s ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s0 descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => s = equivocator_state_descriptor_project s0 descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
n: nat
Heqvi: descriptor = Existing n
Hproper: proper_descriptor X (Existing n) destination
(oitem : option transition_item) (descriptor' : MachineDescriptor X), match equivocator_state_project destination n with | Some sj => match l with | Spawn sn => if decide (n = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing n) | ContinueWith i lx => if decide (i = n) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing n) | ForkWith i lx => if decide (n = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing n) end | None => None end = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (equivocator_state_project destination n) = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X descriptor' s ∧ match descriptor' with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = ( VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (equivocator_state_project destination n) = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
s: state X
Heqvi: descriptor = NewMachine s
Hproper: proper_descriptor X (NewMachine s) destination

(oitem : option transition_item) (descriptor' : MachineDescriptor X), Some (None, NewMachine s) = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ s = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s0 : Equivocators.bounded_state_copies X, equivocator_valid X l (s0, input) → equivocator_transition X l (s0, input) = (destination, output) → proper_descriptor X descriptor' s0 ∧ descriptor' = NewMachine s ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s0 descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => s = equivocator_state_descriptor_project s0 descriptor' end)
by eexists None, _.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
n: nat
Heqvi: descriptor = Existing n
Hproper: proper_descriptor X (Existing n) destination

(oitem : option transition_item) (descriptor' : MachineDescriptor X), match equivocator_state_project destination n with | Some sj => match l with | Spawn sn => if decide (n = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing n) | ContinueWith i lx => if decide (i = n) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing n) | ForkWith i lx => if decide (n = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing n) end | None => None end = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (equivocator_state_project destination n) = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X descriptor' s ∧ match descriptor' with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (equivocator_state_project destination n) = equivocator_state_descriptor_project s descriptor' end)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn

match None with | Some itemx => ( Hex : existing_equivocator_label X (Spawn nsi), l itemx = existing_equivocator_label_extract X (Spawn nsi) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ NewMachine nsi = equivocator_label_descriptor (Spawn nsi) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (NewMachine nsi) s ∧ match NewMachine nsi with | NewMachine _ => equivocator_state_n s ≤ equivocator_state_last destination | Existing id' => id' ≤ equivocator_state_last destination endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (NewMachine nsi) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi) end)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
match None with | Some itemx => ( Hex : existing_equivocator_label X (Spawn nsi), l itemx = existing_equivocator_label_extract X (Spawn nsi) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (Spawn nsi) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) ( sx, VLSM.input itemx) ∧ transition (l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
match Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => ( Hex : existing_equivocator_label X (ContinueWith n li), l itemx = existing_equivocator_label_extract X (ContinueWith n li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ContinueWith n li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ContinueWith n li) ( s, input) → equivocator_transition X (ContinueWith n li) ( s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) ( sx, VLSM.input itemx) ∧ transition (l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
match None with | Some itemx => ( Hex : existing_equivocator_label X (ContinueWith ieqvi li), l itemx = existing_equivocator_label_extract X (ContinueWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ContinueWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ContinueWith ieqvi li) ( s, input) → equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) ( sx, VLSM.input itemx) ∧ transition (l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
match Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => ( Hex : existing_equivocator_label X (ForkWith ieqvi li), l itemx = existing_equivocator_label_extract X (ForkWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing ieqvi = equivocator_label_descriptor (ForkWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) ( s, input) → equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output) → proper_descriptor X (Existing ieqvi) s ∧ match Existing ieqvi with | NewMachine _ => equivocator_state_n s ≤ equivocator_state_last destination | Existing id' => id' ≤ equivocator_state_last destination endmatch Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing ieqvi) → valid (l itemx) ( sx, VLSM.input itemx) ∧ transition (l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing ieqvi) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
match None with | Some itemx => ( Hex : existing_equivocator_label X (ForkWith ieqvi li), l itemx = existing_equivocator_label_extract X (ForkWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ForkWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) ( s, input) → equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) ( sx, VLSM.input itemx) ∧ transition (l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn

match None with | Some itemx => ( Hex : existing_equivocator_label X (Spawn nsi), l itemx = existing_equivocator_label_extract X (Spawn nsi) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ NewMachine nsi = equivocator_label_descriptor (Spawn nsi) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (NewMachine nsi) s ∧ match NewMachine nsi with | NewMachine _ => equivocator_state_n s ≤ equivocator_state_last destination | Existing id' => id' ≤ equivocator_state_last destination endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (NewMachine nsi) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi) end)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn

s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (NewMachine nsi) s ∧ equivocator_state_n s ≤ equivocator_state_last destination ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)

proper_descriptor X (NewMachine nsi) s ∧ equivocator_state_n s ≤ equivocator_state_last destination ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)

equivocator_state_n s ≤ equivocator_state_last destination ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

equivocator_state_n s ≤ equivocator_state_last destination ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

equivocator_state_n s ≤ equivocator_state_last destination ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (NewMachine nsi)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project destination (Existing (equivocator_state_n s))
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

destn = default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s))
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

destn = default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_last destination))
by rewrite Hpr.
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination

match None with | Some itemx => ( Hex : existing_equivocator_label X (Spawn nsi), l itemx = existing_equivocator_label_extract X (Spawn nsi) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (Spawn nsi) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination

s : Equivocators.bounded_state_copies X, equivocator_valid X (Spawn nsi) (s, input) → equivocator_transition X (Spawn nsi) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
proper_descriptor X (Existing n) s
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
_sn: state X
Hpr': equivocator_state_project s n = Some _sn

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
_sn: state X
Hpr': equivocator_state_project s n = Some _sn

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
_sn: state X
Hpr': n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
_sn: state X
Hpr': n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project destination (Existing n)
by simpl; rewrite Hpr.
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

proper_descriptor X (Existing n) s
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hn: n < equivocator_state_n destination

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
nsi: state X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (Spawn nsi) (s, input)
Ht: equivocator_transition X (Spawn nsi) (s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hn: n < equivocator_state_n destination
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

is_Some (equivocator_state_project s n)
by destruct_equivocator_state_project s n _sn Hn'; [eexists | lia].
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn

match Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => ( Hex : existing_equivocator_label X (ContinueWith n li), l itemx = existing_equivocator_label_extract X (ContinueWith n li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ContinueWith n li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ContinueWith n li) (s, input) → equivocator_transition X (ContinueWith n li) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn

(( _ : True, li = li) ∧ input = input ∧ output = output ∧ destn = destn ∧ Existing n = Existing n) ∧ ( s : Equivocators.bounded_state_copies X, match equivocator_state_project s n with | Some si => valid li (si, input) | None => False endmatch equivocator_state_project s n with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output) → is_Some (equivocator_state_project s n) ∧ n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (equivocator_state_project s n) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output)))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn

s : Equivocators.bounded_state_copies X, match equivocator_state_project s n with | Some si => valid li (si, input) | None => False endmatch equivocator_state_project s n with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output) → is_Some (equivocator_state_project s n) ∧ n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (equivocator_state_project s n) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s n with | Some si => valid li (si, input) | None => False end
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output)

is_Some (equivocator_state_project s n) ∧ n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (equivocator_state_project s n) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s n with | Some si => valid li (si, input) | None => False end
Ht: match equivocator_state_project s n with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_update s n si', om') | None => (s, input) end = (destination, output)

is_Some (equivocator_state_project s n) ∧ n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (equivocator_state_project s n) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

is_Some (Some sn) ∧ n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sn) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

n ≤ n ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sn) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

sx : state X, sx = default (equivocator_state_zero s) (Some sn) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)
sx: state X
Hsx: sx = default (equivocator_state_zero s) (Some sn)

valid li (sx, input) ∧ transition li (sx, input) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

valid li (default (equivocator_state_zero s) (Some sn), input) ∧ transition li (default (equivocator_state_zero s) (Some sn), input) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

valid li (sn, input) ∧ transition li (sn, input) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (let (si', om') := transition li (sn, input) in (equivocator_state_update s n si', om')) = (destination, output)

transition li (sn, input) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
si': state X
_output: option message
Ht: (equivocator_state_update s n si', _output) = (destination, output)

(si', _output) = (destn, output)
message: Type
X: VLSM message
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
s: Equivocators.bounded_state_copies X
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
si': state X
_output: option message
Ht: (equivocator_state_update s n si', _output) = (destination, output)
H0: equivocator_state_update s n si' = destination
H1: _output = output

(si', output) = (destn, output)
message: Type
X: VLSM message
li: label X
input, output: option message
n: nat
destn: state X
s: Equivocators.bounded_state_copies X
si': state X
Hpr: equivocator_state_project (equivocator_state_update s n si') n = Some destn
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (equivocator_state_update s n si', output) = (equivocator_state_update s n si', output)

(si', output) = (destn, output)
message: Type
X: VLSM message
li: label X
input, output: option message
n: nat
destn: state X
s: Equivocators.bounded_state_copies X
si': state X
Hpr: equivocator_state_project (equivocator_state_update s n si') n = Some destn
sn: state X
Hpri: equivocator_state_project s n = Some sn
Hv: valid li (sn, input)
Ht: (equivocator_state_update s n si', output) = (equivocator_state_update s n si', output)

n < equivocator_state_n s
by apply equivocator_state_project_Some_rev in Hpri.
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n

match None with | Some itemx => ( Hex : existing_equivocator_label X (ContinueWith ieqvi li), l itemx = existing_equivocator_label_extract X (ContinueWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ContinueWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ContinueWith ieqvi li) (s, input) → equivocator_transition X (ContinueWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n

s : Equivocators.bounded_state_copies X, equivocator_valid X (ContinueWith ieqvi li) (s, input) → equivocator_transition X (ContinueWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (ContinueWith ieqvi li) (s, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s ieqvi with | Some si => valid li (si, input) | None => False end
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
proper_descriptor X (Existing n) s
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
Hproper: proper_descriptor X (Existing n) s

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
Hproper: proper_descriptor X (Existing n) s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: equivocator_state_project s n = Some _sn

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project destination (Existing n)
by simpl; rewrite Hpr.
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = equivocator_state_n s

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = equivocator_state_n s
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: ieqvi ≠ n
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ContinueWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = equivocator_state_n s
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Hn: ¬ n < equivocator_state_n s

is_Some None
by apply equivocator_state_project_Some_rev in Hpr; lia.
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn

match Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => ( Hex : existing_equivocator_label X (ForkWith ieqvi li), l itemx = existing_equivocator_label_extract X (ForkWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing ieqvi = equivocator_label_descriptor (ForkWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) (s, input) → equivocator_transition X (ForkWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing ieqvi) s ∧ match Existing ieqvi with | NewMachine _ => equivocator_state_n s ≤ equivocator_state_last destination | Existing id' => id' ≤ equivocator_state_last destination endmatch Some {| l := li; input := input; destination := destn; output := output |} with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing ieqvi) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing ieqvi) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn

s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) (s, input) → equivocator_transition X (ForkWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing ieqvi) s ∧ ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = equivocator_state_descriptor_project s (Existing ieqvi) → valid (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) ∧ transition (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) = (VLSM.destination {| l := li; input := input; destination := destn; output := output |}, VLSM.output {| l := li; input := input; destination := destn; output := output |}))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (ForkWith ieqvi li) (s, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing ieqvi) s ∧ ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = equivocator_state_descriptor_project s (Existing ieqvi) → valid (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) ∧ transition (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) = (VLSM.destination {| l := li; input := input; destination := destn; output := output |}, VLSM.output {| l := li; input := input; destination := destn; output := output |}))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s ieqvi with | Some si => valid li (si, input) | None => False end
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing ieqvi) s ∧ ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = equivocator_state_descriptor_project s (Existing ieqvi) → valid (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) ∧ transition (l {| l := li; input := input; destination := destn; output := output |}) (sx, VLSM.input {| l := li; input := input; destination := destn; output := output |}) = (VLSM.destination {| l := li; input := input; destination := destn; output := output |}, VLSM.output {| l := li; input := input; destination := destn; output := output |}))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s ieqvi with | Some si => valid li (si, input) | None => False end
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

is_Some (equivocator_state_project s ieqvi) ∧ ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (equivocator_state_project s ieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

is_Some (Some sieqvi) ∧ ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ equivocator_state_descriptor_project destination (Existing (equivocator_state_n s)) = si'

ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ equivocator_state_descriptor_project destination (Existing (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

ieqvi ≤ equivocator_state_last destination ∧ ( sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output))
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ equivocator_state_descriptor_project destination (Existing (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

sx : state X, sx = default (equivocator_state_zero s) (Some sieqvi) → valid li (sx, input) ∧ transition li (sx, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: match equivocator_state_project s ieqvi with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_extend s si', om') | None => (s, input) end = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

sx : state X, sx = sieqvi → valid li (sx, input) ∧ transition li (sx, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: match equivocator_state_project s ieqvi with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_extend s si', om') | None => (s, input) end = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s
sx: state X
Hsx: sx = sieqvi

valid li (sx, input) ∧ transition li (sx, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: match equivocator_state_project s ieqvi with | Some si => let (si', om') := transition li (si, input) in (equivocator_state_extend s si', om') | None => (s, input) end = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

valid li (sieqvi, input) ∧ transition li (sieqvi, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: (let (si', om') := transition li (sieqvi, input) in (equivocator_state_extend s si', om')) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

valid li (sieqvi, input) ∧ transition li (sieqvi, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: (let (si', om') := transition li (sieqvi, input) in (equivocator_state_extend s si', om')) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), transition li (sieqvi, input) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

transition li (sieqvi, input) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
s0: state X
o: option message
Ht: (equivocator_state_extend s s0, o) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Ht_pr: (si' : state X) (_oout : option message), (s0, o) = (si', _oout) → _oout = output ∧ default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = si'
Hlt: ieqvi < equivocator_state_n s

(s0, o) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
s0: state X
o: option message
Ht: (equivocator_state_extend s s0, o) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Hlt: ieqvi < equivocator_state_n s
Heqo: o = output
Heqs0: default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)) = s0

(s0, o) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: (equivocator_state_extend s (default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s))), output) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Hlt: ieqvi < equivocator_state_n s

(default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s)), output) = (destn, output)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
destn: state X
Hpr: equivocator_state_project destination (equivocator_state_last destination) = Some destn
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: (equivocator_state_extend s (default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_n s))), output) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Hlt: ieqvi < equivocator_state_n s

(default (equivocator_state_zero destination) (equivocator_state_project destination (equivocator_state_last destination)), output) = (destn, output)
by rewrite Hpr.
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination

match None with | Some itemx => ( Hex : existing_equivocator_label X (ForkWith ieqvi li), l itemx = existing_equivocator_label_extract X (ForkWith ieqvi li) Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ default (equivocator_state_zero destination) (Some destn) = VLSM.destination itemx ∧ Existing n = equivocator_label_descriptor (ForkWith ieqvi li) | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) (s, input) → equivocator_transition X (ForkWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ match Existing n with | NewMachine _ => equivocator_state_n s ≤ n | Existing id' => id' ≤ n endmatch None with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s (Existing n) → valid (l itemx) (sx, VLSM.input itemx) ∧ transition (l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n) end)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination

s : Equivocators.bounded_state_copies X, equivocator_valid X (ForkWith ieqvi li) (s, input) → equivocator_transition X (ForkWith ieqvi li) (s, input) = (destination, output) → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X (ForkWith ieqvi li) (s, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
Hv: match equivocator_state_project s ieqvi with | Some si => valid li (si, input) | None => False end
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
proper_descriptor X (Existing n) s
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s → proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Hproper: proper_descriptor X (Existing n) s

proper_descriptor X (Existing n) s ∧ n ≤ n ∧ default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Hproper: proper_descriptor X (Existing n) s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: equivocator_state_project s n = Some _sn

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project s (Existing n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
_sn: state X
Hprn: n < equivocator_state_n s

default (equivocator_state_zero destination) (Some destn) = equivocator_state_descriptor_project destination (Existing n)
by cbn; rewrite Hpr.
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

proper_descriptor X (Existing n) s
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)

is_Some (equivocator_state_project s n)
message: Type
X: VLSM message
ieqvi: nat
li: label X
input: option message
destination: state (equivocator_vlsm X)
output: option message
n: nat
destn: state X
Hpr: equivocator_state_project destination n = Some destn
H: n ≠ equivocator_state_last destination
s: Equivocators.bounded_state_copies X
sieqvi: state X
Hpri: equivocator_state_project s ieqvi = Some sieqvi
Hv: valid li (sieqvi, input)
Ht: equivocator_transition X (ForkWith ieqvi li) ( s, input) = (destination, output)
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
Hlst_size: equivocator_state_n destination = S (equivocator_state_last destination)
Hn: ¬ n < equivocator_state_n s

is_Some None
by apply equivocator_state_project_Some_rev in Hpr; lia. Qed.
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project item descriptor = Some (oitem, idescriptor)
s: state (equivocator_vlsm X)
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X (destination item) ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project item descriptor = Some (oitem, idescriptor)
s: state (equivocator_vlsm X)
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X (destination item) ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
item: transition_item
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (destination item)
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project item descriptor = Some (oitem, idescriptor)
s: state (equivocator_vlsm X)
Hv: valid (l item) (s, input item)
Ht: transition (l item) (s, input item) = (destination item, output item)
Hchar: (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (l itemx) (sx, input itemx) ∧ transition (l itemx) (sx, input itemx) = ( destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s descriptor' end)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X (destination item) ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
input: option message
destination: state (equivocator_vlsm X)
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor (VLSM.destination {| l := l; input := input; destination := destination; output := output |})
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: state (equivocator_vlsm X)
Hv: valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |})
Hchar: (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := input; destination := destination; output := output |}), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l {| l := l; input := input; destination := destination; output := output |}) Hex) ∧ VLSM.input {| l := l; input := input; destination := destination; output := output |} = VLSM.input itemx ∧ VLSM.output {| l := l; input := input; destination := destination; output := output |} = VLSM.output itemx ∧ equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) descriptor = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := input; destination := destination; output := output |}) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) → transition (VLSM.l {| l := l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l; input := input; destination := destination; output := output |}, VLSM.output {| l := l; input := input; destination := destination; output := output |}) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = ( VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) descriptor = equivocator_state_descriptor_project s descriptor' end)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
Hchar: (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ descriptor' = equivocator_label_descriptor l | None => True end ∧ ( s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = ( VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s descriptor' end)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
_oitemx: option transition_item
_deqv': MachineDescriptor X
_Hpr: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (_oitemx, _deqv')
Hchar1: match _oitemx with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ _deqv' = equivocator_label_descriptor l | None => True end
Hchar2: s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X _deqv' s ∧ previous_state_descriptor_prop descriptor s _deqv' ∧ match _oitemx with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s _deqv' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s _deqv' end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
_oitemx: option transition_item
_deqv': MachineDescriptor X
_Hpr: Some (oitem, idescriptor) = Some (_oitemx, _deqv')
Hchar1: match _oitemx with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ _deqv' = equivocator_label_descriptor l | None => True end
Hchar2: s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X _deqv' s ∧ previous_state_descriptor_prop descriptor s _deqv' ∧ match _oitemx with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s _deqv' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s _deqv' end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
_oitemx: option transition_item
_deqv': MachineDescriptor X
_Hpr: Some (oitem, idescriptor) = Some (_oitemx, _deqv')
Hchar1: match _oitemx with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ _deqv' = equivocator_label_descriptor l | None => True end
Hchar2: s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X _deqv' s ∧ previous_state_descriptor_prop descriptor s _deqv' ∧ match _oitemx with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s _deqv' → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s _deqv' end
H0: oitem = _oitemx
H1: idescriptor = _deqv'

is_equivocating_state X s ∨ is_newmachine_descriptor X _deqv' → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
_Hpr: Some (oitem, idescriptor) = Some (oitem, idescriptor)
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ idescriptor = equivocator_label_descriptor l | None => True end
Hchar2: s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X idescriptor s ∧ previous_state_descriptor_prop descriptor s idescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s idescriptor → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ idescriptor = equivocator_label_descriptor l | None => True end
Hchar2: s : Equivocators.bounded_state_copies X, equivocator_valid X l (s, input) → equivocator_transition X l (s, input) = (destination, output) → proper_descriptor X idescriptor s ∧ previous_state_descriptor_prop descriptor s idescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s idescriptor → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) ( sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ idescriptor = equivocator_label_descriptor l | None => True end
Hchar2: proper_descriptor X idescriptor s ∧ previous_state_descriptor_prop descriptor s idescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s idescriptor → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
oitem: option transition_item
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} descriptor = Some (oitem, idescriptor)
s: Equivocators.bounded_state_copies X
Hv: equivocator_valid X l (s, input)
Ht: equivocator_transition X l (s, input) = (destination, output)
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X l, VLSM.l itemx = existing_equivocator_label_extract X l Hex) ∧ input = VLSM.input itemx ∧ output = VLSM.output itemx ∧ equivocator_state_descriptor_project destination descriptor = VLSM.destination itemx ∧ idescriptor = equivocator_label_descriptor l | None => True end
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s idescriptor → valid (VLSM.l itemx) (sx, VLSM.input itemx) ∧ transition (VLSM.l itemx) (sx, VLSM.input itemx) = (VLSM.destination itemx, VLSM.output itemx) | None => equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor end

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ContinueWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ContinueWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ContinueWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ContinueWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ContinueWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: match descriptor with | NewMachine _ => Some (None, descriptor) | Existing j0 => match equivocator_state_project destination j0 with | Some sj => if decide (j = j0) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing j) else Some (None, Existing j0) | None => None end end = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
deqvi: nat
Hproper: proper_descriptor X (Existing deqvi) destination
idescriptor: MachineDescriptor X
Hproject: match equivocator_state_project destination deqvi with | Some sj => if decide (j = deqvi) then Some (Some {| l := l; input := input; destination := sj; output := output |}, Existing j) else Some (None, Existing deqvi) | None => None end = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop (Existing deqvi) s idescriptor ∧ equivocator_state_descriptor_project destination (Existing deqvi) = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_newmachine_descriptor X (Existing deqvi)
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
deqvi: nat
Hproper: proper_descriptor X (Existing deqvi) destination
idescriptor: MachineDescriptor X
s0: state X
Hproject: (if decide (j = deqvi) then Some (Some {| l := l; input := input; destination := s0; output := output |}, Existing j) else Some (None, Existing deqvi)) = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop (Existing deqvi) s idescriptor ∧ equivocator_state_descriptor_project destination (Existing deqvi) = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_newmachine_descriptor X (Existing deqvi)
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
deqvi: nat
Hproper: proper_descriptor X (Existing deqvi) destination
idescriptor: MachineDescriptor X
s0: state X
H: j ≠ deqvi
Hproject: Some (None, Existing deqvi) = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ContinueWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop (Existing deqvi) s idescriptor ∧ equivocator_state_descriptor_project destination (Existing deqvi) = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = equivocator_state_n s
Heqv: is_newmachine_descriptor X idescriptor

is_newmachine_descriptor X (Existing deqvi)
by inversion Hproject; subst; inversion Heqv.
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)

is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor → is_equivocating_state X destination ∨ is_newmachine_descriptor X descriptor
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
H: is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor

is_equivocating_state X destination
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
H: is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor

equivocator_state_n destination ≠ 1
message: Type
X: VLSM message
j: nat
l: label X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
descriptor: MachineDescriptor X
Hproper: proper_descriptor X descriptor destination
idescriptor: MachineDescriptor X
Hproject: equivocator_vlsm_transition_item_project {| l := ForkWith j l; input := input; destination := destination; output := output |} descriptor = Some (None, idescriptor)
s: Equivocators.bounded_state_copies X
sj: state X
Hsj: equivocator_state_project s j = Some sj
Hv: valid l (sj, input)
Ht: equivocator_transition X (ForkWith j l) (s, input) = (destination, output)
Hchar1: True
Hdeqv': proper_descriptor X idescriptor s
Hchar2: previous_state_descriptor_prop descriptor s idescriptor ∧ equivocator_state_descriptor_project destination descriptor = equivocator_state_descriptor_project s idescriptor
Ht_size: equivocator_state_n destination = S (equivocator_state_n s)
H: is_equivocating_state X s ∨ is_newmachine_descriptor X idescriptor

S (equivocator_state_n s) ≠ 1
by cbv; lia. Qed.
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')
Hitem': i : nat, descriptor = Existing i ∧ proper_descriptor X descriptor (destination item) ∧ input item = input itemx ∧ output item = output itemx

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')
Hproper: proper_descriptor X descriptor (destination item)

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')
Hproper: (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item descriptor = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (l itemx) (sx, input itemx) ∧ transition (l itemx) (sx, input itemx) = ( destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s descriptor' end)

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')
oitem: option transition_item
odescriptor: MachineDescriptor X
Hpr': equivocator_vlsm_transition_item_project item descriptor = Some (oitem, odescriptor)
H: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ odescriptor = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X odescriptor s ∧ previous_state_descriptor_prop descriptor s odescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s odescriptor → valid (l itemx) ( sx, input itemx) ∧ transition (l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s odescriptor end)

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
oitem: option transition_item
odescriptor: MachineDescriptor X
Hitem: Some (oitem, odescriptor) = Some (Some itemx, descriptor')
Hpr': equivocator_vlsm_transition_item_project item descriptor = Some (oitem, odescriptor)
H: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ odescriptor = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X odescriptor s ∧ previous_state_descriptor_prop descriptor s odescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s odescriptor → valid (l itemx) ( sx, input itemx) ∧ transition (l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s odescriptor end)

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
oitem: option transition_item
odescriptor: MachineDescriptor X
Hitem: Some (oitem, odescriptor) = Some (Some itemx, descriptor')
Hpr': equivocator_vlsm_transition_item_project item descriptor = Some (oitem, odescriptor)
H: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ odescriptor = equivocator_label_descriptor (l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X odescriptor s ∧ previous_state_descriptor_prop descriptor s odescriptor ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s odescriptor → valid (l itemx) ( sx, input itemx) ∧ transition (l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) descriptor = equivocator_state_descriptor_project s odescriptor end)
H1: oitem = Some itemx
H2: odescriptor = descriptor'

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
message: Type
X: VLSM message
item, itemx: transition_item
descriptor, descriptor': MachineDescriptor X
Hitem: Some (Some itemx, descriptor') = Some (Some itemx, descriptor')
Hpr': equivocator_vlsm_transition_item_project item descriptor = Some (Some itemx, descriptor')
H: (( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)) ∧ ( s : state (equivocator_vlsm X), valid (l item) (s, input item) → transition (l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop descriptor s descriptor' ∧ ( sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (l itemx) ( sx, input itemx) ∧ transition (l itemx) ( sx, input itemx) = (destination itemx, output itemx)))

( Hex : existing_equivocator_label X (l item), l itemx = existing_equivocator_label_extract X (l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) descriptor = destination itemx ∧ descriptor' = equivocator_label_descriptor (l item)
by apply H. Qed.
The projection of an equivocator_vlsm trace is obtained by traversing the trace from right to left guided by the descriptors produced by equivocator_vlsm_transition_item_project and gathering all non-empty transition_items it produces.
Definition equivocator_vlsm_trace_project
  (tr : list (transition_item (equivocator_vlsm X)))
  (descriptor : MachineDescriptor X)
  : option (list (transition_item X) * MachineDescriptor X)
  :=
  fold_right
    (fun item result =>
      match result with
      | None => None
      | Some (r, idescriptor) =>
        match equivocator_vlsm_transition_item_project item idescriptor with
        | None => None
        | Some (None, odescriptor) => Some (r, odescriptor)
        | Some (Some item', odescriptor) => Some (item' :: r, odescriptor)
        end
      end)
    (Some ([], descriptor))
    tr.
Projecting on a NewMachine descriptor yields an empty trace and the same descriptor.
message: Type
X: VLSM message
tr: list transition_item
s: state X

equivocator_vlsm_trace_project tr (NewMachine s) = Some ([], NewMachine s)
message: Type
X: VLSM message
tr: list transition_item
s: state X

equivocator_vlsm_trace_project tr (NewMachine s) = Some ([], NewMachine s)
by induction tr; simpl; rewrite ?IHtr. Qed.
equivocator_vlsm_trace_project acts like a morphism w.r.t. concatenation (single element in left operand case).
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr: list transition_item
Hproject: equivocator_vlsm_trace_project ([bprefix] ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr: list transition_item
Hproject: equivocator_vlsm_trace_project ([bprefix] ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr: list transition_item
Hproject: match equivocator_vlsm_trace_project bsuffix dlast with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project bprefix idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr, suffix: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Hproject: match equivocator_vlsm_transition_item_project bprefix dmiddle with | Some (Some item', odescriptor) => Some (item' :: suffix, odescriptor) | Some (None, odescriptor) => Some (suffix, odescriptor) | None => None end = Some (tr, dstart)

(dmiddle0 : MachineDescriptor X) (prefix suffix0 : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle0 = Some (prefix, dstart)) (_ : Some (suffix, dmiddle) = Some (suffix0, dmiddle0)), tr = prefix ++ suffix0
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr, suffix: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Hproject: match equivocator_vlsm_transition_item_project bprefix dmiddle with | Some (Some item', odescriptor) => Some (item' :: suffix, odescriptor) | Some (None, odescriptor) => Some (suffix, odescriptor) | None => None end = Some (tr, dstart)

(prefix suffix0 : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some ( prefix, dstart)) (_ : Some (suffix, dmiddle) = Some (suffix0, dmiddle)), tr = prefix ++ suffix0
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
suffix: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix: transition_item
Hprefix: equivocator_vlsm_transition_item_project bprefix dmiddle = Some (Some prefix, dstart)

(prefix0 suffix0 : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix0, dstart)) (_ : Some (suffix, dmiddle) = Some (suffix0, dmiddle)), prefix :: suffix = prefix0 ++ suffix0
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (tr, dmiddle)
Hprefix: equivocator_vlsm_transition_item_project bprefix dmiddle = Some (None, dstart)
(prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix, dstart)) (_ : Some (tr, dmiddle) = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
suffix: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix: transition_item
Hprefix: equivocator_vlsm_transition_item_project bprefix dmiddle = Some (Some prefix, dstart)

(prefix0 suffix0 : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix0, dstart)) (_ : Some (suffix, dmiddle) = Some (suffix0, dmiddle)), prefix :: suffix = prefix0 ++ suffix0
by exists [prefix], suffix; cbn; rewrite Hprefix.
message: Type
X: VLSM message
bprefix: transition_item
bsuffix: list transition_item
dstart, dlast: MachineDescriptor X
tr: list transition_item
dmiddle: MachineDescriptor X
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (tr, dmiddle)
Hprefix: equivocator_vlsm_transition_item_project bprefix dmiddle = Some (None, dstart)

(prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [bprefix] dmiddle = Some (prefix, dstart)) (_ : Some (tr, dmiddle) = Some (suffix, dmiddle)), tr = prefix ++ suffix
by exists []; exists tr; cbn; rewrite Hprefix. Qed.
equivocator_vlsm_trace_project acts like a morphism w.r.t. concatenation.
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dstart: MachineDescriptor X
tr: list transition_item
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dstart: MachineDescriptor X
tr: list transition_item
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
tr: list transition_item

dstart : MachineDescriptor X, equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X

(tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bsuffix: list transition_item
dlast: MachineDescriptor X
tr: list transition_item
dstart: MachineDescriptor X
Hproject: equivocator_vlsm_trace_project ([] ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
IHbprefix: (tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
Hproject: equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (tr, dstart)
(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
bsuffix: list transition_item
dlast: MachineDescriptor X
tr: list transition_item
dstart: MachineDescriptor X
Hproject: equivocator_vlsm_trace_project ([] ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
by exists dstart, [], tr, eq_refl, Hproject.
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
IHbprefix: (tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
Hproject: equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
IHbprefix: (tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
Hproject: equivocator_vlsm_trace_project (a :: bprefix ++ bsuffix) dlast = Some (tr, dstart)

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
IHbprefix: (tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
Hproject: (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project [a] dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast: MachineDescriptor X
IHbprefix: (tr : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr, dstart) → (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
tr: list transition_item
dstart, da: MachineDescriptor X
prefixa, tr': list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
IHbprefix: (dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, da)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr' = prefix ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(dmiddle : MachineDescriptor X) (prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some ( suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(prefix suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = prefix ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(suffix : list transition_item) (_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefixa ++ prefix', dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = (prefixa ++ prefix') ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

(_ : equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefixa ++ prefix', dstart)) (_ : equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)), tr = (prefixa ++ prefix') ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefixa ++ prefix', dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'
tr = (prefixa ++ prefix') ++ suffix
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefixa ++ prefix', dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

match equivocator_vlsm_trace_project bprefix dmiddle with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project a idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (prefixa ++ prefix', dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

match equivocator_vlsm_transition_item_project a da with | Some (Some item', odescriptor) => Some (item' :: prefix', odescriptor) | Some (None, odescriptor) => Some (prefix', odescriptor) | None => None end = Some (prefixa ++ prefix', dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: match equivocator_vlsm_transition_item_project a da with | Some (Some item', odescriptor) => Some ([item'], odescriptor) | Some (None, odescriptor) => Some ([], odescriptor) | None => None end = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

match equivocator_vlsm_transition_item_project a da with | Some (Some item', odescriptor) => Some (item' :: prefix', odescriptor) | Some (None, odescriptor) => Some (prefix', odescriptor) | None => None end = Some (prefixa ++ prefix', dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
oitem': option transition_item
i: MachineDescriptor X
Ha: match oitem' with | Some item' => Some ([item'], i) | None => Some ([], i) end = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

match oitem' with | Some item' => Some (item' :: prefix', i) | None => Some (prefix', i) end = Some (prefixa ++ prefix', dstart)
by destruct oitem' as [item' |]; inversion Ha; subst.
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, da: MachineDescriptor X
tr': list transition_item
dmiddle: MachineDescriptor X
prefix', suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', da)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
Htr': tr' = prefix' ++ suffix
tr: list transition_item
dstart: MachineDescriptor X
prefixa: list transition_item
Ha: equivocator_vlsm_trace_project [a] da = Some (prefixa, dstart)
Hproject: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (tr', da)
Heq: tr = prefixa ++ tr'

tr = (prefixa ++ prefix') ++ suffix
by subst; rewrite app_assoc. Qed.
equivocator_vlsm_trace_project acts like a morphism w.r.t. concatenation (converse).
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dstart, dmiddle: MachineDescriptor X
prefix, suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)

equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dstart, dmiddle: MachineDescriptor X
prefix, suffix: list transition_item
Hprefix: equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart)
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)

equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
prefix, suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)

dstart : MachineDescriptor X, equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)

(prefix : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: equivocator_vlsm_trace_project [] dmiddle = Some (prefix, dstart)

equivocator_vlsm_trace_project ([] ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
IHbprefix: (prefix : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)
equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: equivocator_vlsm_trace_project [] dmiddle = Some (prefix, dstart)

equivocator_vlsm_trace_project ([] ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
by inversion Hprefix; subst.
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
IHbprefix: (prefix : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: equivocator_vlsm_trace_project (a :: bprefix) dmiddle = Some (prefix, dstart)

equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
IHbprefix: (prefix : list transition_item) (dstart : MachineDescriptor X), equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: match equivocator_vlsm_trace_project bprefix dmiddle with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project a idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (prefix, dstart)

equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix': list transition_item
dstart': MachineDescriptor X
Hprefix': equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', dstart')
IHbprefix: (prefix : list transition_item) (dstart : MachineDescriptor X), Some (prefix', dstart') = Some (prefix, dstart) → equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: match equivocator_vlsm_transition_item_project a dstart' with | Some (Some item', odescriptor) => Some (item' :: prefix', odescriptor) | Some (None, odescriptor) => Some (prefix', odescriptor) | None => None end = Some (prefix, dstart)

equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix': list transition_item
dstart': MachineDescriptor X
Hprefix': equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', dstart')
IHbprefix: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix' ++ suffix, dstart')
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: match equivocator_vlsm_transition_item_project a dstart' with | Some (Some item', odescriptor) => Some (item' :: prefix', odescriptor) | Some (None, odescriptor) => Some (prefix', odescriptor) | None => None end = Some (prefix, dstart)

equivocator_vlsm_trace_project ((a :: bprefix) ++ bsuffix) dlast = Some (prefix ++ suffix, dstart)
message: Type
X: VLSM message
a: transition_item
bprefix, bsuffix: list transition_item
dlast, dmiddle: MachineDescriptor X
suffix: list transition_item
Hsuffix: equivocator_vlsm_trace_project bsuffix dlast = Some (suffix, dmiddle)
prefix': list transition_item
dstart': MachineDescriptor X
Hprefix': equivocator_vlsm_trace_project bprefix dmiddle = Some (prefix', dstart')
IHbprefix: equivocator_vlsm_trace_project (bprefix ++ bsuffix) dlast = Some (prefix' ++ suffix, dstart')
prefix: list transition_item
dstart: MachineDescriptor X
Hprefix: match equivocator_vlsm_transition_item_project a dstart' with | Some (Some item', odescriptor) => Some (item' :: prefix', odescriptor) | Some (None, odescriptor) => Some (prefix', odescriptor) | None => None end = Some (prefix, dstart)

match equivocator_vlsm_transition_item_project a dstart' with | Some (Some item', odescriptor) => Some (item' :: prefix' ++ suffix, odescriptor) | Some (None, odescriptor) => Some (prefix' ++ suffix, odescriptor) | None => None end = Some (prefix ++ suffix, dstart)
by destruct (equivocator_vlsm_transition_item_project a dstart') as [[[item' |] i] |] ; inversion Hprefix; subst. Qed.
Next we prove some inversion properties for equivocator_vlsm_transition_item_project.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
di, di': MachineDescriptor X
item': transition_item
Hitem: equivocator_vlsm_transition_item_project item di = Some (Some item', di')

i : nat, di = Existing i ∧ ( sx : state X, equivocator_state_project s i = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom)))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
di, di': MachineDescriptor X
item': transition_item
Hitem: equivocator_vlsm_transition_item_project item di = Some (Some item', di')

i : nat, di = Existing i ∧ ( sx : state X, equivocator_state_project s i = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom)))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
item': transition_item
Hitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (Some item', di')

i0 : nat, Existing i = Existing i0 ∧ ( sx : state X, equivocator_state_project s i0 = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom)))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
item': transition_item
Hitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (Some item', di')

sx : state X, equivocator_state_project s i = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
item': transition_item
Hitem: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some item', di')

sx : state X, equivocator_state_project s i = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
item': transition_item
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) end = Some (Some item', di')

sx : state X, Some si = Some sx ∧ ( i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := sx; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (sx, oom))))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s', s: state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
item': transition_item
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) end = Some (Some item', di')

i' : nat, di' = Existing i' ∧ ( s'x : state X, equivocator_state_project s' i' = Some s'x ∧ ( Hex : existing_equivocator_label X l, let lx := existing_equivocator_label_extract X l Hex in item' = {| l := lx; input := iom; destination := si; output := oom |} ∧ valid lx (s'x, iom) ∧ transition lx (s'x, iom) = (si, oom)))
message: Type
X: VLSM message
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
i: nat
Heqs'j: equivocator_state_project s' i = Some s'j
Hv: valid lx (s'j, iom)
item:= {| l := ContinueWith i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith i lx) (s', iom) = (s, oom)
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)

(s'j', _oom) = (si, oom)
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)
(s'j', _oom) = (si, oom)
message: Type
X: VLSM message
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
i: nat
Heqs'j: equivocator_state_project s' i = Some s'j
Hv: valid lx (s'j, iom)
item:= {| l := ContinueWith i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith i lx) (s', iom) = (s, oom)
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)

(s'j', _oom) = (si, oom)
message: Type
X: VLSM message
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
i: nat
Heqs'j: equivocator_state_project s' i = Some s'j
Hv: valid lx (s'j, iom)
item:= {| l := ContinueWith i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith i lx) (s', iom) = (s, oom)
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)
Heq_oom: _oom = oom
Heqs'j': equivocator_state_descriptor_project s (Existing i) = s'j'

(s'j', _oom) = (si, oom)
by subst; simpl; rewrite Heqsi.
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)

(s'j', _oom) = (si, oom)
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
s'j': state X
_oom: option message
Hti: transition lx (s'j, iom) = (s'j', _oom)
Heq_oom: _oom = oom
Heqs'j': equivocator_state_descriptor_project s (Existing (equivocator_state_n s')) = s'j'

(s'j', _oom) = (si, oom)
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
Hti: transition lx (s'j, iom) = (equivocator_state_descriptor_project s (Existing (equivocator_state_n s')), oom)

(equivocator_state_descriptor_project s (Existing (equivocator_state_n s')), oom) = (si, oom)
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
Hti: transition lx (s'j, iom) = (equivocator_state_descriptor_project s (Existing (equivocator_state_n s')), oom)

(default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')), oom) = (si, oom)
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
Hti: transition lx (s'j, iom) = (equivocator_state_descriptor_project s (Existing (equivocator_state_n s')), oom)

equivocator_state_last s = equivocator_state_n s'
message: Type
X: VLSM message
j: nat
lx: label X
s', s: state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
item:= {| l := ForkWith j lx; input := iom; destination := s; output := oom |}: transition_item
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) = Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j)
Hti: transition lx (s'j, iom) = (equivocator_state_descriptor_project s (Existing (equivocator_state_n s')), oom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')

equivocator_state_last s = equivocator_state_n s'
by specialize (equivocator_state_last_n X s) as Hs_lst; lia. Qed.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
di, di': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item di = Some (None, di')

match di with | NewMachine _ => di' = di | Existing i => match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ equivocator_state_project s i = Some sn' ∧ initial_state_prop sn' | Existing i' => si : state X, equivocator_state_project s i = Some si ∧ equivocator_state_project s' i' = Some si end end
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
di, di': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item di = Some (None, di')

match di with | NewMachine _ => di' = di | Existing i => match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ equivocator_state_project s i = Some sn' ∧ initial_state_prop sn' | Existing i' => si : state X, equivocator_state_project s i = Some si ∧ equivocator_state_project s' i' = Some si end end
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
i: nat
di': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (None, di')

match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ equivocator_state_project s i = Some sn' ∧ initial_state_prop sn' | Existing i' => si : state X, equivocator_state_project s i = Some si ∧ equivocator_state_project s' i' = Some si end
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
di': MachineDescriptor X
Hitem: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (None, di')

match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ equivocator_state_project s i = Some sn' ∧ initial_state_prop sn' | Existing i' => si : state X, equivocator_state_project s i = Some si ∧ equivocator_state_project s' i' = Some si end
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
di': MachineDescriptor X
Hitem: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, di')

match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ equivocator_state_project s i = Some sn' ∧ initial_state_prop sn' | Existing i' => si : state X, equivocator_state_project s i = Some si ∧ equivocator_state_project s' i' = Some si end
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
di': MachineDescriptor X
si: state X
Heqsi: equivocator_state_project s i = Some si
Hitem: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing i0) else Some (None, Existing i) end = Some (None, di')

match di' with | NewMachine sn' => l = Spawn sn' ∧ i = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ Some si = Some sn' ∧ initial_state_prop sn' | Existing i' => si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i' = Some si0 end
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (None, NewMachine sn) = Some (None, NewMachine sn)

Spawn sn = Spawn sn ∧ equivocator_state_last s = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ Some si = Some sn ∧ initial_state_prop sn
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ContinueWith id lx) (s', iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ForkWith id lx) (s', iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
si: state X
Heqsi: equivocator_state_project s (equivocator_state_last s) = Some si
Hitem: Some (None, NewMachine sn) = Some (None, NewMachine sn)

Spawn sn = Spawn sn ∧ equivocator_state_last s = equivocator_state_last s ∧ iom = None ∧ oom = None ∧ Some si = Some sn ∧ initial_state_prop sn
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (equivocator_state_extend s' sn, None)
si: state X
Heqsi: equivocator_state_project (equivocator_state_extend s' sn) (equivocator_state_last (equivocator_state_extend s' sn)) = Some si
Hitem: Some (None, NewMachine sn) = Some (None, NewMachine sn)

Spawn sn = Spawn sn ∧ equivocator_state_last (equivocator_state_extend s' sn) = equivocator_state_last (equivocator_state_extend s' sn) ∧ iom = None ∧ None = None ∧ Some si = Some sn ∧ initial_state_prop sn
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (equivocator_state_extend s' sn, None)
si: state X
Heqsi: equivocator_state_project (equivocator_state_extend s' sn) (equivocator_state_last (equivocator_state_extend s' sn)) = Some si
Hitem: Some (None, NewMachine sn) = Some (None, NewMachine sn)

Some si = Some sn
by rewrite equivocator_state_extend_lst, equivocator_state_extend_project_2 in Heqsi.
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → equivocator_state_descriptor_project s (Existing i) = equivocator_state_descriptor_project s' (Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (equivocator_state_project s i) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
Hn: default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
Hn: si = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
by destruct_equivocator_state_project s' i s'i Hi; [subst | lia].
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ContinueWith id lx) (s', iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)

si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ContinueWith id lx) (s', iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: match equivocator_state_project s' id with | Some si => valid lx (si, iom) | None => False end
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → equivocator_state_descriptor_project s (Existing i) = equivocator_state_descriptor_project s' (Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → default (equivocator_state_zero s) (equivocator_state_project s i) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = equivocator_state_n s'

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → id ≠ i → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
Hn: id ≠ i → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
Hn: default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ContinueWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: id ≠ i
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
s'i: state X
Hn: default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (Some s'i)
Hi: i < equivocator_state_n s'

Some s'i = Some si
by simpl in Hn; subst.
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ForkWith id lx) (s', iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

si0 : state X, Some si = Some si0 ∧ equivocator_state_project s' i = Some si0
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ForkWith id lx) (s', iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: match equivocator_state_project s' id with | Some si => valid lx (si, iom) | None => False end
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → equivocator_state_descriptor_project s (Existing i) = equivocator_state_descriptor_project s' (Existing i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (equivocator_state_project s i) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: equivocator_state_project s i = Some si
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Hn: i < equivocator_state_n s' → default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
Hn: default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (equivocator_state_project s' i)

equivocator_state_project s' i = Some si
message: Type
X: VLSM message
id: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'id: state X
Hpr: equivocator_state_project s' id = Some s'id
Hv: valid lx (s'id, iom)
Ht: transition (ForkWith id lx) (s', iom) = (s, oom)
i: nat
si: state X
Heqsi: i < equivocator_state_n s
n: i ≠ equivocator_state_last s
Hitem: Some (None, Existing i) = Some (None, Existing i)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)
s'i: state X
Hn: default (equivocator_state_zero s) (Some si) = default (equivocator_state_zero s') (Some s'i)
Hi: i < equivocator_state_n s'

Some s'i = Some si
by simpl in Hn; subst. Qed.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i') = Some (oitem, Existing i'))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i') = Some (oitem, Existing i'))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => match l with | Spawn sn => if decide (i' = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i') | ContinueWith i lx => if decide (i = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i) else Some (None, Existing i') | ForkWith i lx => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i) else Some (None, Existing i') end | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some _ => if decide (i' = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ContinueWith j lx) (s', iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ForkWith j lx) (s', iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some _ => if decide (i' = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
H0: equivocator_state_extend s' sn = s
H1: None = oom

si : state X, equivocator_state_project (equivocator_state_extend s' sn) i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s' sn) i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
Ht: transition (Spawn sn) (s', iom) = (equivocator_state_extend s' sn, None)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project (equivocator_state_extend s' sn) i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s' sn) i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project (equivocator_state_extend s' sn) i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s' sn) i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

si : state X, equivocator_state_project (equivocator_state_extend s' sn) i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project (equivocator_state_extend s' sn) i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

si : state X, equivocator_state_project s' i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s' i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

oitem : option transition_item, match equivocator_state_project s' i' with | Some _ => if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i') | None => None end = Some (oitem, Existing i')
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

oitem : option transition_item, (if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i')) = Some (oitem, Existing i')
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

(if decide (i' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i')) = Some (None, Existing i')
message: Type
X: VLSM message
sn: state X
s': state (equivocator_vlsm X)
iom: option message
Hv: valid (Spawn sn) (s', iom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Hlti': i' < equivocator_state_n s'

i' ≠ equivocator_state_last (equivocator_state_extend s' sn)
by rewrite equivocator_state_extend_lst; lia.
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ContinueWith j lx) (s', iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: match equivocator_state_project s' j with | Some si => valid lx (si, iom) | None => False end
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = equivocator_state_n s'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hlti': i' < equivocator_state_n s'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (j = i') then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hlti': i' < equivocator_state_n s'
si: state X
Hlti: i' < equivocator_state_n s

si0 : state X, Some si = Some si0 ∧ ( oitem : option transition_item, (if decide (j = i') then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) else Some (None, Existing i')) = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ContinueWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = equivocator_state_n s'
Hlti': i' < equivocator_state_n s'
si: state X
Hlti: i' < equivocator_state_n s

oitem : option transition_item, (if decide (j = i') then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) else Some (None, Existing i')) = Some (oitem, Existing i')
by destruct (decide _); subst; eexists _.
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid (ForkWith j lx) (s', iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: match equivocator_state_project s' j with | Some si => valid lx (si, iom) | None => False end
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = S (equivocator_state_n s')

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hlti': i' < equivocator_state_n s'

si : state X, equivocator_state_project s i' = Some si ∧ ( oitem : option transition_item, match equivocator_state_project s i' with | Some sj => if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing j) else Some (None, Existing i') | None => None end = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hlti': i' < equivocator_state_n s'
si: state X
Hlti: i' < equivocator_state_n s

si0 : state X, Some si = Some si0 ∧ ( oitem : option transition_item, (if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) else Some (None, Existing i')) = Some (oitem, Existing i'))
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hlti': i' < equivocator_state_n s'
si: state X
Hlti: i' < equivocator_state_n s

oitem : option transition_item, (if decide (i' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing j) else Some (None, Existing i')) = Some (oitem, Existing i')
message: Type
X: VLSM message
j: nat
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
s'j: state X
Heqs'j: equivocator_state_project s' j = Some s'j
Hv: valid lx (s'j, iom)
Ht: transition (ForkWith j lx) (s', iom) = (s, oom)
i': nat
si': state X
Hi': equivocator_state_project s' i' = Some si'
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Hlti': i' < equivocator_state_n s'
si: state X
Hlti: i' < equivocator_state_n s

i' ≠ equivocator_state_last s
by specialize (equivocator_state_last_n X s); lia. Qed.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
sn: state X
Hnew: l = Spawn sn

(i : nat) (si : option (state X)), equivocator_state_project s i = si ∧ equivocator_vlsm_transition_item_project item (Existing i) = Some (None, NewMachine sn)
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
sn: state X
Hnew: l = Spawn sn

(i : nat) (si : option (state X)), equivocator_state_project s i = si ∧ equivocator_vlsm_transition_item_project item (Existing i) = Some (None, NewMachine sn)
message: Type
X: VLSM message
s, s': state (equivocator_vlsm X)
iom, oom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (Spawn sn) (s', iom) = (s, oom)

(i : nat) (si : option (state X)), equivocator_state_project s i = si ∧ equivocator_vlsm_transition_item_project item (Existing i) = Some (None, NewMachine sn)
message: Type
X: VLSM message
s, s': state (equivocator_vlsm X)
iom, oom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (Spawn sn) (s', iom) = (s, oom)

(i : nat) (si : option (state X)), equivocator_state_project s i = si ∧ match equivocator_state_project s i with | Some _ => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | None => None end = Some (None, NewMachine sn)
message: Type
X: VLSM message
s, s': state (equivocator_vlsm X)
iom, oom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (Spawn sn) (s', iom) = (s, oom)
H0: equivocator_state_extend s' sn = s
H1: None = oom

(i : nat) (si : option (state X)), equivocator_state_project (equivocator_state_extend s' sn) i = si ∧ match equivocator_state_project (equivocator_state_extend s' sn) i with | Some _ => if decide (i = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i) | None => None end = Some (None, NewMachine sn)
message: Type
X: VLSM message
s': state (equivocator_vlsm X)
iom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := equivocator_state_extend s' sn; output := None |}: transition_item
Ht: transition (Spawn sn) (s', iom) = (equivocator_state_extend s' sn, None)

(i : nat) (si : option (state X)), equivocator_state_project (equivocator_state_extend s' sn) i = si ∧ match equivocator_state_project (equivocator_state_extend s' sn) i with | Some _ => if decide (i = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i) | None => None end = Some (None, NewMachine sn)
message: Type
X: VLSM message
s': state (equivocator_vlsm X)
iom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := equivocator_state_extend s' sn; output := None |}: transition_item

(i : nat) (si : option (state X)), equivocator_state_project (equivocator_state_extend s' sn) i = si ∧ match equivocator_state_project (equivocator_state_extend s' sn) i with | Some _ => if decide (i = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing i) | None => None end = Some (None, NewMachine sn)
message: Type
X: VLSM message
s': state (equivocator_vlsm X)
iom: option message
sn: state X
item:= {| l := Spawn sn; input := iom; destination := equivocator_state_extend s' sn; output := None |}: transition_item

si : option (state X), equivocator_state_project (equivocator_state_extend s' sn) (equivocator_state_n s') = si ∧ match equivocator_state_project (equivocator_state_extend s' sn) (equivocator_state_n s') with | Some _ => if decide (equivocator_state_n s' = equivocator_state_last (equivocator_state_extend s' sn)) then Some (None, NewMachine sn) else Some (None, Existing (equivocator_state_n s')) | None => None end = Some (None, NewMachine sn)
by rewrite equivocator_state_extend_lst, equivocator_state_extend_project_2, decide_True ; eauto. Qed.
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
_i: nat
Hsndl: equivocator_label_descriptor l = Existing _i

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
l: label (equivocator_vlsm X)
s, s': state (equivocator_vlsm X)
iom, oom: option message
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
item:= {| l := l; input := iom; destination := s; output := oom |}: transition_item
_i: nat
Hsndl: equivocator_label_descriptor l = Existing _i

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ForkWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ equivocator_state_descriptor_project s (Existing _i) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ default (equivocator_state_zero s) (equivocator_state_project s _i) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht_pr: (si'0 : state X) (_oout : option message), (si', _oom) = (si'0, _oout) → _oout = oom ∧ default (equivocator_state_zero s) (equivocator_state_project s _i) = si'0

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s _i) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s _i) = si'

si : state X, equivocator_state_project s _i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing _i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s _i) = si'

si : state X, equivocator_state_project s _i = Some si ∧ ( itemx : transition_item, match equivocator_state_project s _i with | Some sj => if decide (_i = _i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing _i) else Some (None, Existing _i) | None => None end = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s _i) = si'
Hlti: _i < equivocator_state_n s'

si : state X, equivocator_state_project s _i = Some si ∧ ( itemx : transition_item, match equivocator_state_project s _i with | Some sj => if decide (_i = _i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing _i) else Some (None, Existing _i) | None => None end = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
si: state X
Heqsi': default (equivocator_state_zero s) (Some si) = si'
Hlti: _i < equivocator_state_n s'
Hi: _i < equivocator_state_n s

si0 : state X, Some si = Some si0 ∧ ( itemx : transition_item, (if decide (_i = _i) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing _i) else Some (None, Existing _i)) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
si: state X
Heqsi': si = si'
Hlti: _i < equivocator_state_n s'
Hi: _i < equivocator_state_n s

si0 : state X, Some si = Some si0 ∧ ( itemx : transition_item, (if decide (_i = _i) then Some (Some {| l := lx; input := iom; destination := si; output := oom |}, Existing _i) else Some (None, Existing _i)) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ContinueWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ContinueWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = equivocator_state_n s'
si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Heq_oom: _oom = oom
Hlti: _i < equivocator_state_n s'
Hi: _i < equivocator_state_n s

si : state X, Some si' = Some si ∧ ( itemx : transition_item, (if decide (_i = _i) then Some (Some {| l := lx; input := iom; destination := si'; output := oom |}, Existing _i) else Some (None, Existing _i)) = Some (Some itemx, Existing _i))
by rewrite decide_True; eauto.
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ForkWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ForkWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: transition (ForkWith _i lx) (s', iom) = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ equivocator_state_descriptor_project s (Existing (equivocator_state_n s')) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
Ht: match equivocator_state_project s' _i with | Some si => let (si', om') := transition lx (si, iom) in (equivocator_state_extend s' si', om') | None => (s', iom) end = (s, oom)
s'i: state X
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ equivocator_state_descriptor_project s (Existing (equivocator_state_n s')) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i: state X
Ht: (let (si', om') := transition lx (s'i, iom) in (equivocator_state_extend s' si', om')) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ equivocator_state_descriptor_project s (Existing (equivocator_state_n s')) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i: state X
Ht: (let (si', om') := transition lx (s'i, iom) in (equivocator_state_extend s' si', om')) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Ht_pr: (si' : state X) (_oout : option message), transition lx (s'i, iom) = (si', _oout) → _oout = oom ∧ default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Ht_pr: (si'0 : state X) (_oout : option message), (si', _oom) = (si'0, _oout) → _oout = oom ∧ default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')) = si'0

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')) = si'

(i : nat) (si : state X), equivocator_state_project s i = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')) = si'

si : state X, equivocator_state_project s (equivocator_state_n s') = Some si ∧ ( itemx : transition_item, equivocator_vlsm_transition_item_project item (Existing (equivocator_state_n s')) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Heqsi': default (equivocator_state_zero s) (equivocator_state_project s (equivocator_state_n s')) = si'

si : state X, equivocator_state_project s (equivocator_state_n s') = Some si ∧ ( itemx : transition_item, match equivocator_state_project s (equivocator_state_n s') with | Some sj => if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s')) | None => None end = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
s_lst: state X
Heqsi': default (equivocator_state_zero s) (Some s_lst) = si'
Hlst: equivocator_state_n s' < equivocator_state_n s

si : state X, Some s_lst = Some si ∧ ( itemx : transition_item, (if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := s_lst; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s'))) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
s_lst: state X
Heqsi': s_lst = si'
Hlst: equivocator_state_n s' < equivocator_state_n s

si : state X, Some s_lst = Some si ∧ ( itemx : transition_item, (if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := s_lst; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s'))) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Hlst: equivocator_state_n s' < equivocator_state_n s

si : state X, Some si' = Some si ∧ ( itemx : transition_item, (if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si'; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s'))) = Some (Some itemx, Existing _i))
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Hlst: equivocator_state_n s' < equivocator_state_n s

itemx : transition_item, (if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si'; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s'))) = Some (Some itemx, Existing _i)
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Hlst: equivocator_state_n s' < equivocator_state_n s
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

itemx : transition_item, (if decide (equivocator_state_n s' = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := si'; output := oom |}, Existing _i) else Some (None, Existing (equivocator_state_n s'))) = Some (Some itemx, Existing _i)
message: Type
X: VLSM message
lx: label X
s, s': state (equivocator_vlsm X)
iom, oom: option message
_i: nat
Hsndl: Existing _i = Existing _i
item:= {| l := ForkWith _i lx; input := iom; destination := s; output := oom |}: transition_item
s'i, si': state X
_oom: option message
Hti: transition lx (s'i, iom) = (si', _oom)
Ht: (equivocator_state_extend s' si', _oom) = (s, oom)
Heqs'i: equivocator_state_project s' _i = Some s'i
Hv: valid lx (s'i, iom)
Ht_size: equivocator_state_n s = S (equivocator_state_n s')
Heq_oom: _oom = oom
Hlst: equivocator_state_n s' < equivocator_state_n s
Hs_lst: equivocator_state_n s = S (equivocator_state_last s)

itemx : transition_item, Some (Some {| l := lx; input := iom; destination := si'; output := oom |}, Existing _i) = Some (Some itemx, Existing _i)
by eexists. Qed.
The projection of a segment of an equivocator_vlsm valid trace is defined and a valid trace segment in the original VLSM.
message: Type
X: VLSM message
seed: message → Prop
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X
Hj: equivocator_state_project s j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project [] (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Ht: input_valid_transition (preloaded_vlsm (equivocator_vlsm X) seed) l (s', iom) (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
IHHbtr: equivocator_state_project f j = Some ej → (tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project tl (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end
(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project ({| l := l; input := iom; destination := s; output := oom |} :: tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X
Hj: equivocator_state_project s j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project [] (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X
Hj: equivocator_state_project s j = Some ej

di : MachineDescriptor X, equivocator_vlsm_trace_project [] (Existing j) = Some ([], di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej [] | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej [] end
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X
Hj: equivocator_state_project s j = Some ej

match Existing j with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej [] | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej [] end
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X
Hj: equivocator_state_project s j = Some ej

finite_valid_trace_from_to (preloaded_vlsm X seed) ej ej []
message: Type
X: VLSM message
seed: message → Prop
s: state (preloaded_vlsm (equivocator_vlsm X) seed)
Hs: valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s
j: nat
ej: state X

equivocator_state_project s j = Some ej → valid_state_prop (preloaded_vlsm X seed) ej
by apply preloaded_with_equivocator_state_project_valid_state.
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Ht: input_valid_transition (preloaded_vlsm (equivocator_vlsm X) seed) l (s', iom) (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
IHHbtr: equivocator_state_project f j = Some ej → (tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project tl (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project ({| l := l; input := iom; destination := s; output := oom |} :: tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Ht: input_valid_transition (preloaded_vlsm (equivocator_vlsm X) seed) l (s', iom) (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
IHHbtr: equivocator_state_project f j = Some ej → (tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project tl (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project (item :: tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
IHHbtr: equivocator_state_project f j = Some ej → (tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project tl (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tr end
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project (item :: tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project (item :: tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project ([item] ++ tl) (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), foldr (λ (item : transition_item) (result : option (list transition_item * MachineDescriptor X)), match result with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end) (Some ([], Existing j)) ([item] ++ tl) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), foldr (λ (item : transition_item) (result : option (list transition_item * MachineDescriptor X)), match result with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end) (foldr (λ (item : transition_item) (result : option (list transition_item * MachineDescriptor X)), match result with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end) (Some ([], Existing j)) tl) [item] = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), foldr (λ (item : transition_item) (result : option (list transition_item * MachineDescriptor X)), match result with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end) (Some (tlX, di')) [item] = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
di': MachineDescriptor X
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX end

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item di' with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
Hdi: s0 : state X, equivocator_state_project s i = Some s0 ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s0 ej tlX

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hchar: proper_descriptor X (Existing i) (destination item) → (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l item), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (VLSM.l item) (s, input item) → transition (VLSM.l item) ( s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop (Existing i) s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) (sx, input itemx) = ( destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) (Existing i) = equivocator_state_descriptor_project s descriptor' end)

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hchar: (oitem : option transition_item) (descriptor' : MachineDescriptor X), equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, descriptor') ∧ match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l item), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l item) | None => True end ∧ ( s : state (equivocator_vlsm X), valid (VLSM.l item) (s, input item) → transition (VLSM.l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop (Existing i) s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) (sx, input itemx) = ( destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) (Existing i) = equivocator_state_descriptor_project s descriptor' end)

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
oitem: option transition_item
descriptor': MachineDescriptor X
Hitem_pr: equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, descriptor')
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l item), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l item) | None => True end
Hchar2: s : state (equivocator_vlsm X), valid (VLSM.l item) (s, input item) → transition (VLSM.l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop (Existing i) s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) (Existing i) = equivocator_state_descriptor_project s descriptor' end

(tr : list transition_item) (di : MachineDescriptor X), match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tlX, odescriptor) | Some (None, odescriptor) => Some (tlX, odescriptor) | None => None end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
oitem: option transition_item
descriptor': MachineDescriptor X
Hitem_pr: equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, descriptor')
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l item), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l item) Hex) ∧ input item = input itemx ∧ output item = output itemx ∧ equivocator_state_descriptor_project (destination item) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l item) | None => True end
Hchar2: s : state (equivocator_vlsm X), valid (VLSM.l item) (s, input item) → transition (VLSM.l item) (s, input item) = (destination item, output item) → proper_descriptor X descriptor' s ∧ previous_state_descriptor_prop (Existing i) s descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination item) (Existing i) = equivocator_state_descriptor_project s descriptor' end

(tr : list transition_item) (di : MachineDescriptor X), match oitem with | Some item' => Some (item' :: tlX, descriptor') | None => Some (tlX, descriptor') end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
oitem: option transition_item
descriptor': MachineDescriptor X
Hchar2: s0 : state (equivocator_vlsm X), valid (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) (s0, input {| l := l; input := iom; destination := s; output := oom |}) → transition (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) (s0, input {| l := l; input := iom; destination := s; output := oom |}) = (destination {| l := l; input := iom; destination := s; output := oom |}, output {| l := l; input := iom; destination := s; output := oom |}) → proper_descriptor X descriptor' s0 ∧ previous_state_descriptor_prop (Existing i) s0 descriptor' ∧ match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s0 descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) ( sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = equivocator_state_descriptor_project s0 descriptor' end
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) Hex) ∧ input {| l := l; input := iom; destination := s; output := oom |} = input itemx ∧ output {| l := l; input := iom; destination := s; output := oom |} = output itemx ∧ equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) | None => True end
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (oitem, descriptor')

(tr : list transition_item) (di : MachineDescriptor X), match oitem with | Some item' => Some (item' :: tlX, descriptor') | None => Some (tlX, descriptor') end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
oitem: option transition_item
descriptor': MachineDescriptor X
Hchar1: match oitem with | Some itemx => ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}), VLSM.l itemx = existing_equivocator_label_extract X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) Hex) ∧ input {| l := l; input := iom; destination := s; output := oom |} = input itemx ∧ output {| l := l; input := iom; destination := s; output := oom |} = output itemx ∧ equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = destination itemx ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) | None => True end
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (oitem, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: match oitem with | Some itemx => sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemx) (sx, input itemx) ∧ transition (VLSM.l itemx) (sx, input itemx) = (destination itemx, output itemx) | None => equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = equivocator_state_descriptor_project s' descriptor' end

(tr : list transition_item) (di : MachineDescriptor X), match oitem with | Some item' => Some (item' :: tlX, descriptor') | None => Some (tlX, descriptor') end = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tr | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tr end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
itemX: transition_item
descriptor': MachineDescriptor X
Hchar1: ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}), VLSM.l itemX = existing_equivocator_label_extract X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) Hex) ∧ input {| l := l; input := iom; destination := s; output := oom |} = input itemX ∧ output {| l := l; input := iom; destination := s; output := oom |} = output itemX ∧ equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = destination itemX ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := iom; destination := s; output := oom |})
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (Some itemX, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
descriptor': MachineDescriptor X
Hchar1: True
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (None, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = equivocator_state_descriptor_project s' descriptor'
match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
descriptor': MachineDescriptor X
Hchar1: True
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (None, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = equivocator_state_descriptor_project s' descriptor'

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX end
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
descriptor': MachineDescriptor X
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: match descriptor' with | NewMachine _ => equivocator_state_n s' ≤ i | Existing id' => id' ≤ i end
Hchar2: default (equivocator_state_zero s) (equivocator_state_project s i) = equivocator_state_descriptor_project s' descriptor'

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX end
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
descriptor': MachineDescriptor X
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: match descriptor' with | NewMachine _ => equivocator_state_n s' ≤ i | Existing id' => id' ≤ i end
Hchar2: default (equivocator_state_zero s) (Some si) = equivocator_state_descriptor_project s' descriptor'

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX end
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
sn: state X
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hprevious: equivocator_state_n s' ≤ i
Hchar2: si = sn

initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
Hproper': proper_descriptor X (Existing i') s'
Hprevious: i' ≤ i
Hchar2: si = default (equivocator_state_zero s') (equivocator_state_project s' i')
s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
sn: state X
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hprevious: equivocator_state_n s' ≤ i
Hchar2: si = sn

initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej tlX
by subst si.
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
Hproper': proper_descriptor X (Existing i') s'
Hprevious: i' ≤ i
Hchar2: si = default (equivocator_state_zero s') (equivocator_state_project s' i')

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
s'i': state X
Hproper': equivocator_state_project s' i' = Some s'i'
Hprevious: i' ≤ i
Hchar2: si = default (equivocator_state_zero s') (equivocator_state_project s' i')

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
s'i': state X
Hproper': equivocator_state_project s' i' = Some s'i'
Hprevious: i' ≤ i
Hchar2: si = default (equivocator_state_zero s') (Some s'i')

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
s'i': state X
Hproper': equivocator_state_project s' i' = Some s'i'
Hprevious: i' ≤ i
Hchar2: si = s'i'

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
iom, oom: option message
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: equivocator_valid X l (s', iom)
Ht: equivocator_transition X l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
i': nat
Hchar1: True
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := iom; destination := sj; output := oom |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (None, Existing i')
Hproper': equivocator_state_project s' i' = Some si
Hprevious: i' ≤ i

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej tlX
by exists si.
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
itemX: transition_item
descriptor': MachineDescriptor X
Hchar1: ( Hex : existing_equivocator_label X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}), VLSM.l itemX = existing_equivocator_label_extract X (VLSM.l {| l := l; input := iom; destination := s; output := oom |}) Hex) ∧ input {| l := l; input := iom; destination := s; output := oom |} = input itemX ∧ output {| l := l; input := iom; destination := s; output := oom |} = output itemX ∧ equivocator_state_descriptor_project (destination {| l := l; input := iom; destination := s; output := oom |}) (Existing i) = destination itemX ∧ descriptor' = equivocator_label_descriptor (VLSM.l {| l := l; input := iom; destination := s; output := oom |})
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (Some itemX, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
itemX: transition_item
descriptor': MachineDescriptor X
Hchar1: ( Hex : existing_equivocator_label X l, VLSM.l itemX = existing_equivocator_label_extract X l Hex) ∧ iom = input itemX ∧ oom = output itemX ∧ default (equivocator_state_zero s) (equivocator_state_project s i) = destination itemX ∧ descriptor' = equivocator_label_descriptor l
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (Some itemX, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
itemX: transition_item
descriptor': MachineDescriptor X
Hchar1: ( Hex : existing_equivocator_label X l, VLSM.l itemX = existing_equivocator_label_extract X l Hex) ∧ iom = input itemX ∧ oom = output itemX ∧ default (equivocator_state_zero s) (Some si) = destination itemX ∧ descriptor' = equivocator_label_descriptor l
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (Some itemX, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
iom, oom: option message
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) iom
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
itemX: transition_item
descriptor': MachineDescriptor X
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Heqiom: iom = input itemX
Hoom: oom = output itemX
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
Hdl: descriptor' = equivocator_label_descriptor l
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (Some itemX, descriptor')
Hproper': proper_descriptor X descriptor' s'
Hprevious: previous_state_descriptor_prop (Existing i) s' descriptor'
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' descriptor' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)

match descriptor' with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
Hchar2: sx : state X, sx = equivocator_state_descriptor_project s' (equivocator_label_descriptor l) → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (equivocator_label_descriptor l)
Hproper': proper_descriptor X (equivocator_label_descriptor l) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, equivocator_label_descriptor l)

match equivocator_label_descriptor l with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX) | Existing i => s : state X, equivocator_state_project s' i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX) end
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hchar2: sx : state X, sx = sn → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, NewMachine sn)

initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
Hchar2: sx : state X, sx = default (equivocator_state_zero s') (equivocator_state_project s' i') → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Hproper': proper_descriptor X (Existing i') s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')
s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hchar2: sx : state X, sx = sn → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, NewMachine sn)

initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: previous_state_descriptor_prop (Existing i) s' (NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, NewMachine sn)
HvX: valid (VLSM.l itemX) (sn, input itemX)
HtX: transition (VLSM.l itemX) (sn, input itemX) = (destination itemX, output itemX)

initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: previous_state_descriptor_prop (Existing i) s' (NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, NewMachine sn)
HvX: valid (VLSM.l itemX) (sn, input itemX)
HtX: transition (VLSM.l itemX) (sn, input itemX) = (destination itemX, output itemX)

finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
l0: label X
input: option message
destination: state X
output: option message
Ht: transition l (s', VLSM.input {| l := l0; input := input; destination := destination; output := output |}) = (s, VLSM.output {| l := l0; input := input; destination := destination; output := output |})
Hv: valid l (s', VLSM.input {| l := l0; input := input; destination := destination; output := output |})
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (VLSM.input {| l := l0; input := input; destination := destination; output := output |})
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l {| l := l0; input := input; destination := destination; output := output |} = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = VLSM.destination {| l := l0; input := input; destination := destination; output := output |}
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: previous_state_descriptor_prop (Existing i) s' (NewMachine sn)
Hproper': proper_descriptor X (NewMachine sn) s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := VLSM.input {| l := l0; input := input; destination := destination; output := output |}; destination := s; output := VLSM.output {| l := l0; input := input; destination := destination; output := output |} |} (Existing i) = Some (Some {| l := l0; input := input; destination := destination; output := output |}, NewMachine sn)
HvX: valid (VLSM.l {| l := l0; input := input; destination := destination; output := output |}) (sn, VLSM.input {| l := l0; input := input; destination := destination; output := output |})
HtX: transition (VLSM.l {| l := l0; input := input; destination := destination; output := output |}) (sn, VLSM.input {| l := l0; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l0; input := input; destination := destination; output := output |}, VLSM.output {| l := l0; input := input; destination := destination; output := output |})

finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej ({| l := l0; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
l0: label X
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: l0 = existing_equivocator_label_extract X l Hex
Hsi: si = destination
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := l0; input := input; destination := destination; output := output |}, NewMachine sn)
HvX: valid l0 (sn, input)
HtX: transition l0 (sn, input) = (destination, output)

finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej ({| l := l0; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
l0: label X
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: l0 = existing_equivocator_label_extract X l Hex
Hsi: si = destination
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := si; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := si; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := l0; input := input; destination := destination; output := output |}, NewMachine sn)
HvX: valid l0 (sn, input)
HtX: transition l0 (sn, input) = (destination, output)

finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej ({| l := l0; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)

finite_valid_trace_from_to (preloaded_vlsm X seed) sn ej ({| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)

input_valid_transition (preloaded_vlsm X seed) (existing_equivocator_label_extract X l Hex) (sn, input) (destination, output)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)

valid_state_prop (preloaded_vlsm X seed) sn
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)
option_valid_message_prop (preloaded_vlsm X seed) input
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)

valid_state_prop (preloaded_vlsm X seed) sn
by apply initial_state_is_valid.
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
sn: state X
Hd: equivocator_label_descriptor l = NewMachine sn
Hprevious: equivocator_state_n s' ≤ i
Hproper': initial_state_prop sn
HtX: transition (existing_equivocator_label_extract X l Hex) (sn, input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (sn, input)
Hitem_pr: match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := destination; output := output |}, Existing i0) else Some (None, Existing i) end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, NewMachine sn)

option_valid_message_prop (preloaded_vlsm X seed) input
by apply preloaded_with_equivocator_state_project_valid_message.
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
Hchar2: sx : state X, sx = default (equivocator_state_zero s') (equivocator_state_project s' i') → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Hproper': proper_descriptor X (Existing i') s'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')

s : state X, equivocator_state_project s' i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hchar2: sx : state X, sx = default (equivocator_state_zero s') (Some s'i') → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')

s : state X, Some s'i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hchar2: sx : state X, sx = s'i' → valid (VLSM.l itemX) (sx, input itemX) ∧ transition (VLSM.l itemX) (sx, input itemX) = (destination itemX, output itemX)
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')

s : state X, Some s'i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')
HvX: valid (VLSM.l itemX) (s'i', input itemX)
HtX: transition (VLSM.l itemX) (s'i', input itemX) = (destination itemX, output itemX)

s : state X, Some s'i' = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X seed) s ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
itemX: transition_item
Ht: transition l (s', input itemX) = (s, output itemX)
Hv: valid l (s', input itemX)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (input itemX)
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l itemX = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = destination itemX
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := input itemX; destination := s; output := output itemX |} (Existing i) = Some (Some itemX, Existing i')
HvX: valid (VLSM.l itemX) (s'i', input itemX)
HtX: transition (VLSM.l itemX) (s'i', input itemX) = (destination itemX, output itemX)

finite_valid_trace_from_to (preloaded_vlsm X seed) s'i' ej (itemX :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: state (preloaded_vlsm (equivocator_vlsm X) seed)
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': state (preloaded_vlsm (equivocator_vlsm X) seed)
l: label (preloaded_vlsm (equivocator_vlsm X) seed)
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
l0: label X
input: option message
destination: state X
output: option message
Ht: transition l (s', VLSM.input {| l := l0; input := input; destination := destination; output := output |}) = (s, VLSM.output {| l := l0; input := input; destination := destination; output := output |})
Hv: valid l (s', VLSM.input {| l := l0; input := input; destination := destination; output := output |})
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) (VLSM.input {| l := l0; input := input; destination := destination; output := output |})
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: VLSM.l {| l := l0; input := input; destination := destination; output := output |} = existing_equivocator_label_extract X l Hex
Hsi: default (equivocator_state_zero s) (Some si) = VLSM.destination {| l := l0; input := input; destination := destination; output := output |}
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: previous_state_descriptor_prop (Existing i) s' (Existing i')
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: equivocator_vlsm_transition_item_project {| l := l; input := VLSM.input {| l := l0; input := input; destination := destination; output := output |}; destination := s; output := VLSM.output {| l := l0; input := input; destination := destination; output := output |} |} (Existing i) = Some (Some {| l := l0; input := input; destination := destination; output := output |}, Existing i')
HvX: valid (VLSM.l {| l := l0; input := input; destination := destination; output := output |}) (s'i', VLSM.input {| l := l0; input := input; destination := destination; output := output |})
HtX: transition (VLSM.l {| l := l0; input := input; destination := destination; output := output |}) (s'i', VLSM.input {| l := l0; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := l0; input := input; destination := destination; output := output |}, VLSM.output {| l := l0; input := input; destination := destination; output := output |})

finite_valid_trace_from_to (preloaded_vlsm X seed) s'i' ej ({| l := l0; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
l0: label X
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
si: state X
Heqsi: equivocator_state_project s i = Some si
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) si ej tlX
Hex: existing_equivocator_label X l
Hl: l0 = existing_equivocator_label_extract X l Hex
Hsi: si = destination
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := l0; input := input; destination := destination; output := output |}, Existing i')
HvX: valid l0 (s'i', input)
HtX: transition l0 (s'i', input) = (destination, output)

finite_valid_trace_from_to (preloaded_vlsm X seed) s'i' ej ({| l := l0; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')

finite_valid_trace_from_to (preloaded_vlsm X seed) s'i' ej ({| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |} :: tlX)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')

input_valid_transition (preloaded_vlsm X seed) (existing_equivocator_label_extract X l Hex) (s'i', input) (destination, output)
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')

valid_state_prop (preloaded_vlsm X seed) s'i'
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')
option_valid_message_prop (preloaded_vlsm X seed) input
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')

valid_state_prop (preloaded_vlsm X seed) s'i'
by eapply preloaded_with_equivocator_state_project_valid_state.
message: Type
X: VLSM message
seed: message → Prop
s, f: Equivocators.bounded_state_copies X
tl: list transition_item
Hbtr: finite_valid_trace_from_to (preloaded_vlsm (equivocator_vlsm X) seed) s f tl
s': Equivocators.bounded_state_copies X
l: EquivocatorLabel X
Hs': valid_state_prop (preloaded_vlsm (equivocator_vlsm X) seed) s'
input: option message
destination: state X
output: option message
Ht: equivocator_transition X l (s', input) = (s, output)
Hv: equivocator_valid X l (s', input)
Hiom: option_valid_message_prop (preloaded_vlsm (equivocator_vlsm X) seed) input
j: nat
ej: state X
Hj: equivocator_state_project f j = Some ej
tlX: list transition_item
i: nat
Htl_pr: equivocator_vlsm_trace_project tl (Existing j) = Some (tlX, Existing i)
HltX: finite_valid_trace_from_to (preloaded_vlsm X seed) destination ej tlX
Heqsi: equivocator_state_project s i = Some destination
Hex: existing_equivocator_label X l
i': nat
Hd: equivocator_label_descriptor l = Existing i'
s'i': state X
Hprevious: i' ≤ i
Heqs'i': equivocator_state_project s' i' = Some s'i'
HtX: transition (existing_equivocator_label_extract X l Hex) (s'i', input) = (destination, output)
HvX: valid (existing_equivocator_label_extract X l Hex) (s'i', input)
Hitem_pr: match equivocator_state_project s i with | Some sj => match l with | Spawn sn => if decide (i = equivocator_state_last s) then Some (None, NewMachine sn) else Some (None, Existing i) | ContinueWith i0 lx => if decide (i0 = i) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) | ForkWith i0 lx => if decide (i = equivocator_state_last s) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i0) else Some (None, Existing i) end | None => None end = Some (Some {| l := existing_equivocator_label_extract X l Hex; input := input; destination := destination; output := output |}, Existing i')

option_valid_message_prop (preloaded_vlsm X seed) input
by apply preloaded_with_equivocator_state_project_valid_message. Qed.
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
di: MachineDescriptor X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di)
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr end

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
di: MachineDescriptor X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di)
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr end

match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hdi: initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr

initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
Hdi: s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr
s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hdi: initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr

initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hsn: initial_state_prop sn
Htr: finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr

initial_state_prop sn ∧ finite_valid_trace_from_to X sn ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hsn: initial_state_prop sn
Htr: finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) sn ej tr

finite_valid_trace_from_to X sn ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hsn: initial_state_prop sn
Htr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := X; vlsm_machine := preloaded_vlsm X (λ _ : message, False) |}; vlsm_machine := {| vlsm_type := X; vlsm_machine := X |} |} sn ej tr

finite_valid_trace_from_to X sn ej tr
by clear -Htr; destruct X.
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
Hdi: s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr

s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
s: state X
Hpr_bs_i: equivocator_state_project bs i = Some s
Htr: finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr

s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
s: state X
Hpr_bs_i: equivocator_state_project bs i = Some s
Htr: finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, False)) s ej tr

finite_valid_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := equivocator_vlsm X; vlsm_machine := equivocator_vlsm X |}; vlsm_machine := {| vlsm_type := equivocator_vlsm X; vlsm_machine := preloaded_vlsm (equivocator_vlsm X) (λ _ : message, False) |} |} bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
s: state X
Hpr_bs_i: equivocator_state_project bs i = Some s
Htr: finite_valid_trace_from_to {| vlsm_type := {| vlsm_type := X; vlsm_machine := preloaded_vlsm X (λ _ : message, False) |}; vlsm_machine := {| vlsm_type := X; vlsm_machine := X |} |} s ej tr

finite_valid_trace_from_to X s ej tr
by clear -Htr; destruct X. Qed.
The projection of a segment of a valid trace from the preloaded_with_all_messages_vlsm corresponding to the equivocator_vlsm is defined and it is a valid trace segment in the preloaded_with_all_messages_vlsm corresponding to the original vlsm.
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
di: MachineDescriptor X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di)
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) s ej tr end

(tr : list transition_item) (di : MachineDescriptor X), equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di) ∧ match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
di: MachineDescriptor X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, di)
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) s ej tr end

match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr | Existing i => s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr end
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hdi: initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) sn ej tr

initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
Hdi: s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) s ej tr
s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
sn: state X
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, NewMachine sn)
Hdi: initial_state_prop sn ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) sn ej tr

initial_state_prop sn ∧ finite_constrained_trace_from_to X sn ej tr
by destruct Hdi as [Hsn Htr].
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
Hdi: s : state X, equivocator_state_project bs i = Some s ∧ finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) s ej tr

s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr
message: Type
X: VLSM message
bs, be: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from_to (equivocator_vlsm X) bs be btr
j: nat
ej: state X
Hj: equivocator_state_project be j = Some ej
tr: list transition_item
i: nat
Hbtr_pr: equivocator_vlsm_trace_project btr (Existing j) = Some (tr, Existing i)
s: state X
Hpr_bs_i: equivocator_state_project bs i = Some s
Htr: finite_valid_trace_from_to (preloaded_vlsm X (λ _ : message, True)) s ej tr

s : state X, equivocator_state_project bs i = Some s ∧ finite_constrained_trace_from_to X s ej tr
by eexists. Qed.
If equivocator_vlsm_trace_project does not fail, then the index of the machine descriptor is valid for the last state of the trace argument.
message: Type
X: VLSM message
tr: list transition_item
Hntr: tr ≠ []
j: nat
HtrX: is_Some (equivocator_vlsm_trace_project tr (Existing j))
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj
message: Type
X: VLSM message
tr: list transition_item
Hntr: tr ≠ []
j: nat
HtrX: is_Some (equivocator_vlsm_trace_project tr (Existing j))
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj
message: Type
X: VLSM message
tr: list transition_item
Hntr: {l' : list transition_item & {a : transition_item | tr = l' ++ [a]}}
j: nat
HtrX: is_Some (equivocator_vlsm_trace_project tr (Existing j))
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj
message: Type
X: VLSM message
tr, suffix: list transition_item
x: transition_item
Heq: tr = suffix ++ [x]
j: nat
HtrX: is_Some (equivocator_vlsm_trace_project tr (Existing j))
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
HtrX: is_Some (equivocator_vlsm_trace_project (suffix ++ [x]) (Existing j))
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is (suffix ++ [x])) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
p: (list transition_item * MachineDescriptor X)%type
Htr: equivocator_vlsm_trace_project (suffix ++ [x]) (Existing j) = Some p
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is (suffix ++ [x])) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d: MachineDescriptor X
Htr: equivocator_vlsm_trace_project (suffix ++ [x]) (Existing j) = Some (trX, d)
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is (suffix ++ [x])) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d: MachineDescriptor X
Htr: (dmiddle : MachineDescriptor X) (prefix suffix0 : list transition_item) (_ : equivocator_vlsm_trace_project suffix dmiddle = Some (prefix, d)) (_ : equivocator_vlsm_trace_project [x] (Existing j) = Some (suffix0, dmiddle)), trX = prefix ++ suffix0
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is (suffix ++ [x])) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
lx: list transition_item
Hx: equivocator_vlsm_trace_project [x] (Existing j) = Some (lx, dmiddle)
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (finite_trace_last is (suffix ++ [x])) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
lx: list transition_item
Hx: equivocator_vlsm_trace_project [x] (Existing j) = Some (lx, dmiddle)
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
lx: list transition_item
dj: MachineDescriptor X
Heqdj: dj = Existing j
Hx: equivocator_vlsm_trace_project [x] dj = Some (lx, dmiddle)
is: state (equivocator_vlsm X)

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
lx: list transition_item
dj: MachineDescriptor X
Heqdj: dj = Existing j
Hx: match equivocator_vlsm_transition_item_project x dj with | Some (Some item', odescriptor) => Some ([item'], odescriptor) | Some (None, odescriptor) => Some ([], odescriptor) | None => None end = Some (lx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
lx: list transition_item
dj: MachineDescriptor X
Heqdj: dj = Existing j
_x: option transition_item
_dmiddle: MachineDescriptor X
Hx': equivocator_vlsm_transition_item_project x dj = Some (_x, _dmiddle)
Hx: match _x with | Some item' => Some ([item'], _dmiddle) | None => Some ([], _dmiddle) end = Some (lx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle, dj: MachineDescriptor X
Heqdj: dj = Existing j
itemx: transition_item
Hx': equivocator_vlsm_transition_item_project x dj = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle, dj: MachineDescriptor X
Heqdj: dj = Existing j
Hx': equivocator_vlsm_transition_item_project x dj = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X
sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle, dj: MachineDescriptor X
Heqdj: dj = Existing j
itemx: transition_item
Hx': equivocator_vlsm_transition_item_project x dj = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
itemx: transition_item
Hx': equivocator_vlsm_transition_item_project x (Existing j) = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: label (equivocator_type X)
input: option message
destination: state (equivocator_type X)
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
itemx: transition_item
Hx': equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing j) = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: label (equivocator_type X)
input: option message
destination: state (equivocator_type X)
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
itemx: transition_item
Hx': match equivocator_state_project destination j with | Some sj => match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) end | None => None end = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: label (equivocator_type X)
input: option message
destination: state (equivocator_type X)
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
itemx: transition_item
Hx': match equivocator_state_project destination j with | Some sj => match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) end | None => None end = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project destination j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: label (equivocator_type X)
input: option message
destination: state (equivocator_type X)
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
itemx: transition_item
s: state X
Hx': match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) end = Some (Some itemx, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, Some s = Some sj
by eexists.
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle, dj: MachineDescriptor X
Heqdj: dj = Existing j
Hx': equivocator_vlsm_transition_item_project x dj = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
x: transition_item
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
Hx': equivocator_vlsm_transition_item_project x (Existing j) = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (destination x) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: label (equivocator_type X)
input: option message
destination: state (equivocator_type X)
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
Hx': equivocator_vlsm_transition_item_project {| l := l; input := input; destination := destination; output := output |} (Existing j) = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project (VLSM.destination {| l := l; input := input; destination := destination; output := output |}) j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
Hx': match equivocator_state_project destination j with | Some sj => match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := sj; output := output |}, Existing i) else Some (None, Existing j) end | None => None end = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, equivocator_state_project destination j = Some sj
message: Type
X: VLSM message
suffix: list transition_item
l: EquivocatorLabel X
input: option message
destination: Equivocators.bounded_state_copies X
output: option message
j: nat
trX: list transition_item
d, dmiddle: MachineDescriptor X
s: state X
Hx': match l with | Spawn sn => if decide (j = equivocator_state_last destination) then Some (None, NewMachine sn) else Some (None, Existing j) | ContinueWith i lx => if decide (i = j) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) | ForkWith i lx => if decide (j = equivocator_state_last destination) then Some (Some {| l := lx; input := input; destination := s; output := output |}, Existing i) else Some (None, Existing j) end = Some (None, dmiddle)
is: Equivocators.bounded_state_copies X

sj : state X, Some s = Some sj
by eexists. Qed.
Projecting a valid trace segment on an index which is valid for the first state of the trace does not fail and yields the same index.
message: Type
X: VLSM message
bs: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from (equivocator_vlsm X) bs btr
i: nat
si: state X
Hi: equivocator_state_project bs i = Some si

tr : list transition_item, equivocator_vlsm_trace_project btr (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
bs: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from (equivocator_vlsm X) bs btr
i: nat
si: state X
Hi: equivocator_state_project bs i = Some si

tr : list transition_item, equivocator_vlsm_trace_project btr (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
bs: state (equivocator_vlsm X)
btr: list transition_item
Hbtr: finite_constrained_trace_from (equivocator_vlsm X) bs btr

(i : nat) (si : state X), equivocator_state_project bs i = Some si → tr : list transition_item, equivocator_vlsm_trace_project btr (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hs: valid_state_prop (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s
i: nat
si: state X
Hi: equivocator_state_project s i = Some si

tr : list transition_item, equivocator_vlsm_trace_project [] (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Ht: input_valid_transition (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) l ( s', iom) (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
tr : list transition_item, equivocator_vlsm_trace_project ({| l := l; input := iom; destination := s; output := oom |} :: tl) (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hs: valid_state_prop (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s
i: nat
si: state X
Hi: equivocator_state_project s i = Some si

tr : list transition_item, equivocator_vlsm_trace_project [] (Existing i) = Some (tr, Existing i)
by exists [].
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Ht: input_valid_transition (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) l ( s', iom) (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si

tr : list transition_item, equivocator_vlsm_trace_project ({| l := l; input := iom; destination := s; output := oom |} :: tl) (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Ht: input_valid_transition (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) l ( s', iom) (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}

tr : list transition_item, equivocator_vlsm_trace_project (item :: tl) (Existing i) = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Ht: input_valid_transition (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) l ( s', iom) (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
Hitem: si' : state X, equivocator_state_project s' i = Some si' → si : state X, equivocator_state_project s i = Some si ∧ ( oitem : option transition_item, equivocator_vlsm_transition_item_project {| l := l; input := iom; destination := s; output := oom |} (Existing i) = Some (oitem, Existing i))

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
Hitem: si' : state X, equivocator_state_project s' i = Some si' → si : state X, equivocator_state_project s i = Some si ∧ ( oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i))

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
IHHbtr: (i : nat) (si : state X), equivocator_state_project s i = Some si → tr : list transition_item, equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
Hitem: oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i)

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
Hitem: oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr : list transition_item, match equivocator_vlsm_trace_project tl (Existing i) with | Some (r, idescriptor) => match equivocator_vlsm_transition_item_project item idescriptor with | Some (Some item', odescriptor) => Some (item' :: r, odescriptor) | Some (None, odescriptor) => Some (r, odescriptor) | None => None end | None => None end = Some (tr, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
Hitem: oitem : option transition_item, equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tr, odescriptor) | Some (None, odescriptor) => Some (tr, odescriptor) | None => None end = Some (tr0, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
oitem: option transition_item
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, match equivocator_vlsm_transition_item_project item (Existing i) with | Some (Some item', odescriptor) => Some (item' :: tr, odescriptor) | Some (None, odescriptor) => Some (tr, odescriptor) | None => None end = Some (tr0, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
oitem: option transition_item
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (oitem, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, match oitem with | Some item' => Some (item' :: tr, Existing i) | None => Some (tr, Existing i) end = Some (tr0, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
itemx: transition_item
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, Some (itemx :: tr, Existing i) = Some (tr0, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (None, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)
tr0 : list transition_item, Some (tr, Existing i) = Some (tr0, Existing i)
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
itemx: transition_item
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (Some itemx, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, Some (itemx :: tr, Existing i) = Some (tr0, Existing i)
by exists (itemx :: tr).
message: Type
X: VLSM message
s: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tl: list transition_item
Hbtr: finite_valid_trace_from (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) s tl
s': state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
iom, oom: option message
l: label (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
Hv: valid l (s', iom)
Ht: transition l (s', iom) = (s, oom)
i: nat
si: state X
Hi: equivocator_state_project s' i = Some si
item: transition_item
Heqitem: item = {| l := l; input := iom; destination := s; output := oom |}
si': state X
Hsi': equivocator_state_project s i = Some si'
Hoitem: equivocator_vlsm_transition_item_project item (Existing i) = Some (None, Existing i)
tr: list transition_item
Htr: equivocator_vlsm_trace_project tl (Existing i) = Some (tr, Existing i)

tr0 : list transition_item, Some (tr, Existing i) = Some (tr0, Existing i)
by exists tr. Qed.
An inversion lemma about projections of a valid trace.
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
Hj: is_Some (equivocator_vlsm_trace_project tr (Existing j)) → is : state (equivocator_vlsm X), sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
Hj: is : state (equivocator_vlsm X), sj : state X, equivocator_state_project (finite_trace_last is tr) j = Some sj

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project (finite_trace_last is tr) j = Some fsj

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj

fsj : state X, equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj

equivocator_state_project fs j = Some fsj ∧ match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj

match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
trX': list transition_item
di': MachineDescriptor X
HtrX': equivocator_vlsm_trace_project tr (Existing j) = Some (trX', di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn fsj trX' | Existing i => s : state X, equivocator_state_project is i = Some s ∧ finite_constrained_trace_from_to X s fsj trX' end

match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
trX': list transition_item
di': MachineDescriptor X
HtrX': Some (trX, di) = Some (trX', di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn fsj trX' | Existing i => s : state X, equivocator_state_project is i = Some s ∧ finite_constrained_trace_from_to X s fsj trX' end

match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
trX': list transition_item
di': MachineDescriptor X
HtrX': Some (trX, di) = Some (trX', di')
Hdi: match di' with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn fsj trX' | Existing i => s : state X, equivocator_state_project is i = Some s ∧ finite_constrained_trace_from_to X s fsj trX' end
H0: trX = trX'
H1: di = di'

match di' with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX' | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX' ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
HtrX': Some (trX, di) = Some (trX, di)
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn fsj trX | Existing i => s : state X, equivocator_state_project is i = Some s ∧ finite_constrained_trace_from_to X s fsj trX end

match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j: nat
di: MachineDescriptor X
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, di)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
Hdi: match di with | NewMachine sn => initial_state_prop sn ∧ finite_constrained_trace_from_to X sn fsj trX | Existing i => s : state X, equivocator_state_project is i = Some s ∧ finite_constrained_trace_from_to X s fsj trX end

match di with | NewMachine sn => finite_constrained_trace_init_to X sn fsj trX | Existing i => isi : state X, equivocator_state_project is i = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi) end
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j, n: nat
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, Existing n)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
Hdi: s : state X, equivocator_state_project is n = Some s ∧ finite_constrained_trace_from_to X s fsj trX

isi : state X, equivocator_state_project is n = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi)
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j, n: nat
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, Existing n)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
isi: state X
Hisi: equivocator_state_project is n = Some isi
HtrX': finite_constrained_trace_from_to X isi fsj trX

isi : state X, equivocator_state_project is n = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi)
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j, n: nat
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, Existing n)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
isi: state X
Hisi: equivocator_state_project is n = Some isi
HtrX': finite_constrained_trace_from_to X isi fsj trX

equivocator_state_project is n = Some isi ∧ finite_constrained_trace_from_to X isi fsj trX ∧ (initial_state_prop is → initial_state_prop isi)
message: Type
X: VLSM message
is, fs: state (preloaded_with_all_messages_vlsm (equivocator_vlsm X))
tr: list transition_item
Hntr: tr ≠ []
Htr: finite_constrained_trace_from_to (equivocator_vlsm X) is fs tr
j, n: nat
trX: list transition_item
HtrX: equivocator_vlsm_trace_project tr (Existing j) = Some (trX, Existing n)
fsj: state X
Hfsj: equivocator_state_project fs j = Some fsj
isi: state X
Hisi: equivocator_state_project is n = Some isi
HtrX': finite_constrained_trace_from_to X isi fsj trX

initial_state_prop is → initial_state_prop isi
by apply (equivocator_vlsm_initial_state_preservation_rev X _ _ _ Hisi). Qed. Definition equivocator_label_zero_project (l : equivocator_label X) : option (label X) := match l with | ContinueWith 0 li => Some li | _ => None end.
message: Type
X: VLSM message

VLSM_projection (equivocator_vlsm X) X equivocator_label_zero_project equivocator_state_zero
message: Type
X: VLSM message

VLSM_projection (equivocator_vlsm X) X equivocator_label_zero_project equivocator_state_zero
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
H0: valid lX (s, om)

valid lY (equivocator_state_zero s, om)
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')
transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: equivocator_label_zero_project lX = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
s: state (equivocator_vlsm X)
H: initial_state_prop s
initial_state_prop (equivocator_state_zero s)
message: Type
X: VLSM message
m: message
H: valid_message_prop (equivocator_vlsm X) m
valid_message_prop X m
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
H0: valid lX (s, om)

valid lY (equivocator_state_zero s, om)
by destruct lX as [sn | [| i] lX | [| i] lX]; inversion H; subst.
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition (ContinueWith 0 lY) (s, om) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lY (si, om) in (equivocator_state_update s 0 si', om') | None => (s, om) end = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (let (si', om') := transition lY (equivocator_state_zero s, om) in (equivocator_state_update s 0 si', om')) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
by destruct (transition _ _ _); inversion_clear H0.
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: equivocator_label_zero_project lX = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: match lX with | ContinueWith 0 li => Some li | _ => None end = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
sn: state X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (equivocator_state_extend s sn, None) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_update s (S i) si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
sn: state X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (equivocator_state_extend s sn, None) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by inversion H0.
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_update s (S i) si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by destruct (equivocator_state_project _ _); [destruct (transition _ _ _) |]; inversion H0.
message: Type
X: VLSM message
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (let (si', om') := transition lX (equivocator_state_zero s, om) in (equivocator_state_extend s si', om')) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by destruct (transition _ _ _); inversion_clear H0.
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by destruct (equivocator_state_project _ _); [destruct (transition _ _ _) |]; inversion_clear H0.
message: Type
X: VLSM message
s: state (equivocator_vlsm X)
H: initial_state_prop s

initial_state_prop (equivocator_state_zero s)
by apply H.
message: Type
X: VLSM message
m: message
H: valid_message_prop (equivocator_vlsm X) m

valid_message_prop X m
by apply equivocator_state_project_valid_message. Qed.
message: Type
X: VLSM message

VLSM_projection (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) (preloaded_with_all_messages_vlsm X) equivocator_label_zero_project equivocator_state_zero
message: Type
X: VLSM message

VLSM_projection (preloaded_with_all_messages_vlsm (equivocator_vlsm X)) (preloaded_with_all_messages_vlsm X) equivocator_label_zero_project equivocator_state_zero
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
H0: valid lX (s, om)

valid lY (equivocator_state_zero s, om)
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')
transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: equivocator_label_zero_project lX = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
s: state (equivocator_vlsm X)
H: initial_state_prop s
initial_state_prop (equivocator_state_zero s)
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
H0: valid lX (s, om)

valid lY (equivocator_state_zero s, om)
by destruct lX as [sn | [| i] lX | [| i] lX]; inversion H; subst.
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
lY: label X
H: equivocator_label_zero_project lX = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition (ContinueWith 0 lY) (s, om) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lY (si, om) in (equivocator_state_update s 0 si', om') | None => (s, om) end = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
message: Type
X: VLSM message
lY: label X
H: equivocator_label_zero_project (ContinueWith 0 lY) = Some lY
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (let (si', om') := transition lY (equivocator_state_zero s, om) in (equivocator_state_update s 0 si', om')) = (s', om')

transition lY (equivocator_state_zero s, om) = (equivocator_state_zero s', om')
by destruct (transition _ _ _); inversion_clear H0.
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: equivocator_label_zero_project lX = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
lX: label (equivocator_vlsm X)
H: match lX with | ContinueWith 0 li => Some li | _ => None end = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: transition lX (s, om) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
sn: state X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (equivocator_state_extend s sn, None) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_update s (S i) si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')
equivocator_state_zero s' = equivocator_state_zero s
message: Type
X: VLSM message
sn: state X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: (equivocator_state_extend s sn, None) = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by inversion H0.
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_update s (S i) si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by destruct (equivocator_state_project _ _); [destruct (transition _ _ _) |]; inversion_clear H0.
message: Type
X: VLSM message
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s 0 with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by rewrite equivocator_state_project_zero in H0; destruct (transition _ _ _); inversion_clear H0.
message: Type
X: VLSM message
i: nat
lX: label X
H: None = None
s: state (equivocator_vlsm X)
om: option message
s': state (equivocator_vlsm X)
om': option message
H0: match equivocator_state_project s (S i) with | Some si => let (si', om') := transition lX (si, om) in (equivocator_state_extend s si', om') | None => (s, om) end = (s', om')

equivocator_state_zero s' = equivocator_state_zero s
by destruct (equivocator_state_project _ _); [destruct (transition _ _ _) |]; inversion_clear H0.
message: Type
X: VLSM message
s: state (equivocator_vlsm X)
H: initial_state_prop s

initial_state_prop (equivocator_state_zero s)
apply H. Qed. End sec_equivocator_vlsm_projections.