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.
[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]
[Loading ML file coq-itauto.plugin ... done]
From stdpp Require Import prelude finite. From Coq Require Import FunctionalExtensionality. From VLSM.Lib Require Import Preamble FinSetExtras. From VLSM.Core Require Import VLSM MessageDependencies ProjectionTraces VLSMProjections. From VLSM.Core Require Import Composition SubProjectionTraces ByzantineTraces. From VLSM.Core Require Import Validator Equivocation EquivocationProjections. From VLSM.Core Require Import Equivocation.NoEquivocation Equivocation.FixedSetEquivocation.

Core: VLSM Compositions with a Fixed Set of Byzantine Components

In this module we study a composition in which a fixed subset of components are replaced by byzantine components, i.e., components which can send arbitrary channel_authenticated_messages at any time, while the rest are protocol abiding (non-equivocating) components.
We will show that, if the non-byzantine components are validators for the composition with that specific fixed-set of components as equivocators, then the projections of traces to the non-byzantine components are precisely the projections of traces of the composition with that specific fixed-set of components as equivocators to the non-equivocating components.
That is, non-byzantine validator components do not distinguish between byzantine components and equivocating ones. Therefore, when analyzing the security of a protocol it suffices to consider equivocating components.
Section sec_emit_any_signed_message_vlsm.

A machine which can emit any valid message for a given component

This VLSM is similar to the emit_any_message_vlsm with the exception of the validity predicate which requires that the sender of the emitted message corresponds to the given component.
Context
  {message : Type}
  {index : Type}
  {validator : Type}
  (A : validator -> index)
  (sender : message -> option validator)
  (component_idx : index)
  .

Program Definition emit_any_signed_message_vlsm_s0 :
  {_ : state (@emit_any_message_vlsm_type message) | True} :=
  exist _ tt _.
message, index, validator: Type
A: validator → index
sender: message → option validator
component_idx: index

(λ _ : state emit_any_message_vlsm_type, True) ()
message, index, validator: Type
A: validator → index
sender: message → option validator
component_idx: index

(λ _ : state emit_any_message_vlsm_type, True) ()
done. Defined. #[export] Instance emit_any_signed_message_vlsm_state_inh : Inhabited {_ : state (@emit_any_message_vlsm_type message) | True} := populate emit_any_signed_message_vlsm_s0.
The validity predicate allows sending only signed messages
Definition signed_messages_valid
  (l : label (@emit_any_message_vlsm_type message))
  (som : state (@emit_any_message_vlsm message) * option message)
  : Prop :=
  channel_authenticated_message A sender component_idx l.

Definition emit_any_signed_message_vlsm_machine
  : VLSMMachine emit_any_message_vlsm_type
  :=
  {| initial_state_prop := fun _ => True
   ; initial_message_prop := fun _ => False
   ; s0 := emit_any_signed_message_vlsm_state_inh
   ; transition := fun l _ => (tt, Some l)
   ; valid := signed_messages_valid
  |}.

Definition emit_any_signed_message_vlsm
    := mk_vlsm emit_any_signed_message_vlsm_machine.

End sec_emit_any_signed_message_vlsm.

Section sec_fixed_byzantine_traces.

Context
  {message : Type}
  `{FinSet index Ci}
  `{finite.Finite index}
  (IM : index -> VLSM message)
  `{forall i : index, HasBeenSentCapability (IM i)}
  (byzantine : Ci)
  {validator : Type}
  (A : validator -> index)
  (sender : message -> option validator)
  (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM)
  (can_emit_signed : channel_authentication_prop IM A sender)
  (Hsender_safety : sender_safety_alt_prop IM A sender :=
    channel_authentication_sender_safety IM A sender can_emit_signed)
  .
Given a collection of components indexed by index, we define the associated collection of fixed byzantine components by replacing the components corresponding to the indices in a given byzantine set with byzantine components, i.e., components which can emit any (signed) message.
Definition fixed_byzantine_IM : index -> VLSM message :=
  update_IM IM byzantine (fun i => emit_any_signed_message_vlsm A sender (` i)).

message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

(i : index) (m : message), ¬ initial_message_prop m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

(i : index) (m : message), ¬ initial_message_prop m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

(i : index) (m : message), ¬ initial_message_prop m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

(i : index) (m : message), ¬ initial_message_prop m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
Hm: initial_message_prop m

False
by case_decide; [| destruct (no_initial_messages_in_IM i m)]. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

channel_authentication_prop fixed_byzantine_IM A sender
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

channel_authentication_prop fixed_byzantine_IM A sender
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

channel_authentication_prop (λ i : index, match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end) A sender
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

channel_authentication_prop (λ i : index, if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i) A sender
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
Hm: can_emit (preloaded_with_all_messages_vlsm (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) m

channel_authenticated_message A sender i m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
H9: i ∈ elements byzantine
Hm: can_emit (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i)) m

channel_authenticated_message A sender i m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
H9: i ∈ elements byzantine
s0: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
om: option message
l: label (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
s1: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
Hv: valid l (s0, om)
Ht: transition l (s0, om) = (s1, Some m)

channel_authenticated_message A sender i m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
H9: i ∈ elements byzantine
s0: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
om: option message
l: label (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
s1: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
Hv: signed_messages_valid A sender i l (s0, om)
Ht: ((), Some l) = (s1, Some m)

channel_authenticated_message A sender i m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
i: index
m: message
H9: i ∈ elements byzantine
s0: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
om: option message
l: label (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
s1: state (preloaded_with_all_messages_vlsm (emit_any_signed_message_vlsm A sender i))
Hv: option_map A (sender l) = Some i
Ht: ((), Some l) = (s1, Some m)

channel_authenticated_message A sender i m
by inversion Ht; subst. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

sender_safety_alt_prop fixed_byzantine_IM A sender
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

sender_safety_alt_prop fixed_byzantine_IM A sender
exact (channel_authentication_sender_safety fixed_byzantine_IM A sender fixed_byzantine_IM_preserves_channel_authentication). Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
m: message
i: index
Hi: i ∈ byzantine

composite_label fixed_byzantine_IM
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
m: message
i: index
Hi: i ∈ byzantine

composite_label fixed_byzantine_IM
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
m: message
i: index
Hi: i ∈ byzantine

label (fixed_byzantine_IM i)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
m: message
i: index
Hi: i ∈ byzantine

label match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
m: message
i: index
Hi: i ∈ byzantine

label (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)
by rewrite decide_True; [| apply elem_of_elements, Hi]. Defined. Context (non_byzantine : Ci := difference (list_to_set (enum index)) byzantine) .
Constraint requiring only that the non-byzantine components are not equivocating.
Definition non_byzantine_not_equivocating_constraint
  : composite_label fixed_byzantine_IM ->
    composite_state fixed_byzantine_IM * option message -> Prop :=
      sub_IM_not_equivocating_constraint fixed_byzantine_IM (elements non_byzantine) A sender.

First definition of the fixed byzantine trace property

Fixed byzantine traces are projections to the subset of protocol-following components of traces which are valid for the composition in which a fixed set of components were replaced by byzantine components and the rest are protocol-following (i.e., they are not equivocating).
Definition fixed_byzantine_trace_prop
  (is : composite_state (sub_IM fixed_byzantine_IM (elements non_byzantine)))
  (tr : list (composite_transition_item (sub_IM fixed_byzantine_IM (elements non_byzantine))))
  : Prop :=
  exists bis btr,
    finite_valid_trace
      (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) bis btr /\
    composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) bis = is /\
    finite_trace_sub_projection fixed_byzantine_IM (elements non_byzantine) btr = tr.

Byzantine traces characterization as projections

Section sec_fixed_byzantine_traces_as_projections.

Definition fixed_non_byzantine_projection : VLSM message :=
  pre_induced_sub_projection fixed_byzantine_IM (elements non_byzantine)
    non_byzantine_not_equivocating_constraint.

message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

s : state fixed_non_byzantine_projection, initial_state_prop s ↔ composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

s : state fixed_non_byzantine_projection, initial_state_prop s ↔ composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection

initial_state_prop s → composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s → initial_state_prop s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection

initial_state_prop s → composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: initial_state_prop s
sub_i: sub_index (elements non_byzantine)

initial_state_prop (s sub_i)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
sX: state (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint)
HeqsX: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) sX = s
Hinitial: initial_state_prop sX
sub_i: sub_index (elements non_byzantine)

initial_state_prop (s sub_i)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sX: state (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint)
Hinitial: initial_state_prop sX
sub_i: sub_index (elements non_byzantine)

initial_state_prop (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) sX sub_i)
by apply Hinitial.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection

composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s → initial_state_prop s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s

initial_state_prop s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s

composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) = s ∧ initial_state_prop (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s

composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) = s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s
initial_state_prop (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s

composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) = s
by apply composite_state_sub_projection_lift_to.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s: state fixed_non_byzantine_projection
Hs: composite_initial_state_prop (sub_IM fixed_byzantine_IM (elements non_byzantine)) s

initial_state_prop (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
by apply (lift_sub_state_initial fixed_byzantine_IM (elements non_byzantine)). Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl fixed_non_byzantine_projection (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl fixed_non_byzantine_projection (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
strong_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
strong_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
strong_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by intros s Hincl; apply fixed_non_byzantine_projection_initial_state_preservation.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by right.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by intros l **; cbn; eapply induced_sub_projection_valid_preservation.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
s': state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om': option message

(let (s'X, om') := let (si', om') := transition (projT2 l) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (`(projT1 l)), om) in (state_update fixed_byzantine_IM (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) (`(projT1 l)) si', om') in (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s'X, om')) = (s', om') → (let (i, li) := l in let (si', om') := transition li (s i, om) in (state_update (sub_IM fixed_byzantine_IM (elements non_byzantine)) s i si', om')) = (s', om')
(* an ugly trick to get the forward direction from an iff (<->) lemma *) by eapply proj1; rapply @induced_sub_projection_transition_preservation. Qed.
The induced projection from the composition of fixed_byzantine_IM under the non_byzantine_not_equivocating_constraint to the non-byzantine components has the projection_friendly_property.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

projection_friendly_prop (induced_sub_projection_is_projection fixed_byzantine_IM (elements non_byzantine) non_byzantine_not_equivocating_constraint)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

projection_friendly_prop (induced_sub_projection_is_projection fixed_byzantine_IM (elements non_byzantine) non_byzantine_not_equivocating_constraint)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding (pre_induced_sub_projection fixed_byzantine_IM (elements non_byzantine) non_byzantine_not_equivocating_constraint) (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

s1 s2 : composite_state fixed_byzantine_IM, composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2 → (l : composite_label fixed_byzantine_IM) (om : option message), non_byzantine_not_equivocating_constraint l (s1, om) → non_byzantine_not_equivocating_constraint l (s2, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
om: option message

non_byzantine_not_equivocating_constraint l (s1, om) → non_byzantine_not_equivocating_constraint l (s2, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message

non_byzantine_not_equivocating_constraint l (s1, Some m) → non_byzantine_not_equivocating_constraint l (s2, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message

match option_map A (sender m) with | Some i => match decide (i ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i non_byzantine_i)) (s1 i) m | right _ => True end | None => True endmatch option_map A (sender m) with | Some i => match decide (i ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i non_byzantine_i)) (s2 i) m | right _ => True end | None => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message
v: validator

match option_map A (Some v) with | Some i => match decide (i ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i non_byzantine_i)) (s1 i) m | right _ => True end | None => True endmatch option_map A (Some v) with | Some i => match decide (i ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i non_byzantine_i)) (s2 i) m | right _ => True end | None => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message
v: validator

match decide (A v ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) non_byzantine_i)) (s1 (A v)) m | right _ => True endmatch decide (A v ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) non_byzantine_i)) (s2 (A v)) m | right _ => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message
v: validator
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (s1 (A v)) m → has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (s2 (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1, s2: composite_state fixed_byzantine_IM
Heq: composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s1 = composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) s2
l: composite_label fixed_byzantine_IM
m: message
v: validator
HAv: A v ∈ elements non_byzantine

s1 (A v) = s2 (A v)
exact (equal_f_dep Heq (dexist (A v) HAv)). Qed.
Characterization result for the first definition: the fixed_byzantine_trace_property is equivalent to the finite_valid_trace_property of the pre_induced_sub_projection of the the composition in which a fixed set of components were replaced by byzantine components and the rest are protocol-following to the set of protocol-following components.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
is: composite_state (sub_IM fixed_byzantine_IM (elements non_byzantine))
tr: list (composite_transition_item (sub_IM fixed_byzantine_IM (elements non_byzantine)))

fixed_byzantine_trace_prop is tr ↔ finite_valid_trace fixed_non_byzantine_projection is tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
is: composite_state (sub_IM fixed_byzantine_IM (elements non_byzantine))
tr: list (composite_transition_item (sub_IM fixed_byzantine_IM (elements non_byzantine)))

fixed_byzantine_trace_prop is tr ↔ finite_valid_trace fixed_non_byzantine_projection is tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
is: composite_state (sub_IM fixed_byzantine_IM (elements non_byzantine))
tr: list (composite_transition_item (sub_IM fixed_byzantine_IM (elements non_byzantine)))

(finite_valid_trace fixed_non_byzantine_projection is tr → fixed_byzantine_trace_prop is tr) ∧ (fixed_byzantine_trace_prop is tr → finite_valid_trace fixed_non_byzantine_projection is tr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
is: composite_state (sub_IM fixed_byzantine_IM (elements non_byzantine))
tr: list (composite_transition_item (sub_IM fixed_byzantine_IM (elements non_byzantine)))

projection_friendly_prop (induced_sub_projection_is_projection fixed_byzantine_IM (elements non_byzantine) non_byzantine_not_equivocating_constraint)
by apply fixed_non_byzantine_projection_friendliness. Qed. End sec_fixed_byzantine_traces_as_projections.

Fixed Byzantine traces as traces preloaded with signed messages

In this section we show that byzantine traces correspond to traces of a composition in the non_byzantine set, preloaded with messages signed by the components in the byzantine set.
Section sec_fixed_byzantine_traces_as_preloaded.

Definition fixed_set_signed_message (m : message) : Prop :=
  non_sub_index_authenticated_message (elements non_byzantine) A sender m /\
  (exists i, i ∈ non_byzantine /\
    exists l s, input_constrained (IM i) l (s, Some m)).
Given that we're only looking at the composition indexed by the non_byzantine set, we can define their subset as either a subset of the fixed_byzantine_IM or of the original IM.
As it turns out, the first definition is easier to prove equivalent to the induced fixed_non_byzantine_projection, while the second is easier to work with in general because it makes no reference to byzantine components.
Therefore we'll first be defining them both below and show that they are equivalent to each-other using the generic Lemma same_IM_embedding.
Definition preloaded_fixed_non_byzantine_vlsm' : VLSM message :=
  composite_no_equivocation_vlsm_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine))
    (free_constraint _) fixed_set_signed_message.

Definition preloaded_fixed_non_byzantine_vlsm : VLSM message :=
  composite_no_equivocation_vlsm_with_preloaded (sub_IM IM (elements non_byzantine)) (free_constraint _)
    fixed_set_signed_message.

message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

sub_i : sub_index (elements non_byzantine), sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i = sub_IM IM (elements non_byzantine) sub_i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

sub_i : sub_index (elements non_byzantine), sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i = sub_IM IM (elements non_byzantine) sub_i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sub_i: sub_index (elements non_byzantine)

sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i = sub_IM IM (elements non_byzantine) sub_i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sub_i: sub_index (elements non_byzantine)
i: index
Hi: sub_index_prop (elements non_byzantine) i
Heqsub_i: sub_i = dexist i Hi

sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i = sub_IM IM (elements non_byzantine) sub_i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i

sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi) = sub_IM IM (elements non_byzantine) (dexist i Hi)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i

match decide (`(dexist i Hi) ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist (`(dexist i Hi)) i_in)) | right _ => IM (`(dexist i Hi)) end = IM (`(dexist i Hi))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i

(if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i) = IM i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: i ∉ byzantine

(if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i) = IM i
by rewrite decide_False; [| rewrite elem_of_elements]. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

sub_i : sub_index (elements non_byzantine), sub_IM IM (elements non_byzantine) sub_i = sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

sub_i : sub_index (elements non_byzantine), sub_IM IM (elements non_byzantine) sub_i = sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i
by intro; symmetry; apply non_byzantine_components_same. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm' preloaded_fixed_non_byzantine_vlsm (same_IM_label_rew non_byzantine_components_same) (same_IM_state_rew non_byzantine_components_same)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm' preloaded_fixed_non_byzantine_vlsm (same_IM_label_rew non_byzantine_components_same) (same_IM_state_rew non_byzantine_components_same)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

s1 : state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)))), constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s1 → (l1 : composite_label (sub_IM fixed_byzantine_IM (elements non_byzantine))) (om : option message), no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message l1 (s1, om) → no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements non_byzantine)) (free_constraint (sub_IM IM (elements non_byzantine))) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same l1) (same_IM_state_rew non_byzantine_components_same s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s1
l1: composite_label (sub_IM fixed_byzantine_IM (elements non_byzantine))
om: option message
Hom: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, om)

no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements non_byzantine)) (free_constraint (sub_IM IM (elements non_byzantine))) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same l1) (same_IM_state_rew non_byzantine_components_same s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s1
l1: composite_label (sub_IM fixed_byzantine_IM (elements non_byzantine))
om: option message
Hom: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, om)

composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same l1) (same_IM_state_rew non_byzantine_components_same s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s1
l1: composite_label (sub_IM fixed_byzantine_IM (elements non_byzantine))
m: message
Hom: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, Some m)

composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same l1) (same_IM_state_rew non_byzantine_components_same s1, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s1
l1: composite_label (sub_IM fixed_byzantine_IM (elements non_byzantine))
m: message
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) ( s1, Some m).1 m

composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same l1) (same_IM_state_rew non_byzantine_components_same s1, Some m)
by left; eapply same_IM_composite_has_been_sent_preservation. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm preloaded_fixed_non_byzantine_vlsm' (same_IM_label_rew non_byzantine_components_same_sym) (same_IM_state_rew non_byzantine_components_same_sym)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm preloaded_fixed_non_byzantine_vlsm' (same_IM_label_rew non_byzantine_components_same_sym) (same_IM_state_rew non_byzantine_components_same_sym)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

s1 : state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM (elements non_byzantine)))), constrained_state_prop (free_composite_vlsm (sub_IM IM (elements non_byzantine))) s1 → (l1 : composite_label (sub_IM IM (elements non_byzantine))) (om : option message), no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements non_byzantine)) (free_constraint (sub_IM IM (elements non_byzantine))) fixed_set_signed_message l1 (s1, om) → no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same_sym l1) (same_IM_state_rew non_byzantine_components_same_sym s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM IM (elements non_byzantine))) s1
l1: composite_label (sub_IM IM (elements non_byzantine))
om: option message
Hom: composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, om)

no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same_sym l1) (same_IM_state_rew non_byzantine_components_same_sym s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM IM (elements non_byzantine))) s1
l1: composite_label (sub_IM IM (elements non_byzantine))
om: option message
Hom: composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, om)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same_sym l1) (same_IM_state_rew non_byzantine_components_same_sym s1, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM IM (elements non_byzantine))) s1
l1: composite_label (sub_IM IM (elements non_byzantine))
m: message
Hom: composite_no_equivocations_except_from (sub_IM IM (elements non_byzantine)) fixed_set_signed_message l1 ( s1, Some m)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same_sym l1) (same_IM_state_rew non_byzantine_components_same_sym s1, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
s1: state (preloaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM (elements non_byzantine))))
Hs1: constrained_state_prop (free_composite_vlsm (sub_IM IM (elements non_byzantine))) s1
l1: composite_label (sub_IM IM (elements non_byzantine))
m: message
Hsent: composite_has_been_sent (sub_IM IM (elements non_byzantine)) (s1, Some m).1 m

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (same_IM_label_rew non_byzantine_components_same_sym l1) (same_IM_state_rew non_byzantine_components_same_sym s1, Some m)
by left; eapply same_IM_composite_has_been_sent_preservation. Qed.
We next show that the fixed_non_byzantine_projection and preloaded_fixed_non_byzantine_vlsm' are VLSM_equal.
The validity of fixed_non_byzantine_projection subsumes the constraint of preloaded_fixed_non_byzantine_vlsm'.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

(l : label fixed_non_byzantine_projection) (s : state fixed_non_byzantine_projection) (om : option message), valid l (s, om) → composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

(l : label fixed_non_byzantine_projection) (s : state fixed_non_byzantine_projection) (om : option message), valid l (s, om) → composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
om: option message
Hv: valid l (s, om)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
om: option message
Hv: valid l (s, om)
Hnoequiv: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) (non_sub_index_authenticated_message (elements non_byzantine) A sender) l (s, om)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
Hv: valid l (s, Some m)
Hnoequiv: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) (non_sub_index_authenticated_message (elements non_byzantine) A sender) l (s, Some m)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
Hv: valid l (s, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
Hv: valid l (s, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

fixed_set_signed_message m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
Hv: valid l (s, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
li: label (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
si: state (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
Hv: input_constrained (fixed_byzantine_IM i) li (si, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
li: label (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
si: state (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
Hv: input_constrained (fixed_byzantine_IM i) li (si, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
li: label (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
si: state (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))
Hv: input_constrained (fixed_byzantine_IM i) li (si, Some m)
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

(l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

(li : label (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))) (si : state (preloaded_with_all_messages_vlsm (fixed_byzantine_IM i))), input_constrained (fixed_byzantine_IM i) li (si, Some m) → (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

(li : label (preloaded_with_all_messages_vlsm match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end)) (si : state (preloaded_with_all_messages_vlsm match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end)), input_constrained match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end li (si, Some m) → (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∈ elements non_byzantine
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

(li : label (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) (si : state (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)), valid_state_prop (preloaded_with_all_messages_vlsm (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) si ∧ option_valid_message_prop (preloaded_with_all_messages_vlsm (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) (Some m) ∧ valid li (si, Some m) → (l : label (IM i)) (s : state (IM i)), valid_state_prop (preloaded_with_all_messages_vlsm (IM i)) s ∧ option_valid_message_prop (preloaded_with_all_messages_vlsm (IM i)) (Some m) ∧ valid l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label fixed_non_byzantine_projection
s: state fixed_non_byzantine_projection
m: message
i: index
Hi: i ∉ byzantine
Hseeded: non_sub_index_authenticated_message (elements non_byzantine) A sender m

(li : label (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) (si : state (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)), valid_state_prop (preloaded_with_all_messages_vlsm (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) si ∧ option_valid_message_prop (preloaded_with_all_messages_vlsm (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)) (Some m) ∧ valid li (si, Some m) → (l : label (IM i)) (s : state (IM i)), valid_state_prop (preloaded_with_all_messages_vlsm (IM i)) s ∧ option_valid_message_prop (preloaded_with_all_messages_vlsm (IM i)) (Some m) ∧ valid l (s, Some m)
by rewrite decide_False; [intros; exists li, si | rewrite elem_of_elements]. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl fixed_non_byzantine_projection preloaded_fixed_non_byzantine_vlsm'
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl fixed_non_byzantine_projection preloaded_fixed_non_byzantine_vlsm'
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
by intro; intros; apply fixed_non_byzantine_projection_initial_state_preservation.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} (id s)

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) ( s, Some m).1 m
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} (id s)

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hseeded: fixed_set_signed_message m
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} (id s)
valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) ( s, Some m).1 m
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} (id s)

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} s

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
by eapply preloaded_composite_sent_valid.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} s
Hseeded: fixed_set_signed_message m
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} (id s)

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} m
by apply initial_message_is_valid; right.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)

valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)

valid (id l) (id s, om)
by rapply @induced_sub_projection_valid_preservation.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)

no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)

composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (id l) (id s, om)
by apply fixed_non_byzantine_projection_valid_no_equivocations.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om: option message
s': state {| vlsm_type := projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |}
om': option message
Ht: transition l (s, om) = (s', om')

(let (i, li) := l in let (si', om') := transition li (s i, om) in (state_update (sub_IM fixed_byzantine_IM (elements non_byzantine)) s i si', om')) = (s', om')
by apply induced_sub_projection_transition_preservation. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_valid_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_valid_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sub_i: sub_index (elements non_byzantine)
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) sub_i)
s: state preloaded_fixed_non_byzantine_vlsm'
om: option message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
HomX: option_valid_message_prop preloaded_fixed_non_byzantine_vlsm' om
Hv: valid (existT sub_i li) (s, om)
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT sub_i li) (s, om)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
HomY: option_valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) om

valid (lift_sub_label fixed_byzantine_IM (elements non_byzantine) (existT sub_i li)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
om: option message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
HomX: option_valid_message_prop preloaded_fixed_non_byzantine_vlsm' om
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, om)
Hv: valid (existT (dexist i Hi) li) (s, om)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
HomY: option_valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) om

valid (lift_sub_label fixed_byzantine_IM (elements non_byzantine) (existT (dexist i Hi) li)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
om: option message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
HomX: option_valid_message_prop preloaded_fixed_non_byzantine_vlsm' om
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, om)
Hv: valid (existT (dexist i Hi) li) (s, om)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
HomY: option_valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) om

non_byzantine_not_equivocating_constraint (lift_sub_label fixed_byzantine_IM (elements non_byzantine) (existT (dexist i Hi) li)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
om: option message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, om)

non_byzantine_not_equivocating_constraint (lift_sub_label fixed_byzantine_IM (elements non_byzantine) (existT (dexist i Hi) li)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, Some m)

match option_map A (sender m) with | Some i => match decide (i ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i non_byzantine_i)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s i) m | right _ => True end | None => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, Some m)
v: validator
Hsender: sender m = Some v

match decide (A v ∈ elements non_byzantine) with | left non_byzantine_i => has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) non_byzantine_i)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m | right _ => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hc: composite_no_equivocations_except_from (sub_IM fixed_byzantine_IM (elements non_byzantine)) fixed_set_signed_message (existT (dexist i Hi) li) ( s, Some m)
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hseeded: fixed_set_signed_message m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine
has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state_to fixed_byzantine_IM (elements non_byzantine) (λ n : index, `(vs0 (fixed_byzantine_IM n))) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (s (dexist (A v) HAv)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

constrained_state_prop (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hsent: composite_has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine)) s m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

VLSM_incl_part (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (free_composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by eapply constrained_preloaded_vlsm_incl_preloaded_with_all_messages.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
i: index
Hi: sub_index_prop (elements non_byzantine) i
li: label (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist i Hi))
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
HsX: valid_state_prop preloaded_fixed_non_byzantine_vlsm' s
Hseeded: fixed_set_signed_message m
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements non_byzantine

has_been_sent (sub_IM fixed_byzantine_IM (elements non_byzantine) (dexist (A v) HAv)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
m: message
Hseeded: fixed_set_signed_message m
v: validator
Hsender: sender m = Some v

A v ∉ elements non_byzantine
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
m: message
i: index
Hi: i ∉ elements non_byzantine
Hm: channel_authenticated_message A sender i m
v: validator
Hsender: sender m = Some v

A v ∉ elements non_byzantine
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
m: message
i: index
Hi: i ∉ elements non_byzantine
Hm: option_map A (sender m) = Some i
v: validator
Hsender: sender m = Some v

A v ∉ elements non_byzantine
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
m: message
i: index
Hi: i ∉ elements non_byzantine
v: validator
Hm: option_map A (Some v) = Some i
Hsender: sender m = Some v

A v ∉ elements non_byzantine
by inversion Hm; subst. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_initial_message_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_initial_message_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
HmX: initial_message_prop m

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
sub_i: sub_index (elements non_byzantine)
im: message
Him: initial_message_prop im
Heqm: `(im ↾ Him) = m

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
Hseeded: fixed_set_signed_message m
valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
sub_i: sub_index (elements non_byzantine)
im: message
Him: initial_message_prop im
Heqm: `(im ↾ Him) = m

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
sub_i: sub_index (elements non_byzantine)
im: message
Him: initial_message_prop im
Heqm: `(im ↾ Him) = m

False
message, index, Ci: Type
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sub_i: sub_index (elements non_byzantine)
im: message
Him: initial_message_prop im

False
message, index, Ci: Type
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
sub_i: sub_index (elements non_byzantine)
im: message
Him: initial_message_prop im
i: index
Hi: sub_index_prop (elements non_byzantine) i
Heqsub_i: sub_i = dexist i Hi

False
message, index, Ci: Type
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
im: message
i: index
Hi: sub_index_prop (elements non_byzantine) i
Him: initial_message_prop im

False
message, index, Ci: Type
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
im: message
i: index
Hi: sub_index_prop (elements non_byzantine) i
Him: initial_message_prop im

False
message, index, Ci: Type
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
im: message
i: index
Hi: sub_index_prop (elements non_byzantine) i
Him: initial_message_prop im

False
by case_decide; [| destruct (no_initial_messages_in_IM i im)].
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
Hseeded: fixed_set_signed_message m

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM

valid_state_message_prop X s0 None
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM

valid_state_message_prop X s0 None
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM

initial_state_prop s0
by destruct (composite_s0 fixed_byzantine_IM).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements non_byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: l : label X, valid l (s0, None) → (s' : state X) (om' : option message), transition l (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ byzantine)
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: l : label X, valid l (s0, None) → (s' : state X) (om' : option message), transition l (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: ¬ (i ∈ list_to_set (enum index) ∧ i ∉ byzantine)
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: l : label X, valid l (s0, None) → (s' : state X) (om' : option message), transition l (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: ¬ (i ∉ byzantine)
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: l : label X, valid l (s0, None) → (s' : state X) (om' : option message), transition l (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: l : label X, valid l (s0, None) → (s' : state X) (om' : option message), transition l (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (message_as_byzantine_label m i Hi) (s0, None)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'
valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (message_as_byzantine_label m i Hi) (s0, None)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (message_as_byzantine_label m i Hi) (s0, None)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end), None)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: valid (message_as_byzantine_label m i Hi) (s0, None) → (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)), None)
by rewrite @decide_True.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: (s' : state X) (om' : option message), transition (message_as_byzantine_label m i Hi) (s0, None) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
Hgen: (s' : composite_state fixed_byzantine_IM) (om' : option message), (let (si', om'0) := transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) in (state_update fixed_byzantine_IM (λ n : index, `(vs0 (fixed_byzantine_IM n))) i si', om'0)) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
si': state (fixed_byzantine_IM i)
_om: option message
Ht: transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) = (si', _om)
Hgen: (s' : composite_state fixed_byzantine_IM) (om' : option message), (state_update fixed_byzantine_IM (λ n : index, `(vs0 (fixed_byzantine_IM n))) i si', _om) = (s', om') → valid_state_message_prop X s' om'

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
si': state (fixed_byzantine_IM i)
_om: option message
Ht: transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) = (si', _om)
Hgen: valid_state_message_prop X (state_update fixed_byzantine_IM (λ n : index, `(vs0 (fixed_byzantine_IM n))) i si') _om

valid_message_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label preloaded_fixed_non_byzantine_vlsm'
s: state preloaded_fixed_non_byzantine_vlsm'
m: message
Hv: input_valid preloaded_fixed_non_byzantine_vlsm' l (s, Some m)
HsY: valid_state_prop (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
i: index
Hi: i ∈ byzantine
Hsender: channel_authenticated_message A sender i m
Hvalid: i : index, i ∈ non_byzantine ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
X:= composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint: VLSM message
s0:= `(composite_s0 fixed_byzantine_IM): composite_state fixed_byzantine_IM
Hs0: valid_state_message_prop X s0 None
si': state (fixed_byzantine_IM i)
_om: option message
Ht: transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) = (si', _om)
Hgen: valid_state_message_prop X (state_update fixed_byzantine_IM (λ n : index, `(vs0 (fixed_byzantine_IM n))) i si') _om

Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
si': state (fixed_byzantine_IM i)
_om: option message
Ht: transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) = (si', _om)

Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
_om: option message

si' : state (fixed_byzantine_IM i), transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (fixed_byzantine_IM i)), None) = (si', _om) → Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
_om: option message

si' : state match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end, transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 match decide (i ∈ elements byzantine) with | left i_in => emit_any_signed_message_vlsm A sender (`(dexist i i_in)) | right _ => IM i end), None) = (si', _om) → Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
_om: option message

si' : state (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i), transition (eq_rect_r (λ v : VLSM message, label v) m (decide_True (emit_any_signed_message_vlsm A sender i) (IM i) (match elem_of_elements byzantine i with | conj _ x0 => x0 end Hi))) (`(vs0 (if decide (i ∈ elements byzantine) then emit_any_signed_message_vlsm A sender i else IM i)), None) = (si', _om) → Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
_om: option message

si' : state (emit_any_signed_message_vlsm A sender i), transition (eq_rect_r (λ v : VLSM message, label v) m eq_refl) (`(vs0 (emit_any_signed_message_vlsm A sender i)), None) = (si', _om) → Some m = _om
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
m: message
i: index
Hi: i ∈ byzantine
_om: option message

si' : (), ((), Some m) = (si', _om) → Some m = _om
by inversion 1; itauto. Qed.
Since the fixed_non_byzantine_projection is an induced_validator of the composition of fixed_byzantine_IM with a non_byzantine_not_equivocating_constraint, its initial_messages and validity are derived from valid messages and protocol validity of the larger composition; therefore, the following simple result becomes very important.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_embedding preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_valid_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_embedding_transition_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
strong_projection_initial_state_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_embedding_initial_message_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_valid_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
by intro; intros; apply preloaded_fixed_non_byzantine_vlsm_lift_valid.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_transition_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
by intro; intros * []; rapply lift_sub_transition.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_projection_initial_state_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
by intro; intros; apply (lift_sub_state_initial fixed_byzantine_IM).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_embedding_initial_message_preservation preloaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))
by intro; intros; apply (preloaded_fixed_non_byzantine_vlsm_lift_initial_message l s m). Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl preloaded_fixed_non_byzantine_vlsm' fixed_non_byzantine_projection
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl preloaded_fixed_non_byzantine_vlsm' fixed_non_byzantine_projection
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_initial_message_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_valid_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
weak_incl_transition_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

strong_incl_initial_state_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by intro; intros; apply fixed_non_byzantine_projection_initial_state_preservation.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_initial_message_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by intros l s m Hv _ Him; apply initial_message_is_valid; cbn; right.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_valid_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)

valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s) → option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om → valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

sX : state (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint), InputValidation (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (id l) (id s, om).1 (id s, om).2 (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) sX
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

InputValidation (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (id l) (id s, om).1 (id s, om).2 (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine) (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) = Some (id l)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om
composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) = ( id s, om).1
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om
input_valid (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, ( id s, om).2)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine) (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) = Some (id l)
by apply composite_label_sub_projection_option_lift.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s) = (id s, om).1
by apply composite_state_sub_projection_lift.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
l: label {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
s: state {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |}
om: option message
Hv: input_valid {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True) |} om

input_valid (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (lift_sub_label fixed_byzantine_IM (elements non_byzantine) l) (lift_sub_state fixed_byzantine_IM (elements non_byzantine) s, (id s, om).2)
by apply (VLSM_embedding_input_valid preloaded_fixed_non_byzantine_vlsm_lift).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

weak_incl_transition_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM fixed_byzantine_IM (elements non_byzantine)) (no_equivocations_additional_constraint_with_preloaded (sub_IM fixed_byzantine_IM (elements non_byzantine)) (free_constraint (sub_IM fixed_byzantine_IM (elements non_byzantine))) fixed_set_signed_message)) fixed_set_signed_message) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) (composite_type (sub_IM fixed_byzantine_IM (elements non_byzantine))) (composite_label_sub_projection_option fixed_byzantine_IM (elements non_byzantine)) (composite_state_sub_projection fixed_byzantine_IM (elements non_byzantine)) (lift_sub_label fixed_byzantine_IM (elements non_byzantine)) (lift_sub_state fixed_byzantine_IM (elements non_byzantine))) (λ _ : message, True))
by intros l s om s' om' [_ Ht]; apply induced_sub_projection_transition_preservation. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_eq fixed_non_byzantine_projection preloaded_fixed_non_byzantine_vlsm'
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_eq fixed_non_byzantine_projection preloaded_fixed_non_byzantine_vlsm'
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := fixed_non_byzantine_projection |} {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := preloaded_fixed_non_byzantine_vlsm' |}
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci
VLSM_incl {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := preloaded_fixed_non_byzantine_vlsm' |} {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := fixed_non_byzantine_projection |}
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := fixed_non_byzantine_projection |} {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := preloaded_fixed_non_byzantine_vlsm' |}
by apply fixed_non_byzantine_preloaded_incl.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
byzantine: Ci
validator: Type
A: validator → index
sender: message → option validator
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
non_byzantine:= list_to_set (enum index) ∖ byzantine: Ci

VLSM_incl {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := preloaded_fixed_non_byzantine_vlsm' |} {| vlsm_type := fixed_non_byzantine_projection; vlsm_machine := fixed_non_byzantine_projection |}
by apply preloaded_fixed_non_byzantine_incl. Qed. End sec_fixed_byzantine_traces_as_preloaded.

Second fixed byzantine trace definition

Given the equivalence results from Lemmas fixed_byzantine_trace_char1, fixed_non_byzantine_preloaded_eq, and preloaded_fixed_non_byzantine_VLSM_embedding, we introduce the following alternate definition to the fixed_byzantine_trace_property, this time defined for (not necessarily valid) traces of the full composition, accepting any such trace as long as its projection to the non-byzantine components is indistinguishable from a projection of a valid trace of the composition of fixed_byzantine_IM under the non_byzantine_not_equivocating_constraint.
Definition fixed_byzantine_trace_alt_prop is tr : Prop :=
  finite_valid_trace preloaded_fixed_non_byzantine_vlsm
    (composite_state_sub_projection IM (elements non_byzantine) is)
    (finite_trace_sub_projection IM (elements non_byzantine) tr).

End sec_fixed_byzantine_traces.

Relation between Byzantine components and equivocators

In this section we show that while equivocators can always appear as byzantine to the protocol-abiding components, the converse is also true if the protocol- abiding components satisfy the projection_message_validator_property, which basically allows them to locally verify the authenticity of a received message.
Section sec_fixed_non_equivocating_vs_byzantine.

Context
  {message : Type}
  `{FinSet index Ci}
  `{finite.Finite index}
  (IM : index -> VLSM message)
  `{forall i : index, HasBeenSentCapability (IM i)}
  `{forall i : index, HasBeenReceivedCapability (IM i)}
  (selection : Ci)
  {validator : Type}
  (A : validator -> index)
  (sender : message -> option validator)
  (selection_complement := difference (list_to_set (enum index)) selection)
  (PreNonByzantine : VLSM message := preloaded_fixed_non_byzantine_vlsm IM selection A sender)
  (Fixed : VLSM message := fixed_equivocation_vlsm_composition IM selection)
  (FixedNonEquivocating : VLSM message :=
    pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection))
  (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM)
  (can_emit_signed : channel_authentication_prop IM A sender)
  (Hsender_safety : sender_safety_alt_prop IM A sender :=
    channel_authentication_sender_safety IM A sender can_emit_signed)
  .

message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

VLSM_incl FixedNonEquivocating (pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

VLSM_incl FixedNonEquivocating (pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

input_valid_constraint_subsumption (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
om: option message
Hv: input_valid (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) l (s, om)

sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
om: option message
Hv: input_valid (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) l (s, om)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, om)

sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
om: option message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) om
Hv: valid l (s, om)
Hc: fixed_equivocation_constraint IM selection l (s, om)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, om)

sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender l (s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, Some m)

sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender l (s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, Some m)
v: validator
Hsender: sender m = Some v

match decide (A v ∈ elements selection_complement) with | left non_byzantine_i => has_been_sent (sub_IM IM (elements selection_complement) (dexist (A v) non_byzantine_i)) (s (A v)) m | right _ => True end
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, Some m)
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements selection_complement

has_been_sent (sub_IM IM (elements selection_complement) (dexist (A v) HAv)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, Some m)
v: validator
Hsender: sender m = Some v
HAv: A v ∈ elements selection_complement

has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hstrong_v: strong_fixed_equivocation_constraint IM selection l (s, Some m)
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
i: index
Hi: i ∉ elements selection
Hsent: has_been_sent (IM i) (s i) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hemitted: can_emit (equivocators_composition_for_sent IM selection s) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection
has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
i: index
Hi: i ∉ elements selection
Hsent: has_been_sent (IM i) (s i) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
is: state (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection))
tr: list transition_item
Htr: finite_valid_trace_init_to (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) is s tr
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
i: index
Hi: i ∉ elements selection
Hsent: has_been_sent (IM i) (s i) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

has_been_sent (IM (A v)) (s (A v)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
is: state (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection))
tr: list transition_item
Htr: finite_valid_trace_init_to (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) is s tr
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
i: index
Hi: i ∉ elements selection
Hsent: has_been_sent (IM i) (s i) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

finite_constrained_trace_init_to (free_composite_vlsm IM) ?is s ?tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
is: state (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection))
tr: list transition_item
Htr: finite_valid_trace_init_to (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) is s tr
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
i: index
Hi: i ∉ elements selection
Hsent: has_been_sent (IM i) (s i) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

VLSM_incl_part (constrained_vlsm_machine (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (preloaded_vlsm_machine (free_composite_vlsm IM) (λ _ : message, True))
by apply constrained_preloaded_incl.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label (free_composite_vlsm IM)
s: state (free_composite_vlsm IM)
m: message
Hs: valid_state_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) s
Hom: option_valid_message_prop (constrained_vlsm (free_composite_vlsm IM) (fixed_equivocation_constraint IM selection)) (Some m)
Hv: valid l (s, Some m)
Hc: fixed_equivocation_constraint IM selection l (s, Some m)
Hemitted: can_emit (equivocators_composition_for_sent IM selection s) m
v: validator
Hsender: sender m = Some v
HAv: A v ∉ selection

has_been_sent (IM (A v)) (s (A v)) m
by contradict HAv; apply elem_of_elements; eapply sub_can_emit_sender. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

VLSM_incl FixedNonEquivocating PreNonByzantine
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

VLSM_incl FixedNonEquivocating PreNonByzantine
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
weak_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
weak_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
weak_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

strong_incl_initial_state_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
sX: state (composite_vlsm IM (fixed_equivocation_constraint IM selection))
Hinitial: initial_state_prop sX
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))

initial_state_prop (id (composite_state_sub_projection IM (elements selection_complement) sX) sub_i)
by apply Hinitial.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

weak_incl_initial_message_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hv: input_valid {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hnoeqv: composite_no_equivocations_except_from (sub_IM IM (elements selection_complement)) (non_sub_index_authenticated_message (elements selection_complement) A sender) l (s, Some m)

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hsent: composite_has_been_sent (sub_IM IM (elements selection_complement)) (s, Some m).1 m

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m
valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hsent: composite_has_been_sent (sub_IM IM (elements selection_complement)) (s, Some m).1 m

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
by eapply preloaded_composite_sent_valid.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

initial_message_prop m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

(let (initial_state_prop, initial_state, _, initial_message_prop, initial_message, _, _) := vlsm_machine {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} in initial_message_prop) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

i : index, i ∈ list_to_set (enum index) ∖ selection ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hs: valid_state_prop {| vlsm_type := FixedNonEquivocating; vlsm_machine := pre_induced_sub_projection IM (elements selection_complement) (sub_IM_not_equivocating_constraint IM (elements selection_complement) A sender) |} s
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HmX: initial_message_prop m
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m
i: index
Hi: i ∈ elements selection_complement
Hiv: (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m)

i : index, i ∈ list_to_set (enum index) ∖ selection ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
by apply elem_of_elements in Hi; eexists; split.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

weak_incl_valid_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: input_valid {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} l (s, om)

valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s) → option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om → valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)

valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s) → option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om → valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om

valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om
no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender) (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om

valid (id l) (id s, om)
by eapply induced_sub_projection_valid_preservation in Hv.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om

no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender) (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om

composite_no_equivocations_except_from (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (fixed_set_signed_message IM selection A sender) (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
Hv: valid l (s, om)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} om
Hnoequiv: composite_no_equivocations_except_from (sub_IM IM (elements selection_complement)) (non_sub_index_authenticated_message (elements selection_complement) A sender) l (s, om)

composite_no_equivocations_except_from (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (fixed_set_signed_message IM selection A sender) (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (Some m)
Hnoequiv: composite_no_equivocations_except_from (sub_IM IM (elements selection_complement)) (non_sub_index_authenticated_message (elements selection_complement) A sender) l (s, Some m)

composite_no_equivocations_except_from (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (fixed_set_signed_message IM selection A sender) (id l) (id s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (Some m)
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

fixed_set_signed_message IM selection A sender m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (Some m)
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m

i : index, i ∈ list_to_set (enum index) ∖ selection ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
m: message
Hv: valid l (s, Some m)
HsY: valid_state_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (Some m)
Hseeded: non_sub_index_authenticated_message (elements selection_complement) A sender m
i: index
Hi: i ∈ elements selection_complement
Hiv: (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m)

i : index, i ∈ list_to_set (enum index) ∖ selection ∧ ( (l : label (preloaded_with_all_messages_vlsm (IM i))) (s : state (preloaded_with_all_messages_vlsm (IM i))), input_constrained (IM i) l (s, Some m))
by apply elem_of_elements in Hi; eexists; split.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender

weak_incl_transition_preservation (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True)) (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
s': state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om': option message
Ht: transition l (s, om) = (s', om')

transition (id l) (id s, om) = (id s', om')
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
l: label {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
s: state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om: option message
s': state {| vlsm_type := projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |}
om': option message

transition l (s, om) = (s', om') → transition (id l) (id s, om) = (id s', om')
by apply @induced_sub_projection_transition_preservation. Qed.
As a corollary to the above result, we can conclude that valid traces for the composition with Fixed equivocation have the fixed_byzantine_trace_alt_property.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
is: state Fixed
tr: list transition_item

finite_valid_trace Fixed is tr → fixed_byzantine_trace_alt_prop IM selection A sender is tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
is: state Fixed
tr: list transition_item

finite_valid_trace Fixed is tr → fixed_byzantine_trace_alt_prop IM selection A sender is tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
is: state Fixed
tr: list transition_item
Htr: finite_valid_trace Fixed is tr

fixed_byzantine_trace_alt_prop IM selection A sender is tr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
is: state Fixed
tr: list transition_item
Htr: finite_valid_trace Fixed is tr

finite_valid_trace {| vlsm_type := FixedNonEquivocating; vlsm_machine := FixedNonEquivocating |} (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) is) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) tr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
is: state Fixed
tr: list transition_item
Htr: finite_valid_trace Fixed is tr
Hproj: VLSM_projection (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection)) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement))

finite_valid_trace {| vlsm_type := FixedNonEquivocating; vlsm_machine := FixedNonEquivocating |} (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) is) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) tr)
by apply (VLSM_projection_finite_valid_trace Hproj). Qed. Section sec_validator_fixed_set_byzantine. Context `{FinSet message Cm} (message_dependencies : message -> Cm) `{!Irreflexive (msg_dep_happens_before message_dependencies)} `{forall i, MessageDependencies (IM i) message_dependencies} (Hfull : forall i, message_dependencies_full_node_condition_prop (IM i) message_dependencies) .
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

weak_embedding_valid_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

weak_embedding_valid_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: input_valid PreNonByzantine l (s, om)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed om

valid (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: input_valid PreNonByzantine l (s, om)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed om

valid (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: input_valid PreNonByzantine l (s, om)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed om
fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: input_valid PreNonByzantine l (s, om)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed om

valid (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, om)
by rapply @lift_sub_valid; apply Hv.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: input_valid PreNonByzantine l (s, om)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed om

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hc: option_valid_message_prop PreNonByzantine (Some m) ∧ valid l (s, Some m)

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hc: composite_no_equivocations_except_from (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (fixed_set_signed_message IM selection A sender) l (s, Some m)

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (s, Some m).1 m

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hseeded: fixed_set_signed_message IM selection A sender m
fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (s, Some m).1 m

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (s, Some m).1 m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (s, Some m).1 m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m ∨ composite_has_been_received IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (s, Some m).1 m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) s m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
Hsent: composite_has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) s m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
Hsenti: has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) sub_i) (s sub_i) m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
Hsenti: has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) sub_i) (s sub_i) m
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Heqsub_i: sub_i = dexist i Hi

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Hsenti: has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist i Hi)) (s (dexist i Hi)) m

composite_has_been_sent IM (lift_sub_state IM (elements selection_complement) s) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Hsenti: has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist i Hi)) (s (dexist i Hi)) m

has_been_sent (IM i) (lift_sub_state IM (elements selection_complement) s i) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
s: state PreNonByzantine
m: message
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Hsenti: has_been_sent (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist i Hi)) (s (dexist i Hi)) m

has_been_sent (IM i) (lift_sub_state_to IM (elements selection_complement) (λ n : index, `(vs0 (IM n))) s i) m
by rewrite (lift_sub_state_to_eq _ _ _ _ _ Hi).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hseeded: fixed_set_signed_message IM selection A sender m

fixed_equivocation_constraint IM selection (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, Some m)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: option_valid_message_prop Fixed (Some m)
Hseeded: fixed_set_signed_message IM selection A sender m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HomY: initial_message_prop m ∨ can_emit Fixed m
Hseeded: fixed_set_signed_message IM selection A sender m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hiom: can_emit Fixed m
Hseeded: fixed_set_signed_message IM selection A sender m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hiom: can_emit {| vlsm_type := Fixed; vlsm_machine := preloaded_with_all_messages_vlsm Fixed |} m
Hseeded: fixed_set_signed_message IM selection A sender m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
Hseeded: fixed_set_signed_message IM selection A sender m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
Hsigned: channel_authenticated_message A sender i m

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
Hsigned: option_map A (sender m) = Some i

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: Some (A v) = Some i

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
Heq_v: A v = _v

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
_v: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM _v)) m
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
Heq_v: i = _v

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hiom: can_emit (preloaded_with_all_messages_vlsm (IM i)) m
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hiom: can_emit (preloaded_vlsm (IM i) (λ msg : message, msg ∈ message_dependencies m)) m
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∉ elements (list_to_set (enum index) ∖ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (preloaded_vlsm (IM i) (λ msg : message, msg ∈ message_dependencies m)) m → can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: ¬ (i ∈ list_to_set (enum index) ∧ i ∉ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (preloaded_vlsm (IM i) (λ msg : message, msg ∈ message_dependencies m)) m → can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: ¬ (i ∉ selection)
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (preloaded_vlsm (IM i) (λ msg : message, msg ∈ message_dependencies m)) m → can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

can_emit (preloaded_vlsm (IM i) (λ msg : message, msg ∈ message_dependencies m)) m → can_emit (equivocators_composition_for_directly_observed IM selection (lift_sub_state IM (elements selection_complement) s)) m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i

m0 : message, m0 ∈ message_dependencies m → composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) m0
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: valid l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
sub_j: sub_index (elements (list_to_set (enum index) ∖ selection))
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) sub_j)
s: state PreNonByzantine
m: message
Hv: valid (existT sub_j lj) (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
sub_j: sub_index (elements (list_to_set (enum index) ∖ selection))
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) sub_j)
s: state PreNonByzantine
m: message
Hv: valid (existT sub_j lj) (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
Heqsub_j: sub_j = dexist j Hj

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
s: state PreNonByzantine
m: message
Hv: valid (existT (dexist j Hj) lj) (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
s: state PreNonByzantine
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
s: state PreNonByzantine
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
s: state PreNonByzantine
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
dm: message
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
s: state PreNonByzantine
dm: message
Hfull: has_been_directly_observed (IM j) (s (dexist j Hj)) dm
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
Hdm: dm ∈ message_dependencies m

composite_has_been_directly_observed IM (lift_sub_state IM (elements selection_complement) s) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
s: state PreNonByzantine
dm: message
Hfull: has_been_directly_observed (IM j) (s (dexist j Hj)) dm
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
Hdm: dm ∈ message_dependencies m

has_been_directly_observed (IM j) (lift_sub_state IM (elements selection_complement) s j) dm
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
j: index
Hj: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) j
s: state PreNonByzantine
dm: message
Hfull: has_been_directly_observed (IM j) (s (dexist j Hj)) dm
lj: label (sub_IM IM (elements (list_to_set (enum index) ∖ selection)) (dexist j Hj))
m: message
Hv: valid lj (s (dexist j Hj), Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: i ∈ selection
v: validator
Hsender: sender m = Some v
Hsigned: A v = i
Hdm: dm ∈ message_dependencies m

has_been_directly_observed (IM j) (lift_sub_state_to IM (elements selection_complement) (λ n : index, `(vs0 (IM n))) s j) dm
by rewrite (lift_sub_state_to_eq _ _ _ _ _ Hj). Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

VLSM_embedding PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

VLSM_embedding PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

strong_embedding_valid_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
strong_embedding_transition_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
strong_projection_initial_state_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
strong_embedding_initial_message_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

strong_embedding_valid_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
l: label PreNonByzantine
s: state PreNonByzantine
om: option message
Hv: valid l (s, om)

valid (projT2 l) (lift_sub_state IM (elements selection_complement) s (`(projT1 l)), om)
by apply lift_sub_valid.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

strong_embedding_transition_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
by intro; intros; rapply lift_sub_transition.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

strong_projection_initial_state_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM)) (lift_sub_state IM (elements selection_complement))
by intro; intros; apply (lift_sub_state_initial IM).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies

strong_embedding_initial_message_preservation PreNonByzantine (preloaded_with_all_messages_vlsm (free_composite_vlsm IM))
by right. Qed. Section sec_assuming_initial_messages_lift. Context (Hfixed_non_byzantine_vlsm_lift_initial_message : weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))).
Since FixedNonEquivocating is an induced_validator of Fixed, its initial_messages and validity are derived from valid messages and protocol validity of the larger composition; therefore, the following result becomes very important.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_embedding PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_embedding PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_embedding_valid_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
weak_embedding_transition_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
strong_projection_initial_state_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_embedding_valid_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
by intro; intros; apply fixed_non_byzantine_vlsm_lift_valid.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_embedding_transition_preservation PreNonByzantine Fixed (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
by intro; intros * []; rapply lift_sub_transition.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

strong_projection_initial_state_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
by intro; intros; apply (lift_sub_state_initial IM).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
by intros; apply Hfixed_non_byzantine_vlsm_lift_initial_message. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_incl PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_incl PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

strong_incl_initial_state_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
weak_incl_initial_message_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
weak_incl_valid_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
weak_incl_transition_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

strong_incl_initial_state_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s

initial_state_prop (id s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s

composite_state_sub_projection IM (elements selection_complement) (lift_sub_state IM (elements selection_complement) s) = id s ∧ initial_state_prop (lift_sub_state IM (elements selection_complement) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s

composite_state_sub_projection IM (elements selection_complement) (lift_sub_state IM (elements selection_complement) s) = id s
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s
initial_state_prop (lift_sub_state IM (elements selection_complement) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s

composite_state_sub_projection IM (elements selection_complement) (lift_sub_state IM (elements selection_complement) s) = id s
by apply composite_state_sub_projection_lift_to.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
H19: initial_state_prop s

initial_state_prop (lift_sub_state IM (elements selection_complement) s)
by apply (lift_sub_state_initial IM).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_incl_initial_message_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
by intro; intros; apply initial_message_is_valid; cbn; right.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_incl_valid_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

valid (id l) (id s, om)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

sX : state (composite_vlsm IM (fixed_equivocation_constraint IM selection)), InputValidation (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (id l) (id s, om).1 (id s, om).2 (lift_sub_label IM (elements selection_complement) l) sX
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

InputValidation (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (id l) (id s, om).1 (id s, om).2 (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

composite_label_sub_projection_option IM (elements selection_complement) (lift_sub_label IM (elements selection_complement) l) = Some (id l)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om
composite_state_sub_projection IM (elements selection_complement) (lift_sub_state IM (elements selection_complement) s) = (id s, om).1
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om
input_valid (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, (id s, om).2)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

composite_label_sub_projection_option IM (elements selection_complement) (lift_sub_label IM (elements selection_complement) l) = Some (id l)
by apply composite_label_sub_projection_option_lift.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

composite_state_sub_projection IM (elements selection_complement) (lift_sub_state IM (elements selection_complement) s) = (id s, om).1
by apply composite_state_sub_projection_lift.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

input_valid (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (lift_sub_label IM (elements selection_complement) l) (lift_sub_state IM (elements selection_complement) s, (id s, om).2)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
Hv: input_valid {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} l (s, om)
HsY: valid_state_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} (id s)
HomY: option_valid_message_prop {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True) |} om

VLSM_embedding {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |} (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))
by apply fixed_non_byzantine_vlsm_lift_from_initial.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

weak_incl_transition_preservation (preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender)) (preloaded_vlsm_machine (projection_induced_validator (composite_vlsm IM (fixed_equivocation_constraint IM selection)) (composite_type (sub_IM IM (elements selection_complement))) (composite_label_sub_projection_option IM (elements selection_complement)) (composite_state_sub_projection IM (elements selection_complement)) (lift_sub_label IM (elements selection_complement)) (lift_sub_state IM (elements selection_complement))) (λ _ : message, True))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
l: label {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
s: state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om: option message
s': state {| vlsm_type := composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender)); vlsm_machine := preloaded_vlsm_machine (composite_vlsm (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (no_equivocations_additional_constraint_with_preloaded (sub_IM IM (elements (list_to_set (enum index) ∖ selection))) (free_constraint (sub_IM IM (elements (list_to_set (enum index) ∖ selection)))) (fixed_set_signed_message IM selection A sender))) (fixed_set_signed_message IM selection A sender) |}
om': option message
Ht: transition l (s, om) = (s', om')

transition (id l) (id s, om) = (id s', om')
by apply induced_sub_projection_transition_preservation. Qed.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_eq PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_eq PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_incl {| vlsm_type := PreNonByzantine; vlsm_machine := PreNonByzantine |} {| vlsm_type := PreNonByzantine; vlsm_machine := FixedNonEquivocating |}
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
VLSM_incl {| vlsm_type := PreNonByzantine; vlsm_machine := FixedNonEquivocating |} {| vlsm_type := PreNonByzantine; vlsm_machine := PreNonByzantine |}
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_incl {| vlsm_type := PreNonByzantine; vlsm_machine := PreNonByzantine |} {| vlsm_type := PreNonByzantine; vlsm_machine := FixedNonEquivocating |}
by apply fixed_non_byzantine_incl_fixed_non_equivocating_from_initial.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hfixed_non_byzantine_vlsm_lift_initial_message: weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))

VLSM_incl {| vlsm_type := PreNonByzantine; vlsm_machine := FixedNonEquivocating |} {| vlsm_type := PreNonByzantine; vlsm_machine := PreNonByzantine |}
by apply fixed_non_equivocating_incl_fixed_non_byzantine. Qed. End sec_assuming_initial_messages_lift. Context (Hvalidator : forall i : index, i ∉ selection -> component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i) .
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i

weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i

weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
HmX: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
im: message
Him: initial_message_prop im
Heqm: `(im ↾ Him) = m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hseeded: fixed_set_signed_message IM selection A sender m
valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
im: message
Him: initial_message_prop im
Heqm: `(im ↾ Him) = m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
im: message
Him: initial_message_prop im
Heqm: im = m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
sub_i: sub_index (elements (list_to_set (enum index) ∖ selection))
Him: initial_message_prop m
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Heqsub_i: sub_i = dexist i Hi

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
i: index
Hi: sub_index_prop (elements (list_to_set (enum index) ∖ selection)) i
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H5: Elements index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
m: message
i: index
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H5: Elements index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
m: message
i: index
Him: initial_message_prop m

valid_message_prop Fixed m
message, index, Ci: Type
H5: Elements index Ci
EqDecision1: EqDecision index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
m: message
i: index
Him: initial_message_prop m

initial_message_prop m
by exists i, (exist _ m Him).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hseeded: fixed_set_signed_message IM selection A sender m

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hsigned: non_sub_index_authenticated_message (elements (list_to_set (enum index) ∖ selection)) A sender m
i: index
Hi: i ∈ list_to_set (enum index) ∖ selection
li: label (preloaded_with_all_messages_vlsm (IM i))
si: state (preloaded_with_all_messages_vlsm (IM i))
Hpre_valid: input_constrained (IM i) li (si, Some m)

valid_message_prop Fixed m
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
l: label PreNonByzantine
s: state PreNonByzantine
m: message
Hv: input_valid PreNonByzantine l (s, Some m)
HsY: valid_state_prop Fixed (lift_sub_state IM (elements selection_complement) s)
Hsigned: non_sub_index_authenticated_message (elements (list_to_set (enum index) ∖ selection)) A sender m
i: index
Hi: i ∉ selection
li: label (preloaded_with_all_messages_vlsm (IM i))
si: state (preloaded_with_all_messages_vlsm (IM i))
Hpre_valid: input_constrained (IM i) li (si, Some m)

valid_message_prop Fixed m
by eapply Hvalidator. Qed.
The VLSM corresponding to the induced projection from a fixed-set byzantine composition to the non-byzantine components is trace-equivalent to the VLSM corresponding to the induced projection from the composition of regular components under a fixed-set equivocation constraint to the same components.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i

VLSM_eq PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i

VLSM_eq PreNonByzantine FixedNonEquivocating
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i

weak_embedding_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (elements selection_complement))
by apply validator_fixed_non_byzantine_vlsm_lift_initial_message. Qed.

The main result

Assuming that the non-byzantine components are validators for the fixed_equivocation_constraint we give the following characterization of traces with the fixed_byzantine_trace_alt_property in terms of equivocation:
A trace has the fixed_byzantine_trace_alt_property for a selection of byzantine components iff there exists a valid trace for the Fixed equivocation composition the projections of the two traces to the non-byzantine components, i.e., the components in the selection_complement coincide.
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)

fixed_byzantine_trace_alt_prop IM selection A sender bis btr ↔ ( (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)

fixed_byzantine_trace_alt_prop IM selection A sender bis btr ↔ ( (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)

finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr) ↔ ( (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)

(eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr
finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)

(eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)

finite_valid_trace (pre_induced_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) (fixed_equivocation_constraint IM selection)) (composite_state_sub_projection IM (elements selection_complement) bis) (finite_trace_sub_projection IM (elements selection_complement) btr)
by apply (VLSM_eq_finite_valid_trace validator_fixed_non_byzantine_eq_fixed_non_equivocating).
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr

finite_valid_trace (preloaded_fixed_non_byzantine_vlsm IM selection A sender) (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)
message, index, Ci: Type
H: ElemOf index Ci
H0: Empty Ci
H1: Singleton index Ci
H2: Union Ci
H3: Intersection Ci
H4: Difference Ci
H5: Elements index Ci
EqDecision0: EqDecision index
H6: FinSet index Ci
EqDecision1: EqDecision index
H7: finite.Finite index
IM: index → VLSM message
H8: i : index, HasBeenSentCapability (IM i)
H9: i : index, HasBeenReceivedCapability (IM i)
selection: Ci
validator: Type
A: validator → index
sender: message → option validator
selection_complement:= list_to_set (enum index) ∖ selection: Ci
PreNonByzantine:= preloaded_fixed_non_byzantine_vlsm IM selection A sender: VLSM message
Fixed:= fixed_equivocation_vlsm_composition IM selection: VLSM message
FixedNonEquivocating:= pre_induced_sub_projection IM (elements selection_complement) (fixed_equivocation_constraint IM selection): VLSM message
no_initial_messages_in_IM: no_initial_messages_in_IM_prop IM
can_emit_signed: channel_authentication_prop IM A sender
Hsender_safety:= channel_authentication_sender_safety IM A sender can_emit_signed: sender_safety_alt_prop IM A sender
Cm: Type
H10: ElemOf message Cm
H11: Empty Cm
H12: Singleton message Cm
H13: Union Cm
H14: Intersection Cm
H15: Difference Cm
H16: Elements message Cm
EqDecision2: EqDecision message
H17: FinSet message Cm
message_dependencies: message → Cm
Irreflexive0: Irreflexive (msg_dep_happens_before message_dependencies)
H18: i : index, MessageDependencies (IM i) message_dependencies
Hfull: i : index, message_dependencies_full_node_condition_prop (IM i) message_dependencies
Hvalidator: i : index, i ∉ selection → component_message_validator_prop IM (fixed_equivocation_constraint IM selection) i
bis: composite_state IM
btr: list (composite_transition_item IM)
Htr: (eis : state Fixed) (etr : list transition_item), finite_valid_trace Fixed eis etr ∧ composite_state_sub_projection IM (elements selection_complement) eis = composite_state_sub_projection IM (elements selection_complement) bis ∧ finite_trace_sub_projection IM (elements selection_complement) etr = finite_trace_sub_projection IM (elements selection_complement) btr

finite_valid_trace {| vlsm_type := PreNonByzantine; vlsm_machine := FixedNonEquivocating |} (composite_state_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) bis) (finite_trace_sub_projection IM (elements (list_to_set (enum index) ∖ selection)) btr)
by apply fixed_non_equivocating_traces_char. Qed. End sec_validator_fixed_set_byzantine. End sec_fixed_non_equivocating_vs_byzantine.