From stdpp Require Import prelude finite. From VLSM.Core Require Import VLSM VLSMProjections Composition.From VLSM.Core Require Import ConstrainedVLSM.
Tutorial: Round-Based Asynchronous Muddy Children Puzzle
- we allow an arbitrary but fixed number of children, and
- we allow the children to communicate asynchronously.
- there is at least one muddy child, and
- any child sees all the muddy children except themselves.
Inductive Label : Type := | init | emit | receive. Inductive ChildStatus : Type := | undecided | muddy | clean. Record RoundStatus : Type := mkRS { rs_round : nat; rs_status : ChildStatus; }.
Show RoundStatus using the constructor instead of the record syntax.
Add Printing Constructor RoundStatus. Section sec_muddy. Context (index : Type) `{finite.Finite index} `{Inhabited index} `{FinSet index indexSet} .
Children maintain the set of (indices of) other children they perceive
as being muddy and (except for their initial state) the current round they
perceive themselves to be at and their ChildStatus.
Record State : Type := mkSt
{
st_obs : indexSet;
st_rs : option RoundStatus;
}.
Show State using the constructor instead of the record syntax.
Add Printing Constructor State.
A message carries the identity of its sender, and shares the round number
and their ChildStatus.
Record Message : Type := mkMsg
{
msg_index : index;
msg_round : nat;
msg_status : ChildStatus;
}.
Show Message using the constructor instead of the record syntax.
Add Printing Constructor Message. Definition MCType : VLSMType Message := {| state := State; label := Label; |}. Definition MC_initial_state_prop (s : State) : Prop := st_rs s = None. Equations MC_transition (i : index) (l : Label) (s : State) (om : option Message) : State * option Message := MC_transition i init (mkSt o None) None with size o := { | 0 => (mkSt o (Some (mkRS 0 muddy)), None) | S _ => (mkSt o (Some (mkRS 0 undecided)), None) }; MC_transition i emit (mkSt o (Some (mkRS r status))) None := ((mkSt o (Some (mkRS r status))), Some (mkMsg i r status)); MC_transition i receive (mkSt o (Some (mkRS r undecided))) (Some (mkMsg j r' clean)) with decide (j ∈ o) := { | right Hin => if decide (r' = size o) then (mkSt o (Some (mkRS r' clean)), None) else if decide (r' = size o + 1) then (mkSt o (Some (mkRS (r' - 1) muddy)), None) else (mkSt o (Some (mkRS r undecided)), None) | left Hin => (mkSt o (Some (mkRS r undecided)), None) }; MC_transition i receive (mkSt o (Some (mkRS r undecided))) (Some (mkMsg j r' muddy)) with decide (j ∈ o) := { | left Hin => if decide (r' = size o) then (mkSt o (Some (mkRS r' muddy)), None) else if decide (r' = size o - 1) then (mkSt o (Some (mkRS (r' + 1) clean)), None) else (mkSt o (Some (mkRS r undecided)), None) | right Hin => (mkSt o (Some (mkRS r undecided)), None) }; (** One of the most interesting cases of the transition function is the one when a child [i] who doesn't know their status receives a message from a child [j] who also doesn't know their own status. However, being at round [r'], [j] knows there are more than [r'] muddy children. *) MC_transition i receive (mkSt o (Some (mkRS r undecided))) (Some (mkMsg j r' undecided)) (** If child [i] sees [j] as muddy, they can infer that there are actually more than [r' + 1] muddy children. *) with decide (j ∈ o) := { | left Hin => (** If [r' + 1] is less than or equal to the current round number of [i], then the round doesn't change and the status remains undecided. *) if decide (r' < r) then (mkSt o (Some (mkRS r undecided)), None) (** Else, we update the round number to [r' + 1], and have to compare it to the number of muddy children seen by child [i]. If its less than that, the information gained is not useful enough to infer anything, so the status remains undecided. *) else if decide (r <= r' < size o - 1) then (mkSt o (Some (mkRS (r' + 1) undecided)), None) (** Else, if the updated round ever becomes equal to the number of muddy children seen by [i], they can conclude they are muddy. *) else if decide (r' = size o - 1) then (mkSt o (Some (mkRS (r' + 1) muddy)), None) else (mkSt o (Some (mkRS r undecided)), None) | right Hin => if decide (r' <= r) then (mkSt o (Some (mkRS r undecided)), None) else if decide (r < r' < size o) then (mkSt o (Some (mkRS r' undecided)), None) else if decide (r' = size o) then (mkSt o (Some (mkRS r' muddy)), None) else (mkSt o (Some (mkRS r undecided)), None) }; MC_transition _ _ s _ := (s, None). Definition state_status (s : option RoundStatus) : ChildStatus := from_option rs_status undecided s. Definition state_round (s : option RoundStatus) : nat := from_option rs_round 0 s. Definition state_round_inc (s : State) : nat := match st_rs s with | Some rs => S (rs_round rs) | _ => 0 end. Definition message_status (om : option Message) : ChildStatus := from_option msg_status undecided om. Definition message_round (om : option Message) : nat := from_option msg_round 0 om. Definition message_index (om : option Message) (i : index) : index := from_option msg_index i om. Inductive MC_valid : Label -> State -> option Message -> Prop := | MC_valid_init : forall o : indexSet, MC_valid init (mkSt o None) None | MC_valid_emit : forall (o : indexSet) (rs : RoundStatus), MC_valid emit (mkSt o (Some rs)) None | MC_valid_receive : forall (o : indexSet) (rs : RoundStatus) (m : Message), MC_valid receive (mkSt o (Some rs)) (Some m).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetEqDecision ChildStatusby intros x y; unfold Decision; decide equality. Qed. Definition MC_initial_state_type : Type := {st : State | MC_initial_state_prop st}. Program Definition MC_initial_state : MC_initial_state_type := exist _ (mkSt ∅ None) _.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetEqDecision ChildStatusindexSet, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetMC_initial_state_prop (mkSt ∅ None)done. Qed.indexSet, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetMC_initial_state_prop (mkSt ∅ None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : State, Decision (MC_initial_state_prop s)by typeclasses eauto. Qed. #[export] Instance Inhabited_MC_initial_state_type : Inhabited (MC_initial_state_type) := populate (MC_initial_state). Definition MCMachine (i : index) : VLSMMachine MCType := {| initial_state_prop := MC_initial_state_prop; initial_message_prop := const False; transition := fun l '(st, om) => MC_transition i l st om; valid := fun l '(st, om) => MC_valid l st om; |}. Definition MCVLSM (i : index) : VLSM Message := {| vlsm_type := MCType; vlsm_machine := MCMachine i; |}.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : State, Decision (MC_initial_state_prop s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : composite_state MCVLSM, Decision (composite_initial_state_prop MCVLSM s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : composite_state MCVLSM, Decision (composite_initial_state_prop MCVLSM s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMForall (λ n : index, initial_state_prop (s n)) (enum index) ↔ composite_initial_state_prop MCVLSM sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMDecision (Forall (λ n : index, initial_state_prop (s n)) (enum index))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMForall (λ n : index, initial_state_prop (s n)) (enum index) ↔ composite_initial_state_prop MCVLSM sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM∀ x : index, (x ∈ enum index → initial_state_prop (s x)) ↔ initial_state_prop (s x)by intros Hx; apply Hx, elem_of_enum.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index(x ∈ enum index → initial_state_prop (s x)) → initial_state_prop (s x)by typeclasses eauto. Qed. Definition MuddyUnion (s : composite_state MCVLSM) : indexSet := ⋃ (map (fun i => st_obs (s i)) (enum index)).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMDecision (Forall (λ n : index, initial_state_prop (s n)) (enum index))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: indexi ∈ st_obs (s j) → i ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: indexi ∈ st_obs (s j) → i ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: index
Hobs: i ∈ st_obs (s j)i ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: index
Hobs: i ∈ st_obs (s j)∃ X : indexSet, X ∈ map (λ i : index, st_obs (s i)) (enum index) ∧ i ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: index
Hobs: i ∈ st_obs (s j)st_obs (s j) ∈ map (λ i : index, st_obs (s i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: index
Hobs: i ∈ st_obs (s j)∃ y : index, st_obs (s j) = st_obs (s y) ∧ y ∈ enum indexby apply elem_of_enum. Qed. Definition consistent (s : composite_state MCVLSM) : Prop := MuddyUnion s ≢ ∅ /\ forall n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}. Definition MC_no_equivocation (s : composite_state MCVLSM) (m : Message) : Prop := match m with | mkMsg n r' status' => match s n with | mkSt _ (Some (mkRS r status)) => (status' = status /\ r' = r) \/ (status' = undecided /\ r' < r) | _ => False end end. Inductive MC_no_equivocation_inductive : composite_state MCVLSM -> Message -> Prop := | MC_no_equivocation_inductive_msg_eq : forall s n o r status, s n = mkSt o (Some (mkRS r status)) -> MC_no_equivocation_inductive s (mkMsg n r status) | MC_no_equivocation_inductive_undecided : forall s n o r r' status, s n = mkSt o (Some (mkRS r status)) -> r' < r -> MC_no_equivocation_inductive s (mkMsg n r' undecided).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i, j: index
Hobs: i ∈ st_obs (s j)j ∈ enum indexindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (m : Message), MC_no_equivocation s m ↔ MC_no_equivocation_inductive s mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (m : Message), MC_no_equivocation s m ↔ MC_no_equivocation_inductive s mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
msg_status0: ChildStatusmatch s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end → MC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 msg_status0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
msg_status0: ChildStatusMC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 msg_status0) → match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
msg_status0: ChildStatusmatch s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end → MC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 msg_status0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
st_obs0: indexSet
st_rs0: option RoundStatus
r: RoundStatus
rs_round0: nat
rs_status0: ChildStatus
H11: r = mkRS rs_round0 rs_status0
H10: st_rs0 = Some (mkRS rs_round0 rs_status0)
H9: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))msg_status0 = rs_status0 ∧ msg_round0 = rs_round0 ∨ msg_status0 = undecided ∧ msg_round0 < rs_round0 → MC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 msg_status0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
st_obs0: indexSet
rs_round0: nat
rs_status0: ChildStatus
H9: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))MC_no_equivocation_inductive s (mkMsg msg_index0 rs_round0 rs_status0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
st_obs0: indexSet
rs_round0: nat
rs_status0: ChildStatus
H9: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
H13: msg_round0 < rs_round0MC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 undecided)by eapply MC_no_equivocation_inductive_msg_eq; eauto.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
st_obs0: indexSet
rs_round0: nat
rs_status0: ChildStatus
H9: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))MC_no_equivocation_inductive s (mkMsg msg_index0 rs_round0 rs_status0)by eapply MC_no_equivocation_inductive_undecided; eauto.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
st_obs0: indexSet
rs_round0: nat
rs_status0: ChildStatus
H9: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
H13: msg_round0 < rs_round0MC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 undecided)by repeat case_match; inversion 1; itauto congruence. Qed. Definition MC_constraint (l : composite_label MCVLSM) (sm : composite_state MCVLSM * option Message) : Prop := match l, sm with | existT _ init, (s, _) => consistent s | existT _ receive, (s, Some m) => MC_no_equivocation s m | _, _ => True end. Definition MC_composite_vlsm : VLSM Message := composite_vlsm MCVLSM MC_constraint. Definition MC_final_state (s : composite_state MCVLSM) : Prop := forall n : index, state_status (st_rs (s n)) <> undecided. Definition MC_initial_consistent_state (s : composite_state MCVLSM) : Prop := composite_initial_state_prop MCVLSM s /\ consistent s. Definition MC_non_initial_valid_state (s : composite_state MCVLSM) : Prop := valid_state_prop MC_composite_vlsm s /\ ~ (composite_initial_state_prop MCVLSM s). Definition MC_obs_equivalence (s1 s2 : composite_state MCVLSM) : Prop := forall n : index, st_obs (s1 n) ≡ st_obs (s2 n).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
msg_index0: index
msg_round0: nat
msg_status0: ChildStatusMC_no_equivocation_inductive s (mkMsg msg_index0 msg_round0 msg_status0) → match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetReflexive MC_obs_equivalencedone. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetReflexive MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetSymmetric MC_obs_equivalenceby unfold MC_obs_equivalence; intros x y Hxy. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetSymmetric MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetTransitive MC_obs_equivalenceby unfold MC_obs_equivalence; intros s1 s2 s3 H12 H23; etransitivity; [apply H12 | apply H23]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetTransitive MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetEquivalence MC_obs_equivalenceby split; typeclasses eauto. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSetEquivalence MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
si: Statest_obs (s i) ≡ st_obs si → MC_obs_equivalence s (state_update MCVLSM s i si)by intros Hobs n; destruct (decide (n = i)); subst; state_update_simpl. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
si: Statest_obs (s i) ≡ st_obs si → MC_obs_equivalence s (state_update MCVLSM s i si)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSMMC_obs_equivalence s1 s2 → MuddyUnion s1 ≡ MuddyUnion s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSMMC_obs_equivalence s1 s2 → MuddyUnion s1 ≡ MuddyUnion s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2⋃ map (λ i : index, st_obs (s1 i)) (enum index) ≡ ⋃ map (λ i : index, st_obs (s2 i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i: index(∃ X : indexSet, X ∈ map (λ i : index, st_obs (s1 i)) (enum index) ∧ i ∈ X) ↔ (∃ X : indexSet, X ∈ map (λ i : index, st_obs (s2 i)) (enum index) ∧ i ∈ X)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)∃ X : indexSet, X ∈ map (λ i : index, st_obs (s2 i)) (enum index) ∧ i ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)∃ X : indexSet, X ∈ map (λ i : index, st_obs (s1 i)) (enum index) ∧ i ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)∃ X : indexSet, X ∈ map (λ i : index, st_obs (s2 i)) (enum index) ∧ i ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)st_obs (s2 y) ∈ map (λ i : index, st_obs (s2 i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)i ∈ st_obs (s2 y)by apply elem_of_list_fmap; exists y.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)st_obs (s2 y) ∈ map (λ i : index, st_obs (s2 i)) (enum index)by apply Hobs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s1 y)i ∈ st_obs (s2 y)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)∃ X : indexSet, X ∈ map (λ i : index, st_obs (s1 i)) (enum index) ∧ i ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)st_obs (s1 y) ∈ map (λ i : index, st_obs (s1 i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)i ∈ st_obs (s1 y)by apply elem_of_list_fmap; exists y.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)st_obs (s1 y) ∈ map (λ i : index, st_obs (s1 i)) (enum index)by apply Hobs. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
i, y: index
Hy: y ∈ enum index
Hi: i ∈ st_obs (s2 y)i ∈ st_obs (s1 y)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSMMC_obs_equivalence s1 s2 → consistent s1 → consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSMMC_obs_equivalence s1 s2 → consistent s1 → consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}MuddyUnion s2 ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}∀ n : index, st_obs (s2 n) ≡ MuddyUnion s2 ∖ {[n]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}MuddyUnion s2 ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: ⋃ map (λ i : index, st_obs (s1 i)) (enum index) ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}⋃ map (λ i : index, st_obs (s2 i)) (enum index) ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: ¬ (∀ x : indexSet, x ∈ map (λ i : index, st_obs (s1 i)) (enum index) → x ≡ ∅)
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}¬ (∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅∀ x : indexSet, x ∈ map (λ i : index, st_obs (s1 i)) (enum index) → x ≡ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅
x: indexSet
Hx: x ∈ map (λ i : index, st_obs (s1 i)) (enum index)x ≡ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅
j: indexst_obs (s1 j) ≡ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: ∀ n : index, st_obs (s1 n) ≡ st_obs (s2 n)
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅
j: indexst_obs (s2 j) ≡ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: ∀ n : index, st_obs (s1 n) ≡ st_obs (s2 n)
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅
j: index∃ y : index, st_obs (s2 j) = st_obs (s2 y) ∧ y ∈ enum indexby apply elem_of_enum.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: ∀ n : index, st_obs (s1 n) ≡ st_obs (s2 n)
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
Hnempty: ∀ x : indexSet, x ∈ map (λ i : index, st_obs (s2 i)) (enum index) → x ≡ ∅
j: indexj ∈ enum indexindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}∀ n : index, st_obs (s2 n) ≡ MuddyUnion s2 ∖ {[n]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
j, x: indexx ∈ st_obs (s2 j) ↔ x ∈ MuddyUnion s2 ∖ {[j]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: MC_obs_equivalence s1 s2
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
j, x: indexx ∈ st_obs (s2 j) ↔ x ∈ MuddyUnion s1 ∖ {[j]}by apply Hcons. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
Hobs: ∀ n : index, st_obs (s1 n) ≡ st_obs (s2 n)
Hnempty: MuddyUnion s1 ≢ ∅
Hcons: ∀ n : index, st_obs (s1 n) ≡ MuddyUnion s1 ∖ {[n]}
j, x: indexx ∈ st_obs (s1 j) ↔ x ∈ MuddyUnion s1 ∖ {[j]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM∀ (l : composite_label MCVLSM) (s' : composite_state MCVLSM) (m m' : option Message), composite_transition MCVLSM l (s, m) = (s', m') → MC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM∀ (l : composite_label MCVLSM) (s' : composite_state MCVLSM) (m m' : option Message), composite_transition MCVLSM l (s, m) = (s', m') → MC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
l: composite_label MCVLSM
s': composite_state MCVLSM
m, m': option Message
H9: (let (i, li) := l in let (si', om') := transition li (s i, m) in (state_update MCVLSM s i si', om')) = ( s', m')MC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
li: label (MCVLSM i)
m, m': option Message
s0: state (MCVLSM i)
Ht: MC_transition i li (s i) m = (s0, m')MC_obs_equivalence s (state_update MCVLSM s i s0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
li: label (MCVLSM i)
m, m': option Message
s0: state (MCVLSM i)
Ht: MC_transition i li (s i) m = (s0, m')st_obs (s i) ≡ st_obs s0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
li: label (MCVLSM i)
m, m': option Message
st_obs0: indexSet
st_rs0: option RoundStatus
st_obs1: indexSet
st_rs1: option RoundStatus
Ht: MC_transition i li (mkSt st_obs1 st_rs1) m = (mkSt st_obs0 st_rs0, m')st_obs (mkSt st_obs1 st_rs1) ≡ st_obs (mkSt st_obs0 st_rs0)by apply FunctionalElimination_MC_transition; intros; inversion Ht; subst; try done; repeat case_decide; inversion H10. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
li: label (MCVLSM i)
m, m': option Message
st_obs0: indexSet
st_rs0: option RoundStatus
st_obs1: indexSet
st_rs1: option RoundStatusMC_transition i li (mkSt st_obs1 st_rs1) m = (mkSt st_obs0 st_rs0, m') → st_obs (mkSt st_obs1 st_rs1) ≡ st_obs (mkSt st_obs0 st_rs0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'MC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'MC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'PreOrder MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'input_valid_transition_preserving MC_composite_vlsm MC_obs_equivalenceby split; typeclasses eauto.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'PreOrder MC_obs_equivalenceindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'input_valid_transition_preserving MC_composite_vlsm MC_obs_equivalenceby apply MC_trans_preserves_obs_equiv with l om1 om2. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'
s1, s2: state MC_composite_vlsm
l: label MC_composite_vlsm
om1, om2: option Message
H9: input_valid MC_composite_vlsm l (s1, om1)
H10: transition l (s1, om1) = (s2, om2)MC_obs_equivalence s1 s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'MuddyUnion s ≡ MuddyUnion s'by intros; apply MC_obs_equiv_preserves_muddy, MC_in_futures_preserves_obs_equiv. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'MuddyUnion s ≡ MuddyUnion s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'consistent s ↔ consistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'consistent s ↔ consistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'MuddyUnion s ≢ ∅ ∧ (∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}) ↔ MuddyUnion s' ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'
Hfutures': MC_obs_equivalence s s'MuddyUnion s ≢ ∅ ∧ (∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}) ↔ MuddyUnion s' ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'
Hfutures': ∀ n : index, st_obs (s n) ≡ st_obs (s' n)MuddyUnion s ≢ ∅ ∧ (∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}) ↔ MuddyUnion s' ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: in_futures MC_composite_vlsm s s'
Hfutures': ∀ n : index, st_obs (s n) ≡ st_obs (s' n)MuddyUnion s ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s ∖ {[n]}) ↔ MuddyUnion s' ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]})by setoid_rewrite Hfutures. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, s': composite_state MCVLSM
Hfutures: MuddyUnion s ≡ MuddyUnion s'
Hfutures': ∀ n : index, st_obs (s n) ≡ st_obs (s' n)MuddyUnion s ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s ∖ {[n]}) ↔ MuddyUnion s' ≢ ∅ ∧ (∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : composite_state MCVLSM, MC_non_initial_valid_state s → consistent sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ s : composite_state MCVLSM, MC_non_initial_valid_state s → consistent sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnon_initial: ¬ composite_initial_state_prop MCVLSM sconsistent sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, om) (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
n: ¬ composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid init (s i) om
Hc: MC_constraint (existT i init) (s, om)
Ht: transition (existT i init) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid emit (s i) om
Hc: MC_constraint (existT i emit) (s, om)
Ht: transition (existT i emit) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid receive (s i) om
Hc: MC_constraint (existT i receive) (s, om)
Ht: transition (existT i receive) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid init (s i) om
Hc: MC_constraint (existT i init) (s, om)
Ht: transition (existT i init) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'by eapply MC_trans_preserves_obs_equiv with (s := s) (l := existT i init) (m := om) (m' := om').index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid init (s i) om
Hc: MC_constraint (existT i init) (s, om)
Ht: transition (existT i init) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sMC_obs_equivalence s s'by inversion Hv; specialize (c i); rewrite <- H9 in c.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid emit (s i) om
Hc: MC_constraint (existT i emit) (s, om)
Ht: transition (existT i emit) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'by inversion Hv; specialize (c i); rewrite <- H9 in c.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
i: index
om, om': option Message
s: state MC_composite_vlsm
Hv: MC_valid receive (s i) om
Hc: MC_constraint (existT i receive) (s, om)
Ht: transition (existT i receive) (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
c: composite_initial_state_prop MCVLSM sconsistent s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
n: ¬ composite_initial_state_prop MCVLSM sconsistent s'by apply MC_trans_preserves_obs_equiv with (s := s) (l := l) (m := om) (m' := om'). Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s': state MC_composite_vlsm
l: label MC_composite_vlsm
om, om': option Message
s: state MC_composite_vlsm
Hv: valid l (s, om)
Hc: MC_constraint l (s, om)
Ht: transition l (s, om) = (s', om')
Hnon_initial: ¬ composite_initial_state_prop MCVLSM s'
IHHvalid: ¬ composite_initial_state_prop MCVLSM s → consistent s
n: ¬ composite_initial_state_prop MCVLSM sMC_obs_equivalence s s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∈ MuddyUnion s → size (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∈ MuddyUnion s → size (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∈ MuddyUnion ssize (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∈ MuddyUnion ssize (MuddyUnion s) - size {[j]} = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∈ MuddyUnion s{[j]} ⊆ MuddyUnion sby rewrite size_singleton.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∈ MuddyUnion ssize (MuddyUnion s) - size {[j]} = size (MuddyUnion s) - 1by rewrite singleton_subseteq_l. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∈ MuddyUnion s{[j]} ⊆ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∈ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∈ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}j ∈ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (st_obs (s j)) = size (MuddyUnion s) - 1 → j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (st_obs (s j)) = size (MuddyUnion s) - 1j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) ≠ 0 → j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) ≠ 0 → j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0
Hsizennull: size (MuddyUnion s ∩ {[j]}) ≠ 0j ∈ MuddyUnion sby set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0
Hsizennull: MuddyUnion s ∩ {[j]} ≢ ∅j ∈ MuddyUnion sby lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s) - 1
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∉ MuddyUnion s → size (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∉ MuddyUnion s → size (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∉ MuddyUnion ssize (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∉ MuddyUnion ssize (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)by symmetry; apply size_empty_iff; set_solver. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: j ∉ MuddyUnion s0 = size (MuddyUnion s ∩ {[j]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∉ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: indexconsistent s → j ∉ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}j ∉ MuddyUnion s ↔ size (st_obs (s j)) = size (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (st_obs (s j)) = size (MuddyUnion s) → j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (st_obs (s j)) = size (MuddyUnion s)j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) = 0 → j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) = 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) = 0 → j ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0
Hsize0: size (MuddyUnion s ∩ {[j]}) = 0j ∉ MuddyUnion sby set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0
Hsize0: MuddyUnion s ∩ {[j]} ≡ ∅j ∉ MuddyUnion sby lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
Hnempty: MuddyUnion s ≢ ∅
Hn: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsize: size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]}) = size (MuddyUnion s)
n: size (MuddyUnion s) ≠ 0size (MuddyUnion s ∩ {[j]}) = 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMconsistent s → ∀ n : index, size (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMconsistent s → ∀ n : index, size (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hcons: consistent s
n: indexsize (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hcons: consistent s
n: index
Hdec: n ∈ MuddyUnion ssize (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hcons: consistent s
n: index
Hdec: n ∉ MuddyUnion ssize (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1by apply MC_muddy_number_of_muddy_seen in Hdec as ->; [lia |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hcons: consistent s
n: index
Hdec: n ∈ MuddyUnion ssize (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1by apply MC_clean_number_of_muddy_seen in Hdec as ->; [lia |]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hcons: consistent s
n: index
Hdec: n ∉ MuddyUnion ssize (st_obs (s n)) <= size (MuddyUnion s) <= size (st_obs (s n)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = clean → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) clean)), None)by intros; rewrite H9, H10, MC_transition_equation_9; unfold MC_transition_clause_3; repeat case_decide. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = clean → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) clean)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = clean → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) + 1 → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m - 1) muddy)), None)by intros; rewrite H9, H10, MC_transition_equation_9; unfold MC_transition_clause_3; repeat case_decide; try done; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = clean → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) + 1 → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m - 1) muddy)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = muddy → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None)by intros; rewrite H9, H10, MC_transition_equation_8; unfold MC_transition_clause_4; repeat case_decide. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = muddy → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = muddy → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) - 1 → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) clean)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = muddy → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) - 1 → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) clean)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
s: State
m: Message
H9: state_status (st_rs s) = undecided
H10: msg_status m = muddy
H11: msg_index m ∈ st_obs s
H12: msg_round m = size (st_obs s) - 1
H13: msg_index m ∈ st_obs s
H14: msg_round m = size (st_obs s)(mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) clean)), None)by apply non_empty_inhabited, size_non_empty_iff in H11. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
s: State
m: Message
H9: state_status (st_rs s) = undecided
H10: msg_status m = muddy
H11: msg_index m ∈ st_obs s
H12: msg_round m = size (st_obs s) - 1
H13: msg_index m ∈ st_obs s
H14: msg_round m = size (st_obs s)
e: size (st_obs s) = 0(mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) clean)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) - 1 → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) muddy)), None)by intros; rewrite H9, H10, MC_transition_equation_7; unfold MC_transition_clause_5; repeat case_decide; try done; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∈ st_obs s → msg_round m = size (st_obs s) - 1 → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) muddy)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None)by intros; rewrite H9, H10, MC_transition_equation_7; unfold MC_transition_clause_5; repeat case_decide; try done; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∉ st_obs s → msg_round m = size (st_obs s) → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m) muddy)), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∈ st_obs s → state_round (st_rs s) <= msg_round m < size (st_obs s) - 1 → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) (state_status (st_rs s)))), None)by intros; rewrite H9, H10, MC_transition_equation_7; unfold MC_transition_clause_5; repeat case_decide; try done; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (i : index) (s : State) (m : Message), state_status (st_rs s) = undecided → msg_status m = undecided → msg_index m ∈ st_obs s → state_round (st_rs s) <= msg_round m < size (st_obs s) - 1 → state_round (st_rs s) < size (st_obs s) → MC_transition i receive (mkSt (st_obs s) (Some (mkRS (state_round (st_rs s)) (state_status (st_rs s))))) (Some (mkMsg (msg_index m) (msg_round m) (msg_status m))) = (mkSt (st_obs s) (Some (mkRS (msg_round m + 1) (state_status (st_rs s)))), None)
Definition MC_component_invariant_helper (s : State) (union : indexSet) : Prop := match state_status (st_rs s) with | undecided => state_round (st_rs s) < size (st_obs s) | clean => state_round (st_rs s) = size (st_obs s) /\ size union = size (st_obs s) | muddy => state_round (st_rs s) = size (st_obs s) /\ size union - 1 = size (st_obs s) end. Definition MC_component_invariant (s : composite_state MCVLSM) (i : index) : Prop := MC_component_invariant_helper (s i) (MuddyUnion s). Inductive MC_component_invariant_inductive : composite_state MCVLSM -> index -> Prop := | component_invariant_undecided : forall s i, state_status (st_rs (s i)) = undecided -> state_round (st_rs (s i)) < size (st_obs (s i)) -> MC_component_invariant_inductive s i | component_invariant_clean : forall s i, state_status (st_rs (s i)) = clean -> state_round (st_rs (s i)) = size (st_obs (s i)) -> size (MuddyUnion s) = size (st_obs (s i)) -> MC_component_invariant_inductive s i | component_invariant_muddy : forall s i, state_status (st_rs (s i)) = muddy -> state_round (st_rs (s i)) = size (st_obs (s i)) -> size (MuddyUnion s) - 1 = size (st_obs (s i)) -> MC_component_invariant_inductive s i.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), MC_component_invariant s i ↔ MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), MC_component_invariant s i ↔ MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: indexMC_component_invariant s i ↔ MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: indexmatch state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i)) end ↔ MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = undecided
H10: state_round (st_rs (s i)) < size (st_obs (s i))MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = undecided
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) < size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = muddy
H10: state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i))MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = muddy
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = clean
H10: state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i))MC_component_invariant_inductive s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = clean
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i))by apply component_invariant_undecided.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = undecided
H10: state_round (st_rs (s i)) < size (st_obs (s i))MC_component_invariant_inductive s iby inversion H10; [| congruence..].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = undecided
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) < size (st_obs (s i))by apply component_invariant_muddy; destruct H10.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = muddy
H10: state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i))MC_component_invariant_inductive s iby inversion H10; split; congruence.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = muddy
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i))by apply component_invariant_clean; destruct H10.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = clean
H10: state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i))MC_component_invariant_inductive s iby inversion H10; split; congruence. Qed. Definition MC_composite_invariant (s : composite_state MCVLSM) : Prop := forall (i : index), initial_state_prop (MCVLSM i) (s i) \/ MC_component_invariant s i. Definition MC_composite_invariant_inductive (s : composite_state MCVLSM) : Prop := forall (i : index), initial_state_prop (MCVLSM i) (s i) \/ MC_component_invariant_inductive s i.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
H9: state_status (st_rs (s i)) = clean
H10: MC_component_invariant_inductive s istate_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSeto ≡ st_obs (s i) → j ∈ o → consistent s → (∀ k : index, st_obs (sm k) ≡ st_obs (s k)) → size o - 1 < size (st_obs (sm j)) → size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSeto ≡ st_obs (s i) → j ∈ o → consistent s → (∀ k : index, st_obs (sm k) ≡ st_obs (s k)) → size o - 1 < size (st_obs (sm j)) → size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSet
Heqo: o ≡ st_obs (s i)
Hin: j ∈ o
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size o - 1 < size (st_obs (sm j))size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSet
Heqo: o ≡ st_obs (s i)
Hin: j ∈ o
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size o - 1 < size (st_obs (sm j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size (st_obs (s i)) - 1 < size (st_obs (sm j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (st_obs (sm j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s) - size {[j]}
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}{[j]} ⊆ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s) - size {[j]}
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s) - 1
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = oby destruct Hconsistent; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: size (st_obs (s i)) <= size (MuddyUnion s) <= size (st_obs (s i)) + 1
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s) - 1
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}{[j]} ⊆ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}j ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}∃ X : indexSet, X ∈ map (λ i : index, st_obs (s i)) (enum index) ∧ j ∈ Xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}st_obs (s i) ∈ map (λ i : index, st_obs (s i)) (enum index)by exists i; split; [| apply elem_of_enum]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∈ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o - 1 < size (MuddyUnion s ∖ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}∃ y : index, st_obs (s i) = st_obs (s y) ∧ y ∈ enum indexindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSeto ≡ st_obs (s i) → j ∉ o → consistent s → (∀ k : index, st_obs (sm k) ≡ st_obs (s k)) → size o < size (st_obs (sm j)) → size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSeto ≡ st_obs (s i) → j ∉ o → consistent s → (∀ k : index, st_obs (sm k) ≡ st_obs (s k)) → size o < size (st_obs (sm j)) → size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSet
Heqo: o ≡ st_obs (s i)
Hin: j ∉ o
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size o < size (st_obs (sm j))size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
o: indexSet
Heqo: o ≡ st_obs (s i)
Hin: j ∉ o
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size o < size (st_obs (sm j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size (st_obs (s i)) < size (st_obs (sm j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size (st_obs (s i)) < size (st_obs (s j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}size (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
Hinvs: size (st_obs (s i)) < size (st_obs (s j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jsize (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (st_obs (s j))
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jsize (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jsize (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jsize (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ j0 = size (MuddyUnion s ∩ {[j]})by apply MC_number_of_muddy_seen with (n := i) in Hconsistent; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jsize (MuddyUnion s) - 1 = oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ j0 = size (MuddyUnion s ∩ {[j]})by rewrite Hcons in Hin; set_solver. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s, sm: composite_state MCVLSM
i, j: index
Hin: j ∉ st_obs (s i)
Hconsistent: consistent s
Hobs_equiv: ∀ k : index, st_obs (sm k) ≡ st_obs (s k)
o: nat
Heqo: o = size (st_obs (s i))
Hinvs: o < size (MuddyUnion s) - size (MuddyUnion s ∩ {[j]})
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n: i ≠ jMuddyUnion s ∩ {[j]} ≡ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is, s: composite_state MCVLSM
tr: list (composite_transition_item MCVLSM)finite_valid_trace_init_to MC_composite_vlsm is s tr → ∀ i : index, ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is, s: composite_state MCVLSM
tr: list (composite_transition_item MCVLSM)finite_valid_trace_init_to MC_composite_vlsm is s tr → ∀ i : index, ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is, s: composite_state MCVLSM
tr: list (composite_transition_item MCVLSM)
Htr: finite_valid_trace_init_to MC_composite_vlsm is s tr
i: index
Hnoninit: ¬ MC_initial_state_prop (s i)∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
n: ¬ MC_initial_state_prop (s i)∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
m: MC_initial_state_prop (s i)∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
n: ¬ MC_initial_state_prop (s i)∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initby exists item; rewrite elem_of_app; itauto.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
n: ¬ MC_initial_state_prop (s i)
item: composite_transition_item MCVLSM
Hitem: item ∈ tr
Hinit: projT2 (VLSM.l item) = init∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
m: MC_initial_state_prop (s i)∃ item : composite_transition_item MCVLSM, item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
m: MC_initial_state_prop (s i)?item ∈ tr ++ [{| l := l; input := iom; destination := sf; output := oom |}] ∧ projT2 (VLSM.l ?item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
m: MC_initial_state_prop (s i)(?item ∈ tr ∨ ?item = {| l := l; input := iom; destination := sf; output := oom |}) ∧ projT2 (VLSM.l ?item) = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
l: label MC_composite_vlsm
Ht: input_valid_transition MC_composite_vlsm l (s, iom) (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (VLSM.l item) = init
m: MC_initial_state_prop (s i)projT2 l = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
j: index
lj: label (MCVLSM j)
Hv: MC_valid lj (s j) iom
Ht: (let (si', om') := MC_transition j lj (s j) iom in (state_update MCVLSM s j si', om')) = ( sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
m: MC_initial_state_prop (s i)lj = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
sf: state MC_composite_vlsm
iom, oom: option Message
j: index
lj: label (MCVLSM j)
Hv: MC_valid lj (s j) iom
si': State
om': option Message
Ht: (state_update MCVLSM s j si', om') = (sf, oom)
i: index
Hnoninit: ¬ MC_initial_state_prop (sf i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
m: MC_initial_state_prop (s i)lj = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
iom, oom: option Message
j: index
lj: label (MCVLSM j)
Hv: MC_valid lj (s j) iom
si': State
i: index
Hnoninit: ¬ MC_initial_state_prop (state_update MCVLSM s j si' i)
IHHtr: ¬ MC_initial_state_prop (s i) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
m: MC_initial_state_prop (s i)lj = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
iom, oom: option Message
j: index
lj: label (MCVLSM j)
Hv: MC_valid lj (s j) iom
si': State
m: MC_initial_state_prop (s j)
IHHtr: ¬ MC_initial_state_prop (s j) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
Hnoninit: ¬ MC_initial_state_prop si'lj = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
iom, oom: option Message
j: index
lj: label (MCVLSM j)
Hv: MC_valid lj (s j) iom
si': State
m: st_rs (s j) = None
IHHtr: ¬ MC_initial_state_prop (s j) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
Hnoninit: ¬ MC_initial_state_prop si'lj = initby inversion Hv; subst. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
si, s: composite_state MCVLSM
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm si s tr
iom, oom: option Message
j: index
lj: Label
st_obs0: indexSet
Hv: MC_valid lj (mkSt st_obs0 None) iom
si': State
IHHtr: ¬ MC_initial_state_prop (mkSt st_obs0 None) → ∃ item : composite_transition_item MCVLSM, item ∈ tr ∧ projT2 (l item) = init
Hnoninit: ¬ MC_initial_state_prop si'lj = initindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (l : label MC_composite_vlsm) (s s' : composite_state MCVLSM) (iom oom : option Message), input_valid_transition MC_composite_vlsm l (s, iom) (s', oom) → size (MuddyUnion s) = size (MuddyUnion s')index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (l : label MC_composite_vlsm) (s s' : composite_state MCVLSM) (iom oom : option Message), input_valid_transition MC_composite_vlsm l (s, iom) (s', oom) → size (MuddyUnion s) = size (MuddyUnion s')index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
l: label MC_composite_vlsm
s, s': composite_state MCVLSM
iom, oom: option Message
Hivt: input_valid_transition MC_composite_vlsm l (s, iom) (s', oom)size (MuddyUnion s) = size (MuddyUnion s')index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
l: label MC_composite_vlsm
s, s': composite_state MCVLSM
iom, oom: option Message
Hivt: input_valid_transition MC_composite_vlsm l (s, iom) (s', oom)MC_obs_equivalence s s'by apply Hivt. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
l: label MC_composite_vlsm
s, s': composite_state MCVLSM
iom, oom: option Message
Hivt: input_valid_transition MC_composite_vlsm l (s, iom) (s', oom)composite_transition MCVLSM l (s, iom) = (s', oom)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s1 s2 : composite_state MCVLSM) (tr : list transition_item) (i : index), finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr → st_rs (s2 i) ≠ None → consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s1 s2 : composite_state MCVLSM) (tr : list transition_item) (i : index), finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr → st_rs (s2 i) ≠ None → consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ Noneconsistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
c: composite_initial_state_prop MCVLSM s2consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
n: ¬ composite_initial_state_prop MCVLSM s2consistent s2by specialize (c i).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
c: composite_initial_state_prop MCVLSM s2consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
n: ¬ composite_initial_state_prop MCVLSM s2consistent s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
n: ¬ composite_initial_state_prop MCVLSM s2MC_non_initial_valid_state s2by apply valid_trace_last_pstate in Hfvt. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s1, s2: composite_state MCVLSM
tr: list transition_item
i: index
Hfvt: finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr
Hneq: st_rs (s2 i) ≠ None
n: ¬ composite_initial_state_prop MCVLSM s2valid_state_prop MC_composite_vlsm s2index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), valid_state_prop MC_composite_vlsm s → st_rs (s i) ≠ None → consistent sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), valid_state_prop MC_composite_vlsm s → st_rs (s i) ≠ None → consistent sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hv: valid_state_prop MC_composite_vlsm s
Hneq: st_rs (s i) ≠ Noneconsistent sby constructor. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hv: valid_state_prop MC_composite_vlsm s
Hneq: st_rs (s i) ≠ Nonefinite_valid_trace_from_to MC_composite_vlsm ?s1 s ?trindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index) (r : nat) (status : ChildStatus), MC_composite_invariant s → MC_constraint (existT i receive) (s, Some (mkMsg i r status)) → MC_component_invariant_helper (mkSt (st_obs (s i)) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index) (r : nat) (status : ChildStatus), MC_composite_invariant s → MC_constraint (existT i receive) (s, Some (mkMsg i r status)) → MC_component_invariant_helper (mkSt (st_obs (s i)) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
status: ChildStatus
Hinvs: MC_composite_invariant s
Hc: match s i with | mkSt _ (Some (mkRS r1 status0)) => status = status0 ∧ r = r1 ∨ status = undecided ∧ r < r1 | mkSt _ None => False endMC_component_invariant_helper (mkSt (st_obs (s i)) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
status: ChildStatus
Hinvs: MC_composite_invariant s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsj: s i = mkSt oj (Some (mkRS rj statusj))
Hc: status = statusj ∧ r = rj ∨ status = undecided ∧ r < rjMC_component_invariant_helper (mkSt (st_obs (mkSt oj (Some (mkRS rj statusj)))) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
status: ChildStatus
Hinvs: MC_composite_invariant s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsj: s i = mkSt oj (Some (mkRS rj statusj))
Hc: status = statusj ∧ r = rj ∨ status = undecided ∧ r < rj
Hinvsj: MC_component_invariant s iMC_component_invariant_helper (mkSt (st_obs (mkSt oj (Some (mkRS rj statusj)))) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
status: ChildStatus
Hinvs: MC_composite_invariant s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsj: s i = mkSt oj (Some (mkRS rj statusj))
Hc: status = statusj ∧ r = rj ∨ status = undecided ∧ r < rj
Hinvsj: match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i)) endMC_component_invariant_helper (mkSt (st_obs (mkSt oj (Some (mkRS rj statusj)))) (Some (mkRS r status))) (MuddyUnion s)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
status: ChildStatus
Hinvs: MC_composite_invariant s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsj: s i = mkSt oj (Some (mkRS rj statusj))
Hc: status = statusj ∧ r = rj ∨ status = undecided ∧ r < rj
Hinvsj: match statusj with | undecided => rj < size oj | muddy => rj = size oj ∧ size (MuddyUnion s) - 1 = size oj | clean => rj = size oj ∧ size (MuddyUnion s) = size oj endMC_component_invariant_helper (mkSt oj (Some (mkRS r status))) (MuddyUnion s)by destruct statusj; cbn; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
r: nat
Hinvs: MC_composite_invariant s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsj: s i = mkSt oj (Some (mkRS rj statusj))
H10: r < rj
Hinvsj: match statusj with | undecided => rj < size oj | muddy => rj = size oj ∧ size (MuddyUnion s) - 1 = size oj | clean => rj = size oj ∧ size (MuddyUnion s) = size oj endMC_component_invariant_helper (mkSt oj (Some (mkRS r undecided))) (MuddyUnion s)
The proof of the following lemma proceeds by induction on the length of the trace
to the state s (such a trace exists because of the assumption that the state s
is valid). The induction step analyzes each case of the transition function
and proves the corresponding relations depending on the particular
conditions of the transition.
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → MC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → MC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvsp: valid_state_prop MC_composite_vlsm sMC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
is: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm is s trMC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
is: state MC_composite_vlsm
tr: list transition_item
Htr: finite_valid_trace_init_to MC_composite_vlsm is s tr
len_tr: nat
Heqlen_tr: len_tr = length trMC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
len_tr: nat∀ (s : composite_state MCVLSM) (tr : list transition_item), len_tr = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
len_tr: nat
Hind: ∀ y : nat, y < len_tr → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s∀ (s : composite_state MCVLSM) (tr : list transition_item), len_tr = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr: list transition_item
Hind: ∀ y : nat, y < length tr → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
Htr: finite_valid_trace_init_to MC_composite_vlsm is s trMC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm is s (tr' ++ [lst])
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]MC_composite_invariant sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm is s (tr' ++ [lst])
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: indexinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
Htr': finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is tr') tr'
Hlst: finite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is tr') s [lst]
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: indexinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
s': state MC_composite_vlsm
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: finite_valid_trace_from_to MC_composite_vlsm s' s [lst]
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: indexinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
s': state MC_composite_vlsm
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: finite_valid_trace_from_to MC_composite_vlsm s' s [lst]
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: index
Heqs: finite_trace_last s' [lst] = sinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
s': state MC_composite_vlsm
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: input_valid_transition MC_composite_vlsm (l lst) (s', input lst) (destination lst, output lst)
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: index
Heqs: finite_trace_last s' [lst] = sinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: state MC_composite_vlsm
tr, tr': list transition_item
lst: transition_item
Hind: ∀ y : nat, y < length (tr' ++ [lst]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s: composite_state MCVLSM
s': state MC_composite_vlsm
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: input_valid_transition MC_composite_vlsm (l lst) (s', input lst) ( s, output lst)
Hinit: initial_state_prop is
Htr_lst: tr = tr' ++ [lst]
i: index
Heqs: destination lst = sinitial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
l: composite_label MCVLSM
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := l; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s, s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: input_valid_transition MC_composite_vlsm l (s', input) (s, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := l; input := input; destination := destination; output := output |}]
i: index
Heqs: destination = sMC_initial_state_prop (s i) ∨ MC_component_invariant s iindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
l: composite_label MCVLSM
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := l; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s, s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: input_valid_transition MC_composite_vlsm l (s', input) (s, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := l; input := input; destination := destination; output := output |}]
i: index
Heqs: destination = sMC_initial_state_prop (s i) ∨ match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
l: composite_label MCVLSM
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := l; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s, s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hlst: input_valid_transition MC_composite_vlsm l (s', input) (s, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := l; input := input; destination := destination; output := output |}]
i: index
Heqs: destination = sMC_initial_state_prop (s i) ∨ match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') = size (st_obs (s i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s, s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
Ht: (let (si', om') := MC_transition j lj (s' j) input in (state_update MCVLSM s' j si', om')) = (s, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Heqs: destination = sMC_initial_state_prop (s i) ∨ match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') = size (st_obs (s i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s, s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
o: option Message
Htj: MC_transition j lj (s' j) input = (s0, o)
Ht: (state_update MCVLSM s' j s0, o) = (s, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Heqs: destination = sMC_initial_state_prop (s i) ∨ match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s') = size (st_obs (s i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destinationMC_initial_state_prop (state_update MCVLSM s' j s0 i) ∨ match state_status (st_rs (state_update MCVLSM s' j s0 i)) with | undecided => state_round (st_rs (state_update MCVLSM s' j s0 i)) < size (st_obs (state_update MCVLSM s' j s0 i)) | muddy => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (state_update MCVLSM s' j s0 i)) | clean => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') = size (st_obs (state_update MCVLSM s' j s0 i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destinationMC_composite_invariant s'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destination
Hinvs: MC_composite_invariant s'MC_initial_state_prop (state_update MCVLSM s' j s0 i) ∨ match state_status (st_rs (state_update MCVLSM s' j s0 i)) with | undecided => state_round (st_rs (state_update MCVLSM s' j s0 i)) < size (st_obs (state_update MCVLSM s' j s0 i)) | muddy => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (state_update MCVLSM s' j s0 i)) | clean => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') = size (st_obs (state_update MCVLSM s' j s0 i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destinationMC_composite_invariant s'by rewrite app_length; cbn; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destinationlength tr' < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
j: index
lj: label (MCVLSM j)
input: option Message
destination: composite_state MCVLSM
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
s': composite_state MCVLSM
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: valid (existT j lj) (s', input)
Hc: MC_constraint (existT j lj) (s', input)
s0: State
Ht: (state_update MCVLSM s' j s0, output) = (destination, output)
Htj: MC_transition j lj (s' j) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT j lj; input := input; destination := destination; output := output |}]
i: index
Hdest: state_update MCVLSM s' j s0 = destination
Hinvs: MC_composite_invariant s'MC_initial_state_prop (state_update MCVLSM s' j s0 i) ∨ match state_status (st_rs (state_update MCVLSM s' j s0 i)) with | undecided => state_round (st_rs (state_update MCVLSM s' j s0 i)) < size (st_obs (state_update MCVLSM s' j s0 i)) | muddy => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') - 1 = size (st_obs (state_update MCVLSM s' j s0 i)) | clean => state_round (st_rs (state_update MCVLSM s' j s0 i)) = size (st_obs (state_update MCVLSM s' j s0 i)) ∧ size (MuddyUnion s') = size (st_obs (state_update MCVLSM s' j s0 i)) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr, tr': list transition_item
i: index
destination, s': composite_state MCVLSM
s0: State
Hdest: state_update MCVLSM s' i s0 = destination
lj: label (MCVLSM i)
input, output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Heqs': s' = finite_trace_last is tr'
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hc: MC_constraint (existT i lj) (s', input)
Hv: valid (existT i lj) (s', input)
Htj: MC_transition i lj (s' i) input = (s0, output)
Ht: (state_update MCVLSM s' i s0, output) = (destination, output)
Hinit: composite_initial_state_prop MCVLSM is
Htr_lst: tr = tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]
Hinvs: MC_composite_invariant s'MC_initial_state_prop s0 ∨ match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr': list transition_item
i: index
destination, s': composite_state MCVLSM
s0: State
Hdest: state_update MCVLSM s' i s0 = destination
lj: label (MCVLSM i)
input, output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hc: MC_constraint (existT i lj) (s', input)
Hv: valid (existT i lj) (s', input)
Htj: MC_transition i lj (s' i) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_composite_invariant s'MC_initial_state_prop s0 ∨ match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr': list transition_item
i: index
destination, s': composite_state MCVLSM
s0: State
Hdest: state_update MCVLSM s' i s0 = destination
lj: label (MCVLSM i)
input, output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hc: MC_constraint (existT i lj) (s', input)
Hv: valid (existT i lj) (s', input)
Htj: MC_transition i lj (s' i) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_composite_invariant s'match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr': list transition_item
i: index
destination, s': composite_state MCVLSM
s0: State
Hdest: state_update MCVLSM s' i s0 = destination
lj: label (MCVLSM i)
input, output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hc: MC_constraint (existT i lj) (s', input)
Hv: valid (existT i lj) (s', input)
Htj: MC_transition i lj (s' i) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs, Hinvs': MC_composite_invariant s'match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
tr': list transition_item
i: index
destination, s': composite_state MCVLSM
s0: State
Hdest: state_update MCVLSM s' i s0 = destination
lj: label (MCVLSM i)
input, output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i lj; input := input; destination := destination; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hc: MC_constraint (existT i lj) (s', input)
Hv: valid (existT i lj) (s', input)
Htj: MC_transition i lj (s' i) input = (s0, output)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
status: ChildStatus
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i emit; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i emit (s' i) None = (s0, output)
Hv: valid (existT i emit) (s', None)
Hc: MC_constraint (existT i emit) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o0: indexSet
Heqcall: (mkSt o0 (Some (mkRS r status)), Some (mkMsg i r status)) = ( s0, output)
H10: mkSt o0 (Some (mkRS r status)) = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
m: Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some m) = (s0, output)
Hv: valid (existT i receive) (s', Some m)
Hc: MC_constraint (existT i receive) (s', Some m)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o: indexSet
Heqcall: (mkSt o (Some (mkRS rs_round0 muddy)), None) = (s0, output)
H10: mkSt o (Some (mkRS rs_round0 muddy)) = s' imatch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
m: Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some m) = (s0, output)
Hv: valid (existT i receive) (s', Some m)
Hc: MC_constraint (existT i receive) (s', Some m)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o: indexSet
Heqcall: (mkSt o (Some (mkRS rs_round0 clean)), None) = (s0, output)
H10: mkSt o (Some (mkRS rs_round0 clean)) = s' imatch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
s0: State
output: option Message
Heqcall: (mkSt o0 (Some (mkRS 0 muddy)), None) = (s0, output)
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i init (s' i) None = (s0, output)
Hv: valid (existT i init) (s', None)
Hc: MC_constraint (existT i init) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
n: nat
o0: indexSet
s': composite_state MCVLSM
s0: State
output: option Message
Heqcall: (mkSt o0 (Some (mkRS 0 undecided)), None) = (s0, output)
Heq: size o0 = S n
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i init (s' i) None = (s0, output)
Hv: valid (existT i init) (s', None)
Hc: MC_constraint (existT i init) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (mkSt o0 (Some (mkRS r undecided)), None) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (mkSt o0 (Some (mkRS r undecided)), None) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
status: ChildStatus
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i emit; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i emit (s' i) None = (s0, output)
Hv: valid (existT i emit) (s', None)
Hc: MC_constraint (existT i emit) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o0: indexSet
Heqcall: (mkSt o0 (Some (mkRS r status)), Some (mkMsg i r status)) = ( s0, output)
H10: mkSt o0 (Some (mkRS r status)) = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
status: ChildStatus
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
o0: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i emit; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r status))); output := Some (mkMsg i r status) |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid emit (s' i) None
Hc: True
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o0 (Some (mkRS r status)) = s' i
H12: None = Nonematch status with | undecided => r < size o0 | muddy => r = size o0 ∧ size (MuddyUnion s') - 1 = size o0 | clean => r = size o0 ∧ size (MuddyUnion s') = size o0 endby destruct Hinvs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
status: ChildStatus
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
o0: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i emit; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r status))); output := Some (mkMsg i r status) |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid emit (s' i) None
Hc: True
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r status))) ∨ MC_component_invariant_helper (mkSt o0 (Some (mkRS r status))) (MuddyUnion s')
Hinvs': MC_composite_invariant s'
H10: mkSt o0 (Some (mkRS r status)) = s' i
H12: None = Nonematch status with | undecided => r < size o0 | muddy => r = size o0 ∧ size (MuddyUnion s') - 1 = size o0 | clean => r = size o0 ∧ size (MuddyUnion s') = size o0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
m: Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some m) = (s0, output)
Hv: valid (existT i receive) (s', Some m)
Hc: MC_constraint (existT i receive) (s', Some m)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o: indexSet
Heqcall: (mkSt o (Some (mkRS rs_round0 muddy)), None) = (s0, output)
H10: mkSt o (Some (mkRS rs_round0 muddy)) = s' imatch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
m: Message
o: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i (mkSt o (Some (mkRS rs_round0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some m)
Hc: MC_no_equivocation s' m
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o (Some (mkRS rs_round0 muddy)) = s' irs_round0 = size o ∧ size (MuddyUnion s') - 1 = size oby destruct Hinvs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
m: Message
o: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i (mkSt o (Some (mkRS rs_round0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some m)
Hc: MC_no_equivocation s' m
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o (Some (mkRS rs_round0 muddy))) ∨ rs_round0 = size o ∧ size (MuddyUnion s') - 1 = size o
Hinvs': MC_composite_invariant s'
H10: mkSt o (Some (mkRS rs_round0 muddy)) = s' irs_round0 = size o ∧ size (MuddyUnion s') - 1 = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
m: Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some m) = (s0, output)
Hv: valid (existT i receive) (s', Some m)
Hc: MC_constraint (existT i receive) (s', Some m)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
o: indexSet
Heqcall: (mkSt o (Some (mkRS rs_round0 clean)), None) = (s0, output)
H10: mkSt o (Some (mkRS rs_round0 clean)) = s' imatch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
m: Message
o: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i (mkSt o (Some (mkRS rs_round0 clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some m)
Hc: MC_no_equivocation s' m
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o (Some (mkRS rs_round0 clean)) = s' irs_round0 = size o ∧ size (MuddyUnion s') = size oby destruct Hinvs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
rs_round0: nat
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
m: Message
o: indexSet
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some m; destination := state_update MCVLSM s' i (mkSt o (Some (mkRS rs_round0 clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some m)
Hc: MC_no_equivocation s' m
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o (Some (mkRS rs_round0 clean))) ∨ rs_round0 = size o ∧ size (MuddyUnion s') = size o
Hinvs': MC_composite_invariant s'
H10: mkSt o (Some (mkRS rs_round0 clean)) = s' irs_round0 = size o ∧ size (MuddyUnion s') = size oindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
s0: State
output: option Message
Heqcall: (mkSt o0 (Some (mkRS 0 muddy)), None) = (s0, output)
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i init (s' i) None = (s0, output)
Hv: valid (existT i init) (s', None)
Hc: MC_constraint (existT i init) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hc: consistent s'
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = None0 = size o0 ∧ size (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hc: consistent s'
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonesize (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hunion: MuddyUnion s' ≢ ∅
Hc: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonesize (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hunion: MuddyUnion s' ≢ ∅
Hc: st_obs (mkSt o0 None) ≡ MuddyUnion s' ∖ {[i]}
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonesize (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hunion: MuddyUnion s' ≢ ∅
Hc: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = None{[i]} ⊆ MuddyUnion s'by apply size_empty_iff in Heq; rewrite Heq in Hc; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o0: indexSet
s': composite_state MCVLSM
Heq: size o0 = 0
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS 0 muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid init (s' i) None
Hunion: MuddyUnion s' ≢ ∅
Hc: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinit: composite_initial_state_prop MCVLSM is
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonei ∈ MuddyUnion s'by inversion Heqcall; subst; cbn in *; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
n: nat
o0: indexSet
s': composite_state MCVLSM
s0: State
output: option Message
Heqcall: (mkSt o0 (Some (mkRS 0 undecided)), None) = (s0, output)
Heq: size o0 = S n
is: composite_state MCVLSM
tr': list transition_item
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i init; input := None; destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i init (s' i) None = (s0, output)
Hv: valid (existT i init) (s', None)
Hc: MC_constraint (existT i init) (s', None)
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
H10: mkSt o0 None = s' i
H12: None = Nonematch state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: st_obs (s' i) ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' undecided))) (MuddyUnion s')match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' < r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r <= r' < size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' undecided))) (MuddyUnion s')
H9: o0 ≡ st_obs (s' i)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endby eapply MC_composite_invariant_preservation_muddy_from_undecided.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H15: j ∈ o0
Heq: left H15 = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ size o0 - 1 = r ∨ undecided = undecided ∧ size o0 - 1 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0 - 1) undecided))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0 - 1) undecided)) = (mkSt o0 (Some (mkRS (size o0 - 1 + 1) muddy)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0 - 1) undecided); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0 - 1 + 1) muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H13: ¬ r <= size o0 - 1 < size o0 - 1
H11: ¬ size o0 - 1 < r
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0 - 1) undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: size o0 - 1 < size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)size (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: st_obs (s' i) ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' undecided))) (MuddyUnion s')match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' undecided); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' undecided)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' undecided))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' undecided))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' ≤ r) then (mkSt o0 (Some (mkRS r undecided)), None) else if decide (r < r' < size o0) then (mkSt o0 (Some (mkRS r' undecided)), None) else if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' undecided))) (MuddyUnion s')
H9: o0 ≡ st_obs (s' i)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endby eapply MC_composite_invariant_preservation_muddy_from_clean.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H15: j ∉ o0
Heq: right H15 = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ size o0 = r ∨ undecided = undecided ∧ size o0 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0) undecided))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0) undecided)) = (mkSt o0 (Some (mkRS (size o0) muddy)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0) undecided); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0) muddy))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H13: ¬ r < size o0 < size o0
H11: ¬ size o0 ≤ r
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0) undecided)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: size o0 < size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)size (MuddyUnion s') - 1 = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: st_obs (s' i) ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' muddy))) (MuddyUnion s')match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' muddy)), None) else if decide (r' = size o0 - 1) then (mkSt o0 (Some (mkRS (r' + 1) clean)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' muddy))) (MuddyUnion s')
H9: o0 ≡ st_obs (s' i)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H14: j ∈ o0
Heq: left H14 = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ size o0 - 1 = r ∨ muddy = undecided ∧ size o0 - 1 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0 - 1) muddy))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0 - 1) muddy)) = (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0 - 1) muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H11: size o0 - 1 ≠ size o0
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0 - 1) muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: size o0 - 1 = size (st_obs (s' j)) ∧ size (MuddyUnion s') - 1 = size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)size (MuddyUnion s') = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H14: j ∈ o0
Heq: left H14 = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ size o0 - 1 = r ∨ muddy = undecided ∧ size o0 - 1 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0 - 1) muddy))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0 - 1) muddy)) = (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0 - 1) muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H11: size o0 - 1 ≠ size o0
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0 - 1) muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hstobs: size o0 - 1 = size (st_obs (s' j))
Hmuddy: size (MuddyUnion s') - 1 = size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)size (MuddyUnion s') = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H14: j ∈ o0
Heq: left H14 = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ size o0 - 1 = r ∨ muddy = undecided ∧ size o0 - 1 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0 - 1) muddy))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0 - 1) muddy)) = (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0 - 1) muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H11: size o0 - 1 ≠ size o0
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0 - 1) muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hstobs: size o0 - 1 = size (MuddyUnion s') - 1
Hmuddy: size (MuddyUnion s') - 1 = size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)size (MuddyUnion s') = size o0by apply size_non_empty_iff in HMuddy_s'.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin, H14: j ∈ o0
Heq: left H14 = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ size o0 - 1 = r ∨ muddy = undecided ∧ size o0 - 1 < r | mkSt _ None => False end
Hv: MC_valid receive (s' i) (Some (mkMsg j (size o0 - 1) muddy))
Htj: MC_transition i receive (s' i) (Some (mkMsg j (size o0 - 1) muddy)) = (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean)), None)
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j (size o0 - 1) muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS (size o0 - 1 + 1) clean))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
H11: size o0 - 1 ≠ size o0
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j (size o0 - 1) muddy)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hstobs: size o0 - 1 = size (MuddyUnion s') - 1
Hmuddy: size (MuddyUnion s') - 1 = size (st_obs (s' j))
H9: o0 ≡ st_obs (s' i)
e: size (MuddyUnion s') = 0size (MuddyUnion s') = size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' muddy)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' muddy))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' muddy))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (mkSt o0 (Some (mkRS r undecided)), None) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r undecided))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some (mkMsg j r' muddy))
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ r' = r ∨ muddy = undecided ∧ r' < r | mkSt _ None => False end
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)r < size o0by destruct Hinvs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' muddy); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r undecided))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some (mkMsg j r' muddy))
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ r' = r ∨ muddy = undecided ∧ r' < r | mkSt _ None => False end
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' muddy)r < size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (mkSt o0 (Some (mkRS r undecided)), None) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r undecided))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some (mkMsg j r' clean))
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => clean = status ∧ r' = r ∨ clean = undecided ∧ r' < r | mkSt _ None => False end
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)r < size o0by destruct Hinvs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∈ o0
r': nat
Heq: decide (j ∈ o0) = left Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i (mkSt o0 (Some (mkRS r undecided))); output := None |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Hv: MC_valid receive (s' i) (Some (mkMsg j r' clean))
Hc: match s' j with | mkSt _ (Some (mkRS r status)) => clean = status ∧ r' = r ∨ clean = undecided ∧ r' < r | mkSt _ None => False end
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)r < size o0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: initial_state_prop (s' i) ∨ MC_component_invariant s' i
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
Hinvs: MC_initial_state_prop (mkSt o0 (Some (mkRS r undecided))) ∨ r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htr': finite_valid_trace_from_to MC_composite_vlsm is s' tr'
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: st_obs (s' i) ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' clean))) (MuddyUnion s')match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endby repeat case_decide; inversion Heqcall; subst; cbn in *; clear Heqcall; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
r: nat
j: index
o0: indexSet
Hin: j ∉ o0
r': nat
Heq: decide (j ∈ o0) = right Hin
is: composite_state MCVLSM
tr': list transition_item
s': composite_state MCVLSM
s0: State
output: option Message
Hind: ∀ y : nat, y < length (tr' ++ [{| l := existT i receive; input := Some (mkMsg j r' clean); destination := state_update MCVLSM s' i s0; output := output |}]) → ∀ (s : composite_state MCVLSM) (tr : list transition_item), y = length tr → finite_valid_trace_init_to MC_composite_vlsm is s tr → MC_composite_invariant s
Htj: MC_transition i receive (s' i) (Some (mkMsg j r' clean)) = (s0, output)
Hv: valid (existT i receive) (s', Some (mkMsg j r' clean))
Hc: MC_constraint (existT i receive) (s', Some (mkMsg j r' clean))
Hinit: composite_initial_state_prop MCVLSM is
H0: r < size o0
Hinvs': MC_composite_invariant s'
Heqcall: (if decide (r' = size o0) then (mkSt o0 (Some (mkRS r' clean)), None) else if decide (r' = size o0 + 1) then (mkSt o0 (Some (mkRS (r' - 1) muddy)), None) else (mkSt o0 (Some (mkRS r undecided)), None)) = (s0, output)
m: Message
H10: mkSt o0 (Some (mkRS r undecided)) = s' i
H12: Some m = Some (mkMsg j r' clean)
HMuddy_s': MuddyUnion s' ≢ ∅
Hconsistent: ∀ n : index, st_obs (s' n) ≡ MuddyUnion s' ∖ {[n]}
Hconsi: o0 ≡ MuddyUnion s' ∖ {[i]}
Hinvs: MC_component_invariant_helper (mkSt (st_obs (s' j)) (Some (mkRS r' clean))) (MuddyUnion s')
H9: o0 ≡ st_obs (s' i)match state_status (st_rs s0) with | undecided => state_round (st_rs s0) < size (st_obs s0) | muddy => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') - 1 = size (st_obs s0) | clean => state_round (st_rs s0) = size (st_obs s0) ∧ size (MuddyUnion s') = size (st_obs s0) endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → MC_composite_invariant_inductive sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → MC_composite_invariant_inductive sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hv: valid_state_prop MC_composite_vlsm s
i: indexinitial_state_prop (s i) ∨ MC_component_invariant_inductive s iby right; apply MC_component_invariant_equiv_MC_component_invariant_inductive. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hv: valid_state_prop MC_composite_vlsm s
i: index
H9: MC_component_invariant s iinitial_state_prop (s i) ∨ MC_component_invariant_inductive s i
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Messagevalid_state_prop MC_composite_vlsm s → msg_status m = muddy → MC_no_equivocation s m → msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Messagevalid_state_prop MC_composite_vlsm s → msg_status m = muddy → MC_no_equivocation s m → msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: valid_state_prop MC_composite_vlsm s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s mmsg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: valid_state_prop MC_composite_vlsm s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm smsg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm smsg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm s
Hinit: initial_state_prop (s (msg_index m))msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm s
Hinvariant: MC_component_invariant s (msg_index m)msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm s
Hinit: initial_state_prop (s (msg_index m))msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: match m with | mkMsg n r' status' => match s n with | mkSt _ (Some (mkRS r status)) => status' = status ∧ r' = r ∨ status' = undecided ∧ r' < r | mkSt _ None => False end end
Hs': valid_state_prop MC_composite_vlsm s
Hinit: initial_state_prop (s (msg_index m))msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
H9: m = mkMsg msg_index0 msg_round0 msg_status0
Hmuddy: msg_status0 = muddy
st_obs0: indexSet
st_rs0: option RoundStatus
r: RoundStatus
rs_round0: nat
rs_status0: ChildStatus
H12: r = mkRS rs_round0 rs_status0
H11: st_rs0 = Some (mkRS rs_round0 rs_status0)
H10: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
Hnoequiv: msg_status0 = rs_status0 ∧ msg_round0 = rs_round0 ∨ msg_status0 = undecided ∧ msg_round0 < rs_round0
Hs': valid_state_prop MC_composite_vlsm s
Hinit: MC_initial_state_prop (s msg_index0)msg_round0 = size (MuddyUnion s) - 1 ∧ msg_index0 ∈ MuddyUnion sby rewrite H10 in Hinit.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
H9: m = mkMsg msg_index0 msg_round0 msg_status0
Hmuddy: msg_status0 = muddy
st_obs0: indexSet
st_rs0: option RoundStatus
r: RoundStatus
rs_round0: nat
rs_status0: ChildStatus
H12: r = mkRS rs_round0 rs_status0
H11: st_rs0 = Some (mkRS rs_round0 rs_status0)
H10: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
Hnoequiv: msg_status0 = rs_status0 ∧ msg_round0 = rs_round0 ∨ msg_status0 = undecided ∧ msg_round0 < rs_round0
Hs': valid_state_prop MC_composite_vlsm s
Hinit: st_rs (s msg_index0) = Nonemsg_round0 = size (MuddyUnion s) - 1 ∧ msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm s
Hinvariant: MC_component_invariant s (msg_index m)msg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: MC_no_equivocation s m
Hs': valid_state_prop MC_composite_vlsm s
Hinvariant: match state_status (st_rs (s (msg_index m))) with | undecided => state_round (st_rs (s (msg_index m))) < size (st_obs (s (msg_index m))) | muddy => state_round (st_rs (s (msg_index m))) = size (st_obs (s (msg_index m))) ∧ size (MuddyUnion s) - 1 = size (st_obs (s (msg_index m))) | clean => state_round (st_rs (s (msg_index m))) = size (st_obs (s (msg_index m))) ∧ size (MuddyUnion s) = size (st_obs (s (msg_index m))) endmsg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: MC_composite_invariant s
Hmuddy: msg_status m = muddy
Hnoequiv: match m with | mkMsg n r' status' => match s n with | mkSt _ (Some (mkRS r status)) => status' = status ∧ r' = r ∨ status' = undecided ∧ r' < r | mkSt _ None => False end end
Hs': valid_state_prop MC_composite_vlsm s
Hinvariant: match state_status (st_rs (s (msg_index m))) with | undecided => state_round (st_rs (s (msg_index m))) < size (st_obs (s (msg_index m))) | muddy => state_round (st_rs (s (msg_index m))) = size (st_obs (s (msg_index m))) ∧ size (MuddyUnion s) - 1 = size (st_obs (s (msg_index m))) | clean => state_round (st_rs (s (msg_index m))) = size (st_obs (s (msg_index m))) ∧ size (MuddyUnion s) = size (st_obs (s (msg_index m))) endmsg_round m = size (MuddyUnion s) - 1 ∧ msg_index m ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
Hs': valid_state_prop MC_composite_vlsm s
H9: muddy = muddy
Hinvariant: state_round (st_rs (s msg_index0)) = size (st_obs (s msg_index0)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s msg_index0))rs_round0 = size (MuddyUnion s) - 1 ∧ msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
Hs': valid_state_prop MC_composite_vlsm s
H9: muddy = muddy
Hinvariant: rs_round0 = size st_obs0 ∧ size (MuddyUnion s) - 1 = size st_obs0rs_round0 = size (MuddyUnion s) - 1 ∧ msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
Hs': valid_state_prop MC_composite_vlsm s
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size st_obs0rs_round0 = size (MuddyUnion s) - 1 ∧ msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
Hs': valid_state_prop MC_composite_vlsm s
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size st_obs0msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size st_obs0
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (st_obs (s msg_index0))
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1
Hintersectleq1: MuddyUnion s ∩ {[msg_index0]} ⊆ {[msg_index0]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1
Hintersectleq1: size (MuddyUnion s ∩ {[msg_index0]}) ≤ size {[msg_index0]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1
Hintersectleq1: size (MuddyUnion s ∩ {[msg_index0]}) ≤ 1msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1
Hintersectleq1: size (MuddyUnion s ∩ {[msg_index0]}) ≤ 1
Hintersecteq1: size (MuddyUnion s ∩ {[msg_index0]}) = 1msg_index0 ∈ MuddyUnion sby set_solver. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_composite_invariant s
msg_index0: index
st_obs0: indexSet
rs_round0: nat
H11: s msg_index0 = mkSt st_obs0 (Some (mkRS rs_round0 muddy))
H9: muddy = muddy
Hn: rs_round0 = size st_obs0
Hstobs: size (MuddyUnion s) - 1 = size (MuddyUnion s) - size (MuddyUnion s ∩ {[msg_index0]})
Hnempty: size (MuddyUnion s) ≠ 0
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hintersectgeq1: size (MuddyUnion s ∩ {[msg_index0]}) ≥ 1
Hintersectleq1: size (MuddyUnion s ∩ {[msg_index0]}) ≤ 1
Hintersecteq1: ∃ x : index, MuddyUnion s ∩ {[msg_index0]} ≡ {[x]}msg_index0 ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: indexvalid_state_prop MC_composite_vlsm s → ¬ MC_initial_state_prop (s i) → state_status (st_rs (s i)) = undecided → state_round (st_rs (s i)) < size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: indexvalid_state_prop MC_composite_vlsm s → ¬ MC_initial_state_prop (s i) → state_status (st_rs (s i)) = undecided → state_round (st_rs (s i)) < size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hvalid: valid_state_prop MC_composite_vlsm s
Hundecided: ¬ MC_initial_state_prop (s i)state_status (st_rs (s i)) = undecided → state_round (st_rs (s i)) < size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hvalid: MC_composite_invariant_inductive s
Hundecided: ¬ MC_initial_state_prop (s i)state_status (st_rs (s i)) = undecided → state_round (st_rs (s i)) < size (st_obs (s i))by inversion H9; [| congruence ..]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hvalid: MC_composite_invariant_inductive s
Hundecided: ¬ MC_initial_state_prop (s i)
H9: MC_component_invariant_inductive s istate_status (st_rs (s i)) = undecided → state_round (st_rs (s i)) < size (st_obs (s i))
The definitions MC_build_muddy_muddy_trace and MC_build_clean_muddy_trace are
useful in the proof of the lemma MC_build_valid_message, where, having a valid
state, we aim to obtain a valid message with status undecided, emitted earlier
on a trace leading to that state.
To achieve this, we take two indexes i and j, with child j being seen as muddy
by i, and want to build a valid trace which corresponds to an exchange of messages
between these two children.
We then use these two function definitions in order to build these traces between
two specific children.
Fixpoint MC_build_muddy_muddy_trace (is : composite_state MCVLSM) (target helper : index) (round : nat) : list (composite_transition_item MCVLSM) := match round with | 0 => let s := state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) in let item0 := Build_transition_item (T := composite_type MCVLSM) (existT helper init) None s None in let item1 := Build_transition_item (T := composite_type MCVLSM) (existT target init) None (state_update MCVLSM s target (mkSt (st_obs (s target)) (Some (mkRS 0 undecided)))) None in [item0; item1] | S n => let tr := MC_build_muddy_muddy_trace is helper target n in let s := finite_trace_last is tr in let item := Build_transition_item (T := composite_type MCVLSM) (existT target receive) (Some (mkMsg helper n undecided)) (state_update MCVLSM s target (mkSt (st_obs (s target)) (Some (mkRS (S n) undecided)))) None in tr ++ [item] end.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is target helper round) target = mkSt obs (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is target helper round) target = mkSt obs (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) target = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, List.last (map destination (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is target = mkSt obs (Some (mkRS (S round) undecided))by state_update_simpl; eexists.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) target = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, List.last (map destination (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is target = mkSt obs (Some (mkRS (S round) undecided))by state_update_simpl; eexists. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))) target = mkSt obs (Some (mkRS (S round) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nathelper ≠ target → ∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nathelper ≠ target → ∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, List.last (map destination (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is helper = mkSt obs (Some (mkRS (round - 0) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))) helper = mkSt obs (Some (mkRS (round - 0) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs (Some (mkRS (round - 0) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target
obs: indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs (Some (mkRS round undecided))∃ obs : indexSet, finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs (Some (mkRS (round - 0) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target
obs: indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs (Some (mkRS round undecided))∃ obs0 : indexSet, mkSt obs (Some (mkRS round undecided)) = mkSt obs0 (Some (mkRS (round - 0) undecided))by eexists. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target
obs: indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs (Some (mkRS round undecided))∃ obs0 : indexSet, mkSt obs (Some (mkRS round undecided)) = mkSt obs0 (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → ∀ (i : index) (obs : indexSet) (round : nat) (status : ChildStatus), s i = mkSt obs (Some (mkRS round status)) → valid_message_prop MC_composite_vlsm (mkMsg i round status)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → ∀ (i : index) (obs : indexSet) (round : nat) (status : ChildStatus), s i = mkSt obs (Some (mkRS round status)) → valid_message_prop MC_composite_vlsm (mkMsg i round status)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))valid_message_prop MC_composite_vlsm (mkMsg i round status)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))input_valid_transition MC_composite_vlsm (existT i emit) (s, None) (s, Some (mkMsg i round status))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))MC_valid emit (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))(let (si', om') := MC_transition i emit (s i) None in (state_update MCVLSM s i si', om')) = (s, Some (mkMsg i round status))by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))option_valid_message_prop MC_composite_vlsm Noneby rewrite Hsi; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))MC_valid emit (s i) Noneby rewrite !Hsi, MC_transition_equation_5, state_update_id. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: index
obs: indexSet
round: nat
status: ChildStatus
Hsi: s i = mkSt obs (Some (mkRS round status))(let (si', om') := MC_transition i emit (s i) None in (state_update MCVLSM s i si', om')) = (s, Some (mkMsg i round status))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s1 s2 : composite_state MCVLSM) (tr : list transition_item) (i : index), finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr → st_obs (s1 i) ≡ st_obs (s2 i)by intros; apply MC_in_futures_preserves_obs_equiv; exists tr. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s1 s2 : composite_state MCVLSM) (tr : list transition_item) (i : index), finite_valid_trace_from_to MC_composite_vlsm s1 s2 tr → st_obs (s1 i) ≡ st_obs (s2 i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: natcomposite_initial_state_prop MCVLSM is → consistent is → target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: natcomposite_initial_state_prop MCVLSM is → consistent is → target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: round < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is∀ target helper : index, target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
IHround: ∀ target helper : index, target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: initial_state_prop (is helper)finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: initial_state_prop (is target)finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2finite_valid_trace_init_to_alt MC_composite_vlsm is ?f (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2message_valid_transitions_from_to MC_composite_vlsm is ?f (MC_build_muddy_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition helper init (is helper) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is helper) None ∧ consistent isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition target init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2message_valid_transitions_from_to MC_composite_vlsm (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided)))) ?f []by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition helper init (is helper) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition helper init (mkSt o_helper None) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := match size o_helper with | 0 => (mkSt o_helper (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_helper (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size o_helper with | 0 => (mkSt o_helper (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_helper (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (st_obs (is helper)) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[helper]}) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o){[helper]} ⊆ MuddyUnion isby case_match; inversion Hso; subst; [lia |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)by rewrite <- elem_of_subseteq_singleton.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[helper]}) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o){[helper]} ⊆ MuddyUnion isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is helper) None ∧ consistent isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is helper) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
st_obs0: indexSet
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs0 ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (mkSt st_obs0 None) Noneby apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition target init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition target init (is target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (is target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt o_target None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := MC_transition target init (mkSt o_target None) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt o_target None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt o_target None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target s, o) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt (st_obs (is target)) None
Hinittarget: st_rs (mkSt (st_obs (is target)) None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt (st_obs (is target)) None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (st_obs (is target)) with | 0 => (mkSt (st_obs (is target)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is target)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target s, o) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt (st_obs (is target)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt (st_obs (is target)) None
Hinittarget: st_rs (mkSt (st_obs (is target)) None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt (st_obs (is target)) None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is target)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is target)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target s, o) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt (st_obs (is target)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt (st_obs (is target)) None
Hinittarget: st_rs (mkSt (st_obs (is target)) None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt (st_obs (is target)) None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[target]}) with | 0 => (mkSt (st_obs (is target)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is target)) (Some (mkRS 0 undecided)), None) end = (s, o){[target]} ⊆ MuddyUnion isby case_match; inversion Hso; subst; [lia |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt (st_obs (is target)) None
Hinittarget: st_rs (mkSt (st_obs (is target)) None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt (st_obs (is target)) None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is target)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is target)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target s, o) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt (st_obs (is target)) None)) (Some (mkRS 0 undecided))), None)by rewrite <- elem_of_subseteq_singleton.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
o_target: indexSet
Histarget: is target = mkSt (st_obs (is target)) None
Hinittarget: st_rs (mkSt (st_obs (is target)) None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt (st_obs (is target)) None) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[target]}) with | 0 => (mkSt (st_obs (is target)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is target)) (Some (mkRS 0 undecided)), None) end = (s, o){[target]} ⊆ MuddyUnion isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is target) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (is target) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
st_obs0: indexSet
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs0 ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_valid init (mkSt st_obs0 None) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))by apply MC_state_update_preserves_obs_equiv.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2MC_obs_equivalence is (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))by apply mvt_empty.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: 0 < size (MuddyUnion is) - 1
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
H9: size (MuddyUnion is) ≥ 2message_valid_transitions_from_to MC_composite_vlsm (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided)))) ?f []index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
IHround: ∀ target helper : index, target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
IHround: ∀ target helper : index, target ∈ MuddyUnion is → helper ∈ MuddyUnion is → target ≠ helper → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is target helper round)
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
IHround: helper ≠ target → round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round)
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IHround: round < size (MuddyUnion is) - 1 → finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round)finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IHround: finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round)finite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop isfinite_valid_trace MC_composite_vlsm is (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop isfinite_valid_trace_init_to MC_composite_vlsm is ?f (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop isfinite_valid_trace_from_to MC_composite_vlsm is ?f (MC_build_muddy_muddy_trace is helper target round ++ [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}])index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop isfinite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) ?f [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop isinput_valid_transition MC_composite_vlsm (existT target receive) (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round), Some (mkMsg helper round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))input_valid_transition MC_composite_vlsm (existT target receive) (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round), Some (mkMsg helper round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))input_valid_transition MC_composite_vlsm (existT target receive) (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round), Some (mkMsg helper round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))MC_valid receive (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target) (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))match finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition target receive (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target) (Some (mkMsg helper round undecided)) in (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))), None)by apply valid_trace_last_pstate in IH.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round))option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))by apply MC_valid_message_from_valid_state with (s := final) (obs := obs').index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
final: state (composite_type MCVLSM)
Heqfinal: final = finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)
IH: finite_valid_trace_from_to MC_composite_vlsm is final (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: final target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: final helper = mkSt obs' (Some (mkRS round undecided))
Hfinal: valid_state_prop MC_composite_vlsm finaloption_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))MC_valid receive (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target) (Some (mkMsg helper round undecided))by constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))MC_valid receive (mkSt obs (Some (mkRS (round - 1) undecided))) (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))match finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endby left.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))undecided = undecided ∧ round = round ∨ undecided = undecided ∧ round < roundindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition target receive (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target) (Some (mkMsg helper round undecided)) in (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt (st_obs (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition_clause_5 target obs (round - 1) helper (decide (helper ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))(let (si', om') := (if decide (helper ∈ obs) then λ r' : nat, if decide (r' < round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 <= r' < size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None) else λ r' : nat, if decide (r' ≤ round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 < r' < size obs) then (mkSt obs (Some (mkRS r' undecided)), None) else if decide (r' = size obs) then (mkSt obs (Some (mkRS r' muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None)) round in (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs(let (si', om') := (if decide (helper ∈ obs) then λ r' : nat, if decide (r' < round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 <= r' < size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None) else λ r' : nat, if decide (r' ≤ round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 < r' < size obs) then (mkSt obs (Some (mkRS r' undecided)), None) else if decide (r' = size obs) then (mkSt obs (Some (mkRS r' muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None)) round in (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs(state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (round + 1) undecided))), None) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obsround - 1 <= round < size obs - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs¬ round < round - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obshelper ∈ obsby replace (round + 1) with (S round) by lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs(state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (round + 1) undecided))), None) = (state_update MCVLSM (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) target (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obsround - 1 <= round < size obs - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obsround < size obs - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs
Hround: S round < size (MuddyUnion (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round))) - 1round < size obs - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs
Hround: S round < size (MuddyUnion is) - 1round < size obs - 1by rewrite <- Hobsequiv; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs
Hround: S round < size (MuddyUnion is) - 1
Hnrmuddy: size (st_obs (is target)) = size (MuddyUnion is) - 1round < size obs - 1by lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obs¬ round < round - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obshelper ∈ obsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obshelper ∈ st_obs (is target)by set_solver. Qed. Fixpoint MC_build_clean_muddy_trace (is : composite_state MCVLSM) (target helper : index) (round : nat) : list (composite_transition_item MCVLSM) := match round with | 0 => let s := state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) in let item0 := Build_transition_item (T := composite_type MCVLSM) (existT helper init) None s None in let item1 := Build_transition_item (T := composite_type MCVLSM) (existT target init) None (state_update MCVLSM s target (mkSt (st_obs (s target)) (Some (mkRS 0 undecided)))) None in [item0; item1] | S n => let tr := MC_build_clean_muddy_trace is target helper n in let s := finite_trace_last is tr in let item0 := Build_transition_item (T := composite_type MCVLSM) (existT helper receive) (Some (mkMsg target n undecided)) (state_update MCVLSM s helper (mkSt (st_obs (s helper)) (Some (mkRS n undecided)))) None in let item1 := Build_transition_item (T := composite_type MCVLSM) (existT target receive) (Some (mkMsg helper n undecided)) (state_update MCVLSM (destination item0) target (mkSt (st_obs (s target)) (Some (mkRS (S n) undecided)))) None in tr ++ [item0; item1] end.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
target, helper: index
Htarget: target ∈ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is) - 1
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_muddy_muddy_trace is helper target round)) (MC_build_muddy_muddy_trace is helper target round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) target = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_muddy_muddy_trace is helper target round) helper = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is target) ≡ obshelper ∈ MuddyUnion is ∖ {[target]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) target = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, List.last (map destination (MC_build_clean_muddy_trace is target helper round ++ [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is target = mkSt obs (Some (mkRS (S round) undecided))by state_update_simpl; eexists.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) target = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, List.last (map destination (MC_build_clean_muddy_trace is target helper round ++ [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is target = mkSt obs (Some (mkRS (S round) undecided))by state_update_simpl; eexists. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))) target = mkSt obs (Some (mkRS (S round) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nathelper ≠ target → ∃ obs : indexSet, finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nathelper ≠ target → ∃ obs : indexSet, finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hneq: helper ≠ target∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) helper = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, List.last (map destination (MC_build_clean_muddy_trace is target helper round ++ [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is helper = mkSt obs (Some (mkRS (round - 0) undecided))by state_update_simpl; eexists.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hneq: helper ≠ target∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))) helper = mkSt obs (Some (mkRS 0 undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, List.last (map destination (MC_build_clean_muddy_trace is target helper round ++ [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}])) is helper = mkSt obs (Some (mkRS (round - 0) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))) helper = mkSt obs (Some (mkRS (round - 0) undecided))by state_update_simpl; eexists. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hneq: helper ≠ target∃ obs : indexSet, state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))) helper = mkSt obs (Some (mkRS round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: natcomposite_initial_state_prop MCVLSM is → consistent is → target ∉ MuddyUnion is → helper ∈ MuddyUnion is → round < size (st_obs (is target)) → size (MuddyUnion is) ≥ 2 → finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: natcomposite_initial_state_prop MCVLSM is → consistent is → target ∉ MuddyUnion is → helper ∈ MuddyUnion is → round < size (st_obs (is target)) → size (MuddyUnion is) ≥ 2 → finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hround: round < size (st_obs (is target))
Hsizemuddy: size (MuddyUnion is) ≥ 2finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hround: round < size (st_obs (is target))
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helperfinite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helperround < size (st_obs (is target)) → finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
IHround: round < size (st_obs (is target)) → finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)
Hround: S round < size (st_obs (is target))finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: initial_state_prop (is helper)finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: initial_state_prop (is target)finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = Nonefinite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}finite_valid_trace_init_to_alt MC_composite_vlsm is ?f (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}message_valid_transitions_from_to MC_composite_vlsm is ?f (MC_build_clean_muddy_trace is target helper 0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition helper init (is helper) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is helper) None ∧ consistent isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition target init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}message_valid_transitions_from_to MC_composite_vlsm (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided)))) ?f []by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition helper init (is helper) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition helper init (mkSt o_helper None) None in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := match size o_helper with | 0 => (mkSt o_helper (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_helper (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM is helper si', om')) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt o_helper None
Hinithelper: st_rs (mkSt o_helper None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt o_helper None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size o_helper with | 0 => (mkSt o_helper (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_helper (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt o_helper None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size (st_obs (is helper)) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[helper]}) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o){[helper]} ⊆ MuddyUnion isby case_match; inversion Hso; subst; [lia |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size (MuddyUnion is) - 1 with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o)(state_update MCVLSM is helper s, o) = (state_update MCVLSM is helper (mkSt (st_obs (mkSt (st_obs (is helper)) None)) (Some (mkRS 0 undecided))), None)by rewrite <- elem_of_subseteq_singleton.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
o_helper: indexSet
Hishelper: is helper = mkSt (st_obs (is helper)) None
Hinithelper: st_rs (mkSt (st_obs (is helper)) None) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (mkSt (st_obs (is helper)) None) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}
s: State
o: option Message
Hso: match size (MuddyUnion is ∖ {[helper]}) with | 0 => (mkSt (st_obs (is helper)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)), None) end = (s, o){[helper]} ⊆ MuddyUnion isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is helper) None ∧ consistent isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is helper) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
st_obs0: indexSet
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs0 ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (mkSt st_obs0 None) Noneby apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition target init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition target init (is target) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (is target)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := MC_transition target init (mkSt o_target None) None in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
n: size o_target ≠ 0(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size o_target = 0(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)by cbn; repeat case_match; [lia | inversion H9].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
n: size o_target ≠ 0(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size o_target = 0(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (st_obs (is target)) = 0(let (si', om') := match size o_target with | 0 => (mkSt o_target (Some (mkRS 0 muddy)), None) | S _ => (mkSt o_target (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (mkSt o_target None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is ∖ {[target]}) = 0{[target]} ⊆ MuddyUnion isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is ∖ {[target]}) = 0Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]}) = 0Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]}) = 0
Hsizeintersect: size (MuddyUnion is ∩ {[target]}) ≥ 2Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]}) = 0
Hsizeintersect: size (MuddyUnion is ∩ {[target]}) ≥ 2
Hintersectleq1: MuddyUnion is ∩ {[target]} ⊆ {[target]}Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]}) = 0
Hsizeintersect: size (MuddyUnion is ∩ {[target]}) ≥ 2
Hintersectleq1: size (MuddyUnion is ∩ {[target]}) ≤ size {[target]}Falseby lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
o_target: indexSet
Histarget: is target = mkSt o_target None
Hround: 0 < size (st_obs (mkSt o_target None))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (mkSt o_target None) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (mkSt o_target None) ≡ MuddyUnion is ∖ {[target]}
e: size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]}) = 0
Hsizeintersect: size (MuddyUnion is ∩ {[target]}) ≥ 2
Hintersectleq1: size (MuddyUnion is ∩ {[target]}) ≤ 1Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is target) None ∧ consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is target) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_valid init (is target) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
st_obs0: indexSet
Hround: 0 < size st_obs0
Hinithelper: st_rs (is helper) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs0 ≡ MuddyUnion is ∖ {[target]}MC_valid init (mkSt st_obs0 None) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}consistent (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))by apply MC_state_update_preserves_obs_equiv.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}MC_obs_equivalence is (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))))by apply mvt_empty.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
Hinit: composite_initial_state_prop MCVLSM is
Hmuddyunion: MuddyUnion is ≢ ∅
Hconsobs: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: 0 < size (st_obs (is target))
Hinithelper: st_rs (is helper) = None
Hinittarget: st_rs (is target) = None
Hconshelper: st_obs (is helper) ≡ MuddyUnion is ∖ {[helper]}
Hconstarget: st_obs (is target) ≡ MuddyUnion is ∖ {[target]}message_valid_transitions_from_to MC_composite_vlsm (state_update MCVLSM (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided)))) target (mkSt (st_obs (state_update MCVLSM is helper (mkSt (st_obs (is helper)) (Some (mkRS 0 undecided))) target)) (Some (mkRS 0 undecided)))) ?f []index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
IHround: round < size (st_obs (is target)) → finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)
Hround: S round < size (st_obs (is target))finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IHround: finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper round)finite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop isfinite_valid_trace MC_composite_vlsm is (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop isfinite_valid_trace_init_to MC_composite_vlsm is ?f (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop isfinite_valid_trace_from_to MC_composite_vlsm is ?f (MC_build_clean_muddy_trace is target helper (S round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop isfinite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) ?f [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))finite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) ?f [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))finite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) ?f [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)finite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) ?f [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))option_valid_message_prop MC_composite_vlsm (Some (mkMsg target round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))MC_valid receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper) (Some (mkMsg target round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))match finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition helper receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper) (Some (mkMsg target round undecided)) in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)by apply valid_trace_last_pstate in IH.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))option_valid_message_prop MC_composite_vlsm (Some (mkMsg target round undecided))by eapply MC_valid_message_from_valid_state.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))option_valid_message_prop MC_composite_vlsm (Some (mkMsg target round undecided))by rewrite Hlasthelper; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))MC_valid receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper) (Some (mkMsg target round undecided))by subst; rewrite Hlast; cbn; left.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))match finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition helper receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper) (Some (mkMsg target round undecided)) in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))(let (si', om') := MC_transition_clause_5 helper obs (round - 1) target (decide (target ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs(let (si', om') := MC_transition_clause_5 helper obs (round - 1) target (decide (target ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs(let (si', om') := MC_transition_clause_5 helper obs (round - 1) target (decide (target ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obstarget ∉ obsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obs(let (si', om') := MC_transition_clause_5 helper obs (round - 1) target (decide (target ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obstarget ∉ obsby set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obstarget ∉ MuddyUnion is ∖ {[helper]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obs(let (si', om') := MC_transition_clause_5 helper obs (round - 1) target (decide (target ∈ obs)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obs(let (si', om') := (if decide (target ∈ obs) then λ r' : nat, if decide (r' < round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 <= r' < size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs - 1) then (mkSt obs (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None) else λ r' : nat, if decide (r' ≤ round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 < r' < size obs) then (mkSt obs (Some (mkRS r' undecided)), None) else if decide (r' = size obs) then (mkSt obs (Some (mkRS r' muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None)) round in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obs(let (si', om') := if decide (round ≤ round - 1) then (mkSt obs (Some (mkRS (round - 1) undecided)), None) else if decide (round - 1 < round < size obs) then (mkSt obs (Some (mkRS round undecided)), None) else if decide (round = size obs) then (mkSt obs (Some (mkRS round muddy)), None) else (mkSt obs (Some (mkRS (round - 1) undecided)), None) in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S (S round) < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) (MC_build_clean_muddy_trace is target helper (S round))
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) helper = mkSt obs (Some (mkRS (S round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) target = mkSt obs' (Some (mkRS (S round) undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obs(let (si', om') := if decide (S round ≤ S round - 1) then (mkSt obs (Some (mkRS (S round - 1) undecided)), None) else if decide (S round - 1 < S round < size obs) then (mkSt obs (Some (mkRS (S round) undecided)), None) else if decide (S round = size obs) then (mkSt obs (Some (mkRS (S round) muddy)), None) else (mkSt obs (Some (mkRS (S round - 1) undecided)), None) in (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) helper (mkSt obs (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S (S round) < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) (MC_build_clean_muddy_trace is target helper (S round))
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) helper = mkSt obs (Some (mkRS (S round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) target = mkSt obs' (Some (mkRS (S round) undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obsS round - 1 < S round < size obsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S (S round) < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) (MC_build_clean_muddy_trace is target helper (S round))
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) helper = mkSt obs (Some (mkRS (S round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) target = mkSt obs' (Some (mkRS (S round) undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obsS round < size obsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S (S round) < size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]})
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) (MC_build_clean_muddy_trace is target helper (S round))
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) helper = mkSt obs (Some (mkRS (S round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) target = mkSt obs' (Some (mkRS (S round) undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obsS round < size obsby lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S (S round) < size (MuddyUnion is) - size (MuddyUnion is ∩ {[target]})
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round))) (MC_build_clean_muddy_trace is target helper (S round))
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) helper = mkSt obs (Some (mkRS (S round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper (S round)) target = mkSt obs' (Some (mkRS (S round) undecided))
Hobsequiv: st_obs (is helper) ≡ obs
Htargetobs: target ∉ obsS round < size (MuddyUnion is) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)finite_valid_trace_from_to MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) ?f [{| l := existT helper receive; input := Some (mkMsg target round undecided); destination := state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))); output := None |}; {| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) ?f [{| l := existT target receive; input := Some (mkMsg helper round undecided); destination := state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: initial_state_prop is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)input_valid_transition MC_composite_vlsm (existT target receive) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), Some (mkMsg helper round undecided)) (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)MC_valid receive (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) target) (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)match state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)(let (si', om') := MC_transition target receive (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) target) (Some (mkMsg helper round undecided)) in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))), None)by eapply input_valid_transition_destination.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
final: ∀ j : index, state (MCVLSM j)
Heqfinal: final = state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (final, None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))option_valid_message_prop MC_composite_vlsm (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
final: ∀ j : index, state (MCVLSM j)
Heqfinal: final = state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (final, None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))input_valid_transition MC_composite_vlsm (existT helper emit) (final, None) (final, Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))valid_state_prop MC_composite_vlsm (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))MC_valid emit (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))(let (si', om') := MC_transition helper emit (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper) None in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), Some (mkMsg helper round undecided))by eapply input_valid_transition_destination.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))valid_state_prop MC_composite_vlsm (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))))by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))option_valid_message_prop MC_composite_vlsm Noneby state_update_simpl; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))MC_valid emit (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))(let (si', om') := MC_transition helper emit (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper) None in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), Some (mkMsg helper round undecided))by rewrite MC_transition_equation_5, state_update_twice.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hfinal: valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is target helper round))(let (si', om') := MC_transition helper emit (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) None in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) helper si', om')) = (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)MC_valid receive (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) target) (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)MC_valid receive (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt obs (Some (mkRS round undecided))) target) (Some (mkMsg helper round undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)MC_valid receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target) (Some (mkMsg helper round undecided))by constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)MC_valid receive (mkSt obs' (Some (mkRS round undecided))) (Some (mkMsg helper round undecided))by state_update_simpl; left.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)match state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) helper with | mkSt _ (Some (mkRS r status)) => undecided = status ∧ round = r ∨ undecided = undecided ∧ round < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)(let (si', om') := MC_transition target receive (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))) target) (Some (mkMsg helper round undecided)) in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)(let (si', om') := MC_transition target receive (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target) (Some (mkMsg helper round undecided)) in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target)) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)(let (si', om') := MC_transition_clause_5 target obs' round helper (decide (helper ∈ obs')) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)(let (si', om') := (if decide (helper ∈ obs') then λ r' : nat, if decide (r' < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= r' < size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) else λ r' : nat, if decide (r' ≤ round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round < r' < size obs') then (mkSt obs' (Some (mkRS r' undecided)), None) else if decide (r' = size obs') then (mkSt obs' (Some (mkRS r' muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None)) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hcons: consistent is
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'(let (si', om') := (if decide (helper ∈ obs') then λ r' : nat, if decide (r' < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= r' < size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) else λ r' : nat, if decide (r' ≤ round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round < r' < size obs') then (mkSt obs' (Some (mkRS r' undecided)), None) else if decide (r' = size obs') then (mkSt obs' (Some (mkRS r' muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None)) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'(let (si', om') := (if decide (helper ∈ obs') then λ r' : nat, if decide (r' < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= r' < size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) else λ r' : nat, if decide (r' ≤ round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round < r' < size obs') then (mkSt obs' (Some (mkRS r' undecided)), None) else if decide (r' = size obs') then (mkSt obs' (Some (mkRS r' muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None)) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'helper ∈ obs'index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(let (si', om') := (if decide (helper ∈ obs') then λ r' : nat, if decide (r' < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= r' < size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) else λ r' : nat, if decide (r' ≤ round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round < r' < size obs') then (mkSt obs' (Some (mkRS r' undecided)), None) else if decide (r' = size obs') then (mkSt obs' (Some (mkRS r' muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None)) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'helper ∈ obs'by set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'helper ∈ MuddyUnion is ∖ {[target]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(let (si', om') := (if decide (helper ∈ obs') then λ r' : nat, if decide (r' < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= r' < size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size obs' - 1) then (mkSt obs' (Some (mkRS (r' + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) else λ r' : nat, if decide (r' ≤ round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round < r' < size obs') then (mkSt obs' (Some (mkRS r' undecided)), None) else if decide (r' = size obs') then (mkSt obs' (Some (mkRS r' muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None)) round in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(let (si', om') := if decide (round < round) then (mkSt obs' (Some (mkRS round undecided)), None) else if decide (round <= round < size obs' - 1) then (mkSt obs' (Some (mkRS (round + 1) undecided)), None) else if decide (round = size obs' - 1) then (mkSt obs' (Some (mkRS (round + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(let (si', om') := if decide (round <= round < size obs' - 1) then (mkSt obs' (Some (mkRS (round + 1) undecided)), None) else if decide (round = size obs' - 1) then (mkSt obs' (Some (mkRS (round + 1) muddy)), None) else (mkSt obs' (Some (mkRS round undecided)), None) in (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target si', om')) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt obs' (Some (mkRS (round + 1) undecided))), None) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'round <= round < size obs' - 1by replace (round + 1) with (S round) by lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'(state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt obs' (Some (mkRS (round + 1) undecided))), None) = (state_update MCVLSM (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided)))) target (mkSt (st_obs (mkSt obs' (Some (mkRS round undecided)))) (Some (mkRS (S round) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'round <= round < size obs' - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'round < size obs' - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (st_obs (is target))
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'round < size (MuddyUnion is ∖ {[target]}) - 1by lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
is: composite_state MCVLSM
target, helper: index
round: nat
Hinit: composite_initial_state_prop MCVLSM is
Hnempty: MuddyUnion is ≢ ∅
Hcons: ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Htarget: target ∉ MuddyUnion is
Hhelper: helper ∈ MuddyUnion is
Hsizemuddy: size (MuddyUnion is) ≥ 2
Hdiff: target ≠ helper
Hround: S round < size (MuddyUnion is ∖ {[target]})
IH: finite_valid_trace_from_to MC_composite_vlsm is (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) (MC_build_clean_muddy_trace is target helper round)
IHinit: composite_initial_state_prop MCVLSM is
obs: indexSet
Hlasthelper: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper = mkSt obs (Some (mkRS (round - 1) undecided))
obs': indexSet
Hlast: finite_trace_last is (MC_build_clean_muddy_trace is target helper round) target = mkSt obs' (Some (mkRS round undecided))
Hvalidtr1: input_valid_transition MC_composite_vlsm (existT helper receive) (finite_trace_last is (MC_build_clean_muddy_trace is target helper round), Some (mkMsg target round undecided)) (state_update MCVLSM (finite_trace_last is (MC_build_clean_muddy_trace is target helper round)) helper (mkSt (st_obs (finite_trace_last is (MC_build_clean_muddy_trace is target helper round) helper)) (Some (mkRS round undecided))), None)
Hobsequiv: st_obs (is target) ≡ obs'
Htargetobs: helper ∈ obs'round < size (MuddyUnion is ∖ {[target]}) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: natMC_non_initial_valid_state s → length (enum index) > 1 → ∀ i : index, round < size (st_obs (s i)) → valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: natMC_non_initial_valid_state s → length (enum index) > 1 → ∀ i : index, round < size (st_obs (s i)) → valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
Hvalid: MC_non_initial_valid_state s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
Hvalid: MC_non_initial_valid_state s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hcons: consistent svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hcons: consistent svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hcons: consistent svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hcons: consistent s
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}MuddyUnion is ≡ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)by apply MC_in_futures_preserves_muddy; exists tr.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}MuddyUnion is ≡ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (st_obs (s i))
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion svalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s∃ j : index, j ∈ st_obs (is i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ st_obs (is i)valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s∃ j : index, j ∈ st_obs (is i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s∃ j : index, j ∈ MuddyUnion is ∖ {[i]}by lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion ssize (MuddyUnion is ∖ {[i]}) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ st_obs (is i)valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ st_obs (is i)
Hhelper: j ∈ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ jvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is i j round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))round < size (MuddyUnion is) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))
H9: finite_valid_trace_from MC_composite_vlsm is (MC_build_muddy_muddy_trace is i j round)
H10: initial_state_prop isvalid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is i j round))by rewrite size_difference, size_singleton in Hround; [| set_solver].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))round < size (MuddyUnion is) - 1by apply finite_valid_trace_last_pstate.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
e: i ∈ MuddyUnion is
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_muddy_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))
H9: finite_valid_trace_from MC_composite_vlsm is (MC_build_muddy_muddy_trace is i j round)
H10: initial_state_prop isvalid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_muddy_muddy_trace is i j round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion isvalid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))valid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is i j round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))round < size (st_obs (is i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))
H9: finite_valid_trace_from MC_composite_vlsm is (MC_build_clean_muddy_trace is i j round)
H10: initial_state_prop isvalid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is i j round))by rewrite <- Hcons' in Hround.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))round < size (st_obs (is i))by apply finite_valid_trace_last_pstate.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: ¬ size (MuddyUnion is) ≤ 1
obsf: indexSet
Hlasti: finite_trace_last is (MC_build_clean_muddy_trace is i j round) i = mkSt obsf (Some (mkRS round undecided))
H9: finite_valid_trace_from MC_composite_vlsm is (MC_build_clean_muddy_trace is i j round)
H10: initial_state_prop isvalid_state_prop MC_composite_vlsm (finite_trace_last is (MC_build_clean_muddy_trace is i j round))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid_message_prop MC_composite_vlsm (mkMsg i round undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid_message_prop MC_composite_vlsm (mkMsg i 0 undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid_state_prop MC_composite_vlsm (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1input_valid_transition MC_composite_vlsm (existT i init) (is, None) (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid (existT i init) (is, None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1MuddyUnion is ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1
n0, x: indexx ∈ st_obs (is n0) → x ∈ MuddyUnion is ∖ {[n0]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1
n0, x: indexx ∈ MuddyUnion is ∖ {[n0]} → x ∈ st_obs (is n0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1transition (existT i init) (is, None) = (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid (existT i init) (is, None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
Hinit: initial_state_prop (is i)
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1valid (existT i init) (is, None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
Hinit: MC_initial_state_prop (is i)
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1MC_valid init (is i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
Hinit: st_rs (is i) = None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1MC_valid init (is i) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1MC_valid init (mkSt st_obs0 None) Nonedone.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1MuddyUnion is ≢ ∅by cbn in *; apply Hcons'.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1
n0, x: indexx ∈ st_obs (is n0) → x ∈ MuddyUnion is ∖ {[n0]}by apply Hcons'.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1
n0, x: indexx ∈ MuddyUnion is ∖ {[n0]} → x ∈ st_obs (is n0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1transition (existT i init) (is, None) = (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
Hinit: initial_state_prop is
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
i: index
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := MC_transition i init (is i) None in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
Hinit: initial_state_prop (is i)
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := MC_transition i init (is i) None in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: state MC_composite_vlsm
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
st_rs0: option RoundStatus
Hisi: is i = mkSt st_obs0 st_rs0
Hinit: initial_state_prop (mkSt st_obs0 st_rs0)
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := MC_transition i init (mkSt st_obs0 st_rs0) None in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt (st_obs (mkSt st_obs0 st_rs0)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := MC_transition_clause_1 i st_obs0 (size st_obs0) in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := match size st_obs0 with | 0 => (mkSt st_obs0 (Some (mkRS 0 muddy)), None) | S _ => (mkSt st_obs0 (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 1(let (si', om') := match size (st_obs (is i)) with | 0 => (mkSt (st_obs (is i)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (is i)) (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM is i si', om')) = (state_update MCVLSM is i (mkSt (st_obs (is i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 11 = size (st_obs (is i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 11 = size (MuddyUnion is) - size (MuddyUnion is ∩ {[i]})index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
l: size (MuddyUnion is) ≤ 11 = size (MuddyUnion is) - size ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
Heq: size (MuddyUnion is) = 0
l: 0 ≤ 11 = 0 - size ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: nat
Heq: size (MuddyUnion is) = S n0
l: S n0 ≤ 11 = S n0 - size ∅by apply size_empty_inv in Heq.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
Heq: size (MuddyUnion is) = 0
l: 0 ≤ 11 = 0 - size ∅by rewrite size_empty; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
round: nat
is: composite_state MCVLSM
tr: list transition_item
Hfromto: finite_valid_trace_from_to MC_composite_vlsm is s tr
i: index
st_obs0: indexSet
Hisi: is i = mkSt st_obs0 None
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
Hlength: length (enum index) > 1
Hround: round < size (MuddyUnion is ∖ {[i]})
Hnemptys: MuddyUnion s ≢ ∅
Hconss: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hcons: consistent is
Hnempty: MuddyUnion is ≢ ∅
Hcons': ∀ n : index, st_obs (is n) ≡ MuddyUnion is ∖ {[n]}
Hmuddyis: MuddyUnion is ≡ MuddyUnion s
j: index
Hj: j ∈ MuddyUnion is ∖ {[i]}
Hhelper: j ∈ MuddyUnion is
Hdiff: i ≠ j
n: i ∉ MuddyUnion is
n0: nat
Heq: size (MuddyUnion is) = S n0
l: S n0 ≤ 11 = S n0 - size ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), st_rs (s i) ≠ None → ¬ composite_initial_state_prop MCVLSM sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet∀ (s : composite_state MCVLSM) (i : index), st_rs (s i) ≠ None → ¬ composite_initial_state_prop MCVLSM sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hneq: st_rs (s i) ≠ None
Hforall: composite_initial_state_prop MCVLSM sFalseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hneq: st_rs (s i) ≠ None
Hforall: Forall (λ n : index, initial_state_prop (s n)) (enum index)Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hneq: st_rs (s i) ≠ None¬ Forall (λ n : index, initial_state_prop (s n)) (enum index)by exists i; split; cbn; [apply elem_of_enum |]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
i: index
Hneq: st_rs (s i) ≠ None∃ x : index, x ∈ enum index ∧ (not ∘ (λ n : index, initial_state_prop (s n))) xindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Messagevalid_state_prop MC_composite_vlsm s → MC_no_equivocation s m → length (enum index) > 1 → valid_message_prop MC_composite_vlsm mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Messagevalid_state_prop MC_composite_vlsm s → MC_no_equivocation s m → length (enum index) > 1 → valid_message_prop MC_composite_vlsm mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: valid_state_prop MC_composite_vlsm s
Hnoequiv: MC_no_equivocation s m
Hlength: length (enum index) > 1valid_message_prop MC_composite_vlsm mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
m: Message
Hs: valid_state_prop MC_composite_vlsm s
Hnoequiv: match m with | mkMsg n r' status' => match s n with | mkSt _ (Some (mkRS r status)) => status' = status ∧ r' = r ∨ status' = undecided ∧ r' < r | mkSt _ None => False end end
Hlength: length (enum index) > 1valid_message_prop MC_composite_vlsm mindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
statusm: ChildStatus
Hs: valid_state_prop MC_composite_vlsm s
Hnoequiv: match s j with | mkSt _ (Some (mkRS r status)) => statusm = status ∧ rm = r ∨ statusm = undecided ∧ rm < r | mkSt _ None => False end
Hlength: length (enum index) > 1valid_message_prop MC_composite_vlsm (mkMsg j rm statusm)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
statusm: ChildStatus
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hnoequiv: statusm = statusj ∧ rm = rj ∨ statusm = undecided ∧ rm < rj
Hlength: length (enum index) > 1valid_message_prop MC_composite_vlsm (mkMsg j rm statusm)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1valid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1MC_non_initial_valid_state sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state svalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1MC_non_initial_valid_state sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1¬ composite_initial_state_prop MCVLSM sby rewrite Hsjf.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1st_rs (s ?i) ≠ Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state svalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hcons: consistent svalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}valid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant svalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant s
Hinvariant: MC_component_invariant s jvalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant s
Hinvariant: match state_status (st_rs (s j)) with | undecided => state_round (st_rs (s j)) < size (st_obs (s j)) | muddy => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j)) | clean => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) = size (st_obs (s j)) endvalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant s
Hinvariant: match statusj with | undecided => rj < size oj | muddy => rj = size oj ∧ size (MuddyUnion s) - 1 = size oj | clean => rj = size oj ∧ size (MuddyUnion s) = size oj endvalid_message_prop MC_composite_vlsm (mkMsg j rm undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt oj (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant s
Hinvariant: match statusj with | undecided => rj < size oj | muddy => rj = size oj ∧ size (MuddyUnion s) - 1 = size oj | clean => rj = size oj ∧ size (MuddyUnion s) = size oj endrm < size (st_obs (s j))by case_match; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
j: index
rm: nat
Hs: valid_state_prop MC_composite_vlsm s
oj: indexSet
rj: nat
statusj: ChildStatus
Hsjf: s j = mkSt (st_obs (s j)) (Some (mkRS rj statusj))
Hrm: rm < rj
Hlength: length (enum index) > 1
Hnivs: MC_non_initial_valid_state s
Hnempty: MuddyUnion s ≢ ∅
Hcons': ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hinvariants: MC_composite_invariant s
Hinvariant: match statusj with | undecided => rj < size (st_obs (s j)) | muddy => rj = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j)) | clean => rj = size (st_obs (s j)) ∧ size (MuddyUnion s) = size (st_obs (s j)) endrm < size (st_obs (s j))
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) → input_valid_transition_item MC_composite_vlsm s item → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) → input_valid_transition_item MC_composite_vlsm s item → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: message_round (input item) = size (st_obs (s i))
Hvalid: input_valid_transition_item MC_composite_vlsm s itemstate_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: message_round (input item) = size (st_obs (s i))
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Htr: MC_transition i receive (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s i))) (state_status (st_rs (s i)))))) (Some (mkMsg (msg_index (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))) (msg_round (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))) (msg_status (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))))) = (mkSt (st_obs (s i)) (Some (mkRS (msg_round (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))) muddy)), None)state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = muddy
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i))
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Htr: MC_transition i receive (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s i))) (state_status (st_rs (s i)))))) (Some (mkMsg (msg_index (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))) (msg_round (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))) (msg_status (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))))) = (mkSt (st_obs (s i)) (Some (mkRS (msg_round (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))) muddy)), None)state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x))
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input) muddy)), None)state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x))
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input) muddy)), None)state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x))
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input) muddy)), None)
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (destination x)) > state_round (st_rs (s x))by apply MC_valid_noninitial_state_undecided_round_less_obs; [| inversion Hv |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x))
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input) muddy)), None)state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x))
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input) muddy)), None)
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hsundecided: state_status (st_rs (s x)) = undecided
m: Message
Hmround: from_option msg_round 0 (Some m) = size (st_obs (s x))
Hmindex: from_option msg_index x (Some m) ∈ st_obs (s x)
Hmmuddy: from_option msg_status undecided (Some m) = muddy
Hs: valid_state_prop MC_composite_vlsm s
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x (Some m)) (from_option msg_round 0 (Some m)) (from_option msg_status undecided (Some m)))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 (Some m)) muddy)), None)
Hv: MC_valid receive (s x) (Some m)
Hc: MC_no_equivocation s m
Ht: (let (si', om') := MC_transition x receive (s x) (Some m) in (state_update MCVLSM s x si', om')) = (destination, output)
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
o: indexSet
rs: RoundStatus
H9: mkSt o (Some rs) = s xstate_round (st_rs (destination x)) > state_round (st_rs (mkSt o (Some rs)))by rewrite Htr in Ht; inversion Ht; subst; state_update_simpl; cbn; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
o: indexSet
rs_round0: nat
rs_status0: ChildStatus
Hsundecided: rs_status0 = undecided
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
Hmround: msg_round0 = size o
Hmindex: msg_index0 ∈ o
Hmmuddy: msg_status0 = muddy
Hs: valid_state_prop MC_composite_vlsm s
Htr: MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) = (mkSt o (Some (mkRS msg_round0 muddy)), None)
Hv: MC_valid receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0))
Hc: match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end
Ht: (let (si', om') := MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) in (state_update MCVLSM s x si', om')) = (destination, output)
Hrounds: rs_round0 < size ostate_round (st_rs (destination x)) > rs_round0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: message_round (input item) = size (st_obs (s i)) - 1state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: message_round (input item) = size (st_obs (s i)) - 1
Htr: MC_transition i receive (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s i))) (state_status (st_rs (s i)))))) (Some (mkMsg (msg_index (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))) (msg_round (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))) (msg_status (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item)))))) = (mkSt (st_obs (s i)) (Some (mkRS (msg_round (mkMsg (message_index (input item) i) (message_round (input item)) (message_status (input item))) + 1) clean)), None)state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = muddy
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1
Htr: MC_transition i receive (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s i))) (state_status (st_rs (s i)))))) (Some (mkMsg (msg_index (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))) (msg_round (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))) (msg_status (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})))))) = (mkSt (st_obs (s i)) (Some (mkRS (msg_round (mkMsg (message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i) (message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})) (message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}))) + 1) clean)), None)state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input + 1) clean)), None)state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input + 1) clean)), None)state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input + 1) clean)), None)
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (destination x)) > state_round (st_rs (s x))by apply MC_valid_noninitial_state_undecided_round_less_obs; [| inversion Hv |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input + 1) clean)), None)state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = muddy
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x input) (from_option msg_round 0 input) (from_option msg_status undecided input))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 input + 1) clean)), None)
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
m: Message
Hv: MC_valid receive (s x) (Some m)
Hc: MC_no_equivocation s m
Ht: (let (si', om') := MC_transition x receive (s x) (Some m) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: state_status (st_rs (s x)) = undecided
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (from_option msg_index x (Some m)) (from_option msg_round 0 (Some m)) (from_option msg_status undecided (Some m)))) = (mkSt (st_obs (s x)) (Some (mkRS (from_option msg_round 0 (Some m) + 1) clean)), None)
Hmround: from_option msg_round 0 (Some m) = size (st_obs (s x)) - 1
Hmindex: from_option msg_index x (Some m) ∈ st_obs (s x)
Hmmuddy: from_option msg_status undecided (Some m) = muddy
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
o: indexSet
rs: RoundStatus
H9: mkSt o (Some rs) = s xstate_round (st_rs (destination x)) > state_round (st_rs (mkSt o (Some rs)))by rewrite Htr in Ht; inversion Ht; subst; state_update_simpl; cbn; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
Hs: valid_state_prop MC_composite_vlsm s
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
o: indexSet
rs_round0: nat
rs_status0: ChildStatus
Hv: MC_valid receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0))
Hc: match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end
Ht: (let (si', om') := MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: rs_status0 = undecided
Htr: MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) = (mkSt o (Some (mkRS (msg_round0 + 1) clean)), None)
Hmround: msg_round0 = size o - 1
Hmindex: msg_index0 ∈ o
Hmmuddy: msg_status0 = muddy
Hrounds: rs_round0 < size ostate_round (st_rs (destination x)) > rs_round0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 ∨ message_round (input item) = size (st_obs (s i)) → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = muddy → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 ∨ message_round (input item) = size (st_obs (s i)) → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
H9: message_round (input item) = size (st_obs (s i)) - 1state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
H9: message_round (input item) = size (st_obs (s i))state_round (st_rs (destination item i)) > state_round (st_rs (s i))by apply MC_undecided_muddy_to_clean_increase_round.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
H9: message_round (input item) = size (st_obs (s i)) - 1state_round (st_rs (destination item i)) > state_round (st_rs (s i))by apply MC_undecided_muddy_to_muddy_increase_round. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = muddy
Hmindex: message_index (input item) i ∈ st_obs (s i)
H9: message_round (input item) = size (st_obs (s i))state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = undecided → message_index (input item) i ∈ st_obs (s i) → state_round (st_rs (s i)) <= message_round (input item) < size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = undecided → message_index (input item) i ∈ st_obs (s i) → state_round (st_rs (s i)) <= message_round (input item) < size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = undecided
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (input item) < size (st_obs (s i)) - 1state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) < size (st_obs (s i)) - 1state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) < size (st_obs (s i)) - 1state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) < size (st_obs (s i)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))by apply MC_valid_noninitial_state_undecided_round_less_obs; cbn in *; subst; [| inversion Hv |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) < size (st_obs (s i)) - 1state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: state_round (st_rs (s i)) <= message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) < size (st_obs (s i)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = undecided
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: state_round (st_rs (s x)) <= from_option msg_round 0 input < size (st_obs (s x)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (message_index input x) (message_round input) (message_status input))) = (mkSt (st_obs (s x)) (Some (mkRS (message_round input + 1) (state_status (st_rs (s x))))), None)state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
m: Message
Hv: MC_valid receive (s x) (Some m)
Hc: MC_no_equivocation s m
Ht: (let (si', om') := MC_transition x receive (s x) (Some m) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: state_status (st_rs (s x)) = undecided
Hmround: state_round (st_rs (s x)) <= from_option msg_round 0 (Some m) < size (st_obs (s x)) - 1
Hmindex: from_option msg_index x (Some m) ∈ st_obs (s x)
Hmmuddy: from_option msg_status undecided (Some m) = undecided
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (message_index (Some m) x) (message_round (Some m)) (message_status (Some m)))) = (mkSt (st_obs (s x)) (Some (mkRS (message_round (Some m) + 1) (state_status (st_rs (s x))))), None)
o: indexSet
rs: RoundStatus
H9: mkSt o (Some rs) = s xstate_round (st_rs (destination x)) > rs_round rsby rewrite Htr in Ht; inversion Ht; subst; state_update_simpl; cbn; lia. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
o: indexSet
rs_round0: nat
rs_status0: ChildStatus
Hv: MC_valid receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0))
Hc: match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end
Ht: (let (si', om') := MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: rs_status0 = undecided
Hmround: rs_round0 <= msg_round0 < size o - 1
Hmindex: msg_index0 ∈ o
Hmmuddy: msg_status0 = undecided
Hrounds: rs_round0 < size o
Htr: MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) = (mkSt o (Some (mkRS (msg_round0 + 1) rs_status0)), None)state_round (st_rs (destination x)) > rs_round0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = undecided → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): indexprojT2 (l item) = receive → input_valid_transition_item MC_composite_vlsm s item → state_status (st_rs (s i)) = undecided → message_status (input item) = undecided → message_index (input item) i ∈ st_obs (s i) → message_round (input item) = size (st_obs (s i)) - 1 → state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
item: composite_transition_item MCVLSM
i:= projT1 (l item): index
Hl: projT2 (l item) = receive
Hvalid: input_valid_transition_item MC_composite_vlsm s item
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (input item) = undecided
Hmindex: message_index (input item) i ∈ st_obs (s i)
Hmround: message_round (input item) = size (st_obs (s i)) - 1state_round (st_rs (destination item i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))by apply MC_valid_noninitial_state_undecided_round_less_obs; cbn in *; subst; [| inversion Hv |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1state_round (st_rs (s x)) < size (st_obs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
l: label (MCVLSM x)
input: option Message
destination: state (composite_type MCVLSM)
output: option Message
i:= projT1 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}): index
Hl: projT2 (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) = receive
Hs: valid_state_prop MC_composite_vlsm s
Hv: valid (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Hc: MC_constraint (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hsundecided: state_status (st_rs (s i)) = undecided
Hmmuddy: message_status (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = undecided
Hmindex: message_index (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) i ∈ st_obs (s i)
Hmround: message_round (VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = size (st_obs (s i)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))state_round (st_rs (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |} i)) > state_round (st_rs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
input: option Message
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
Ht: (let (si', om') := MC_transition x receive (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hc: match input with | Some m => MC_no_equivocation s m | None => True end
Hv: MC_valid receive (s x) input
Hsundecided: state_status (st_rs (s x)) = undecided
Hmmuddy: from_option msg_status undecided input = undecided
Hmindex: from_option msg_index x input ∈ st_obs (s x)
Hmround: from_option msg_round 0 input = size (st_obs (s x)) - 1
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (message_index input x) (message_round input) (message_status input))) = (mkSt (st_obs (s x)) (Some (mkRS (message_round input + 1) muddy)), None)state_round (st_rs (destination x)) > state_round (st_rs (s x))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
m: Message
Hv: MC_valid receive (s x) (Some m)
Hc: MC_no_equivocation s m
Ht: (let (si', om') := MC_transition x receive (s x) (Some m) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: state_status (st_rs (s x)) = undecided
Hmround: from_option msg_round 0 (Some m) = size (st_obs (s x)) - 1
Hmindex: from_option msg_index x (Some m) ∈ st_obs (s x)
Hmmuddy: from_option msg_status undecided (Some m) = undecided
Hrounds: state_round (st_rs (s x)) < size (st_obs (s x))
Htr: MC_transition x receive (mkSt (st_obs (s x)) (Some (mkRS (state_round (st_rs (s x))) (state_status (st_rs (s x)))))) (Some (mkMsg (message_index (Some m) x) (message_round (Some m)) (message_status (Some m)))) = (mkSt (st_obs (s x)) (Some (mkRS (message_round (Some m) + 1) muddy)), None)
o: indexSet
rs: RoundStatus
H9: mkSt o (Some rs) = s xstate_round (st_rs (destination x)) > state_round (st_rs (mkSt o (Some rs)))by rewrite Htr in Ht; inversion Ht; subst; state_update_simpl; cbn; lia. Qed. Definition MC_transition_item_update s j i st rs : transition_item := {| l := existT i receive : composite_label MCVLSM; input := Some (mkMsg j (state_round (st_rs (s j))) st); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some rs)); output := None |}.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
x: index
destination: composite_state MCVLSM
output: option Message
i:= x: index
Hs: valid_state_prop MC_composite_vlsm s
msg_index0: index
msg_round0: nat
msg_status0: ChildStatus
o: indexSet
rs_round0: nat
rs_status0: ChildStatus
Hv: MC_valid receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0))
Hc: match s msg_index0 with | mkSt _ (Some (mkRS r status)) => msg_status0 = status ∧ msg_round0 = r ∨ msg_status0 = undecided ∧ msg_round0 < r | mkSt _ None => False end
Ht: (let (si', om') := MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) in (state_update MCVLSM s x si', om')) = (destination, output)
Hsundecided: rs_status0 = undecided
Hmround: msg_round0 = size o - 1
Hmindex: msg_index0 ∈ o
Hmmuddy: msg_status0 = undecided
Hrounds: rs_round0 < size o
Htr: MC_transition x receive (mkSt o (Some (mkRS rs_round0 rs_status0))) (Some (mkMsg msg_index0 msg_round0 msg_status0)) = (mkSt o (Some (mkRS (msg_round0 + 1) muddy)), None)state_round (st_rs (destination x)) > rs_round0
The progress lemma MC_progress ensures that from any valid non-initial
composite state, there is a valid transition to a new composite state, in
which at least one component has greater round than before the transition.
This result will be further used in the proof of the correctness theorem.
To prove the progress lemma, we proceed by case analysis. We distinguish
two main cases:
- at least one component is in an initial state, so we can use an init transition to prove our goal;
- else, we have to obtain two convenient components and build a transition between them, that results in an increased round number.
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMMC_non_initial_valid_state s → ¬ MC_final_state s → length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, let i := projT1 (l item) in input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item i) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMMC_non_initial_valid_state s → ¬ MC_final_state s → length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, let i := projT1 (l item) in input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item i) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMMC_non_initial_valid_state s → ¬ (∀ n : index, state_status (st_rs (s n)) ≠ undecided) → length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMMC_non_initial_valid_state s → ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index) → length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
n: ¬ (∃ i : index, MC_initial_state_prop (s i))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
Hobssi: size (st_obs (s i)) = 0length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
n: nat
Hobssi: size (st_obs (s i)) = S nlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
Hobssi: size (st_obs (s i)) = 0length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |} ∧ state_round_inc (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))) i) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |} ∧ 1 > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |} ∧ 1 > match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |} ∧ 1 > 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1MuddyUnion s ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1
n, x: indexx ∈ st_obs (s n) → x ∈ MuddyUnion s ∖ {[n]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1
n, x: indexx ∈ MuddyUnion s ∖ {[n]} → x ∈ st_obs (s n)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1MC_valid init (s i) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1MC_valid init (mkSt st_obs0 None) Noneby apply MC_non_initial_valid_consistent in Hs as [].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1MuddyUnion s ≢ ∅by apply MC_non_initial_valid_consistent in Hs as []; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1
n, x: indexx ∈ st_obs (s n) → x ∈ MuddyUnion s ∖ {[n]}by apply MC_non_initial_valid_consistent in Hs as []; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1
n, x: indexx ∈ MuddyUnion s ∖ {[n]} → x ∈ st_obs (s n)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
Hobssi: size (st_obs (s i)) = 0
H9: length (enum index) > 1(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
Heq: size o = 0
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 muddy)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
n: nat
Heq: size o = S n
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 undecided)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 muddy))), None)by rewrite <- Heqcall; inversion H11.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
Heq: size o = 0
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 muddy)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 muddy))), None)by inversion H10; congruence.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
n: nat
Heq: size o = S n
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
Hobssi: size st_obs0 = 0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 undecided)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
n: nat
Hobssi: size (st_obs (s i)) = S nlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |} ∧ state_round_inc (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))) i) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |} ∧ 1 > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: MC_initial_state_prop (s i)
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |} ∧ 1 > match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |} ∧ 1 > 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1MuddyUnion s ≢ ∅index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1
n0, x: indexx ∈ st_obs (s n0) → x ∈ MuddyUnion s ∖ {[n0]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1
n0, x: indexx ∈ MuddyUnion s ∖ {[n0]} → x ∈ st_obs (s n0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)by apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1MC_valid init (s i) Noneby constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
st_obs0: indexSet
n: nat
Hobssi: size st_obs0 = S n
H9: length (enum index) > 1MC_valid init (mkSt st_obs0 None) Noneby apply MC_non_initial_valid_consistent in Hs as [].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1MuddyUnion s ≢ ∅apply MC_non_initial_valid_consistent in Hs as []; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1
n0, x: indexx ∈ st_obs (s n0) → x ∈ MuddyUnion s ∖ {[n0]}apply MC_non_initial_valid_consistent in Hs as []; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1
n0, x: indexx ∈ MuddyUnion s ∖ {[n0]} → x ∈ st_obs (s n0)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
He: st_rs (s i) = None
n: nat
Hobssi: size (st_obs (s i)) = S n
H9: length (enum index) > 1(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
i: index
st_obs0: indexSet
n: nat
Hobssi: size st_obs0 = S n
H9: length (enum index) > 1(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
Heq: size o = 0
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
n: nat
Hobssi: size st_obs0 = S n
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 muddy)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
n: nat
Heq: size o = S n
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
n0: nat
Hobssi: size st_obs0 = S n0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 undecided)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)by inversion H10; congruence.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
Heq: size o = 0
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
n: nat
Hobssi: size st_obs0 = S n
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 muddy)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)by rewrite <- Heqcall; inversion H11.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
i: index
o: indexSet
n: nat
Heq: size o = S n
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
st_obs0: indexSet
n0: nat
Hobssi: size st_obs0 = S n0
H9: length (enum index) > 1
H10: init = init
H11: mkSt o None = mkSt st_obs0 None
Heqcall: (mkSt o (Some (mkRS 0 undecided)), None) = MC_transition i init (mkSt st_obs0 None) None(let (si', om') := MC_transition i init (mkSt st_obs0 None) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt st_obs0 (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: ¬ Forall (λ x : index, state_status (st_rs (s x)) ≠ undecided) (enum index)
n: ¬ (∃ i : index, MC_initial_state_prop (s i))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
Hall: Exists (not ∘ (λ x : index, state_status (st_rs (s x)) ≠ undecided)) (enum index)
n: ¬ (∃ i : index, MC_initial_state_prop (s i))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))state_status (st_rs (s i)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))by destruct (decide (state_status (st_rs (s i)) = undecided)).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))state_status (st_rs (s i)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddylength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddylength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ state_round_inc (destination {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} (projT1 (l {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |}))) > state_round_inc (s (projT1 (l {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: valid_state_prop MC_composite_vlsm s ∧ ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)))
Hc: MC_constraint (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)))input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy)))
Hc: let (x, l0) := let (l, _, _, _) := MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy) in l in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hc: state_round (st_rs (s j)) = size (MuddyUnion s) - 1 ∧ j ∈ MuddyUnion sinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))); output := None |} ∧ S (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: state_round (st_rs (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))) i)) > state_round (st_rs (s i))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sstate_round (st_rs (s j)) = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: state_round (st_rs (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))) i)) > state_round (st_rs (s i))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j))) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: state_round (st_rs (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy)))) > state_round (st_rs (s i))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j))) > state_round_inc (s i)by case_match; cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: state_round (st_rs (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy)))) > state_round (st_rs (s i))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j))) > match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion s
n0: i ≠ jj ∈ st_obs (s i)by clear - Hj n e; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion s
n0: i ≠ jj ∈ MuddyUnion s ∖ {[i]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sstate_round (st_rs (s j)) = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion si ∈ MuddyUnion sby clear - e; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hv: MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion si ∈ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)by repeat case_match; [rewrite <- H10; left; rewrite H10, <- He |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1match s j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ state_round (st_rs (s j)) = r ∨ muddy = undecided ∧ state_round (st_rs (s j)) < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
st_rs0: option RoundStatus
Hsi: s i = mkSt st_obs0 st_rs0
Hi: ¬ state_status (st_rs (mkSt st_obs0 st_rs0)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 st_rs0)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
r: RoundStatus
Hsi: s i = mkSt st_obs0 (Some r)
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some r))) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 (Some r))) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
sti: ChildStatus
Hsi: s i = mkSt st_obs0 (Some (mkRS ri sti))
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri sti)))) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 (Some (mkRS ri sti)))) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition i receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy)) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)by eapply MC_valid_noequiv_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))by rewrite Hsi; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition i receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy)) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition_clause_4 i st_obs0 ri j (decide (j ∈ st_obs0)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: MC_component_invariant s j(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: match state_status (st_rs (s j)) with | undecided => state_round (st_rs (s j)) < size (st_obs (s j)) | muddy => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j)) | clean => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) = size (st_obs (s j)) end(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j))(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (MuddyUnion s) - 1 = size (st_obs (s j))(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j))) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (size (MuddyUnion s) - 1) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs (s i)) then λ r' : nat, if decide (r' = size (st_obs (s i))) then (mkSt (st_obs (s i)) (Some (mkRS r' muddy)), None) else if decide (r' = size (st_obs (s i)) - 1) then (mkSt (st_obs (s i)) (Some (mkRS (r' + 1) clean)), None) else (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None)) (size (MuddyUnion s) - 1) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := if decide (size (MuddyUnion s) - 1 = size (st_obs (s i))) then (mkSt (st_obs (s i)) (Some (mkRS (size (MuddyUnion s) - 1) muddy)), None) else if decide (size (MuddyUnion s) - 1 = size (st_obs (s i)) - 1) then (mkSt (st_obs (s i)) (Some (mkRS (size (MuddyUnion s) - 1 + 1) clean)), None) else (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := if decide (size (MuddyUnion s) - 1 = size (st_obs (s i))) then (mkSt (st_obs (s i)) (Some (mkRS (size (MuddyUnion s) - 1) muddy)), None) else if decide (size (MuddyUnion s) - 1 = size (st_obs (s i)) - 1) then (mkSt (st_obs (s i)) (Some (mkRS (size (MuddyUnion s) - 1 + 1) clean)), None) else (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1) muddy))), None)by symmetry; apply MC_muddy_number_of_muddy_seen_iff; [| by apply MuddyUnion_elem in e].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ st_obs (s i)by destruct (decide (i = j)); [subst; rewrite Hsi in He | set_solver].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hinit: MuddyUnion s ≢ ∅
j: index
He: state_status (st_rs (s j)) = muddy
e: i ∈ st_obs (s j)
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ MuddyUnion s ∖ {[i]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ state_round_inc (destination {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} (projT1 (l {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |}))) > state_round_inc (s (projT1 (l {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: valid_state_prop MC_composite_vlsm s ∧ ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)) → input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)) → input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hc: MC_constraint (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))input_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hc: let (x, l0) := let (l, _, _, _) := MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean) in l in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hc: state_round (st_rs (s j)) = size (MuddyUnion s) - 1 ∧ j ∈ MuddyUnion sinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sinput_valid_transition_item MC_composite_vlsm s {| l := existT i receive; input := Some (mkMsg j (state_round (st_rs (s j))) muddy); destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: state_round (st_rs (destination (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)) (projT1 (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))))) > state_round (st_rs (s (projT1 (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j)) + 1) > state_round_inc (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sstate_round (st_rs (s j)) = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: state_round (st_rs (destination (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)) (projT1 (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))))) > state_round (st_rs (s (projT1 (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j)) + 1) > state_round_inc (s i)by case_match; cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: state_round (st_rs (s j)) + 1 > state_round (st_rs (s i))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sS (state_round (st_rs (s j)) + 1) > match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion s
n1: i ≠ jj ∈ st_obs (s i)by clear - Hj n n1; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion s
n1: i ≠ jj ∈ MuddyUnion s ∖ {[i]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion sstate_round (st_rs (s j)) = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion si ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion si ∉ MuddyUnion sby clear - n0 n1; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hv: valid (l (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))) (s, input (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean)))
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hj: j ∈ MuddyUnion s
n1: i ≠ ji ∉ MuddyUnion sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)by repeat case_match; [rewrite <- H10; left; rewrite H10, <- He |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}match s j with | mkSt _ (Some (mkRS r status)) => muddy = status ∧ state_round (st_rs (s j)) = r ∨ muddy = undecided ∧ state_round (st_rs (s j)) < r | mkSt _ None => False endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
st_rs0: option RoundStatus
Hsi: s i = mkSt st_obs0 st_rs0
Hi: ¬ state_status (st_rs (mkSt st_obs0 st_rs0)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 st_rs0)) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
r: RoundStatus
Hsi: s i = mkSt st_obs0 (Some r)
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some r))) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 (Some r))) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
sti: ChildStatus
Hsi: s i = mkSt st_obs0 (Some (mkRS ri sti))
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri sti)))) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (mkSt st_obs0 (Some (mkRS ri sti)))) = undecided
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j i muddy (mkRS (state_round (st_rs (s j)) + 1) clean))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition i receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy)) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)by eapply MC_valid_noequiv_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)option_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) muddy))by rewrite Hsi; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)MC_valid receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition i receive (s i) (Some (mkMsg j (state_round (st_rs (s j))) muddy)) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := MC_transition_clause_4 i st_obs0 ri j (decide (j ∈ st_obs0)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: MC_component_invariant s j(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: match state_status (st_rs (s j)) with | undecided => state_round (st_rs (s j)) < size (st_obs (s j)) | muddy => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j)) | clean => state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) = size (st_obs (s j)) end(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hinvariant: state_round (st_rs (s j)) = size (st_obs (s j)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s j))(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (MuddyUnion s) - 1 = size (st_obs (s j))(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) muddy)
Hroundj: state_round (st_rs (s j)) = size (st_obs (s j))
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) clean)), None) else (mkSt st_obs0 (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt st_obs0 (Some (mkRS ri undecided)), None)) (size (MuddyUnion s) - 1) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt st_obs0 (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1 + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s(let (si', om') := (if decide (j ∈ st_obs (s i)) then λ r' : nat, if decide (r' = size (st_obs (s i))) then (mkSt (st_obs (s i)) (Some (mkRS r' muddy)), None) else if decide (r' = size (st_obs (s i)) - 1) then (mkSt (st_obs (s i)) (Some (mkRS (r' + 1) clean)), None) else (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None) else λ _ : nat, (mkSt (st_obs (s i)) (Some (mkRS ri undecided)), None)) (size (MuddyUnion s) - 1) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) (Some (mkRS ri undecided)))) (Some (mkRS (size (MuddyUnion s) - 1 + 1) clean))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 ≠ size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s j)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1by rewrite Hsi in He.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s j)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1by apply MC_clean_number_of_muddy_seen in HinotinMuddy; [rewrite HinotinMuddy |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ j
HinotinMuddy: i ∉ MuddyUnion ssize (MuddyUnion s) - 1 = size (st_obs (s i)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 ≠ size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 ≠ size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 ≠ size (st_obs (s i))by rewrite Hsi in He.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion ssize (MuddyUnion s) - 1 ≠ size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 ≠ size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jsize (MuddyUnion s) - 1 ≠ size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ j
HinotinMuddy: i ∉ MuddyUnion ssize (MuddyUnion s) - 1 ≠ size (st_obs (s i))by destruct (decide (size (MuddyUnion s) = 0)); [apply size_empty_iff in e | lia].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ MuddyUnion s ∖ {[j]}
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ j
HinotinMuddy: size (st_obs (s i)) = size (MuddyUnion s)size (MuddyUnion s) - 1 ≠ size (st_obs (s i))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ st_obs (s i)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ MuddyUnion s ∖ {[i]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ MuddyUnion s ∖ {[j]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jj ∈ MuddyUnion s ∖ {[i]}by rewrite Hsi in He.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
j: index
Hsi: s j = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
He: state_status (st_rs (s j)) = muddy
n0: j ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion sj ∈ MuddyUnion s ∖ {[j]}by set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninitial: ¬ composite_initial_state_prop MCVLSM s
i: index
st_obs0: indexSet
ri: nat
Hi: ¬ state_status (st_rs (mkSt st_obs0 (Some (mkRS ri undecided)))) ≠ undecided
Hsi: s i = mkSt st_obs0 (Some (mkRS ri undecided))
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hcons: consistent s
j: index
He: state_status (st_rs (s j)) = muddy
n0: i ∉ st_obs (s j)
H9: length (enum index) > 1
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hnoequiv: MC_no_equivocation s (mkMsg j (size (MuddyUnion s) - 1) muddy)
Hroundj: state_round (st_rs (s j)) = size (MuddyUnion s) - 1
Hobsj: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hjmuddy: j ∈ MuddyUnion s
n1: i ≠ jj ∈ MuddyUnion s ∖ {[i]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hcons: consistent s
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: MC_non_initial_valid_state s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hs: valid_state_prop MC_composite_vlsm s ∧ ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': valid_state_prop MC_composite_vlsm slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive slength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s∃ k : index, k ∈ st_obs (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s∃ k : index, k ∈ st_obs (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive ssize (st_obs (s j)) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
Hinvariant: MC_component_invariant_inductive s jsize (st_obs (s j)) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
Hinvariant: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jsize (st_obs (s j)) ≠ 0by apply size_non_empty_iff in Hmuddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hvalid': MC_composite_invariant_inductive s
Hinvariant: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jsize (st_obs (s j)) ≠ 0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)state_status (st_rs (s k)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)state_status (st_rs (s k)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hinvariantk: MC_component_invariant_inductive s kstate_status (st_rs (s k)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hinvariantk: MC_component_invariant_inductive s kstate_status (st_rs (s k)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hinvariantk: MC_component_invariant_inductive s k
s0: index → State
i0: index
H9: state_status (st_rs (s k)) = clean
H10: state_round (st_rs (s k)) = size (st_obs (s k))
H11: size (MuddyUnion s) = size (st_obs (s k))
H12: s0 = s
H13: i0 = kstate_status (st_rs (s k)) = undecidedby apply size_non_empty_iff in Hmuddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: size (st_obs (s k)) = size (MuddyUnion s) - 1
Hinvariantk: MC_component_invariant_inductive s k
s0: index → State
i0: index
H9: state_status (st_rs (s k)) = clean
H10: state_round (st_rs (s k)) = size (st_obs (s k))
H11: size (MuddyUnion s) = size (st_obs (s k))
H12: s0 = s
H13: i0 = kstate_status (st_rs (s k)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecidedstate_round (st_rs (s k)) < size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecidedstate_round (st_rs (s k)) < size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: MC_component_invariant_inductive s kstate_round (st_rs (s k)) < size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: MC_component_invariant_inductive s k
s0: index → State
i0: index
H9: state_status (st_rs (s k)) = clean
H10: state_round (st_rs (s k)) = size (st_obs (s k))
H11: size (MuddyUnion s) = size (st_obs (s k))
H12: s0 = s
H13: i0 = kstate_round (st_rs (s k)) < size (st_obs (s k))by rewrite Hkobs in H11; apply size_non_empty_iff in Hmuddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: size (st_obs (s k)) = size (MuddyUnion s) - 1
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: MC_component_invariant_inductive s k
s0: index → State
i0: index
H9: state_status (st_rs (s k)) = clean
H10: state_round (st_rs (s k)) = size (st_obs (s k))
H11: size (MuddyUnion s) = size (st_obs (s k))
H12: s0 = s
H13: i0 = kstate_round (st_rs (s k)) < size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))state_status (st_rs (s j)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))state_status (st_rs (s j)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hinvariantj: MC_component_invariant_inductive s jstate_status (st_rs (s j)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hinvariantj: MC_component_invariant_inductive s jstate_status (st_rs (s j)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hinvariantj: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jstate_status (st_rs (s j)) = undecidedby apply size_non_empty_iff in Hmuddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hinvariantj: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jstate_status (st_rs (s j)) = undecidedindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecidedlength (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecidedstate_round (st_rs (s j)) < size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecidedstate_round (st_rs (s j)) < size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: MC_component_invariant_inductive s jstate_round (st_rs (s j)) < size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jstate_round (st_rs (s j)) < size (st_obs (s j))by rewrite Hjmuddyunion in H11; apply size_non_empty_iff in Hmuddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: MC_component_invariant_inductive s j
s0: index → State
i0: index
H9: state_status (st_rs (s j)) = clean
H10: state_round (st_rs (s j)) = size (st_obs (s j))
H11: size (MuddyUnion s) = size (st_obs (s j))
H12: s0 = s
H13: i0 = jstate_round (st_rs (s j)) < size (st_obs (s j))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))size (st_obs (s j)) = size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))size (st_obs (s j)) = size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))size (st_obs (s j)) = size (st_obs (s k))by rewrite <- Hjmuddyunion in Hkobs.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: size (st_obs (s j)) = size (MuddyUnion s) - 1
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: size (st_obs (s k)) = size (MuddyUnion s) - 1
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))size (st_obs (s j)) = size (st_obs (s k))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))j ∈ st_obs (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))j ∈ st_obs (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))j ∈ MuddyUnion s ∖ {[k]}by rewrite Hobs in Hkobs; set_solver.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
k: index
Hkobs: k ∈ st_obs (s k)
Hjmuddyunion: k ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjk_eq_size_muddy: size (st_obs (s k)) = size (st_obs (s k))
Hjinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s k)) = undecidedk ∈ MuddyUnion s ∖ {[k]}index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ state_round_inc (destination {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} (projT1 (VLSM.l {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |}))) > state_round_inc (s (projT1 (VLSM.l {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hc: MC_constraint (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hc: let (x, l0) := VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)) in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hc: valid_message_prop MC_composite_vlsm (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)) (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))))) > state_round (st_rs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s k)) + 1 > state_round (st_rs (s j))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > match st_rs (s j) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
r: RoundStatus
H10: st_rs (s j) = Some r
Hjundecided: state_status (Some r) = undecided
Hjinvariant: state_round (Some r) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (Some r) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s k)) + 1 > state_round (Some r)
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > S (rs_round r)by cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
roundk: nat
Heqroundk: roundk = state_round (st_rs (s k))
Hkinvariant: roundk < size (st_obs (s k))
r: RoundStatus
H10: st_rs (s j) = Some r
Hjundecided: state_status (Some r) = undecided
Hjinvariant: state_round (Some r) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (Some r) < roundk
Heqobsminus1: roundk = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: roundk + 1 > from_option rs_round 0 (Some r)
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) muddy)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (roundk + 1) muddy))) (s, input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) muddy)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k roundk undecided))S (roundk + 1) > S (rs_round r)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))by unfold MC_no_equivocation; repeat case_match; cbn in *; [left; split | lia].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hjundecided': state_status (st_rs (s j)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
st_rs0: option RoundStatus
Hsj: s j = mkSt st_obs0 st_rs0
Hjundecided': state_status (st_rs (mkSt st_obs0 st_rs0)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
r: RoundStatus
Hsj: s j = mkSt st_obs0 (Some r)
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some r))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided(let (si', om') := MC_transition j receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition j receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition j receive (mkSt st_obs0 (Some (mkRS rj undecided))) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition_clause_5 j st_obs0 rj k (decide (k ∈ st_obs0)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (k ∈ st_obs0) then λ r' : nat, if decide (r' < rj) then (mkSt st_obs0 (Some (mkRS rj undecided)), None) else if decide (rj <= r' < size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt st_obs0 (Some (mkRS rj undecided)), None) else λ r' : nat, if decide (r' ≤ rj) then (mkSt st_obs0 (Some (mkRS rj undecided)), None) else if decide (rj < r' < size st_obs0) then (mkSt st_obs0 (Some (mkRS r' undecided)), None) else if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else (mkSt st_obs0 (Some (mkRS rj undecided)), None)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)by rewrite decide_True, decide_False, decide_False, decide_True; [| lia .. |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Heqobsminus1: state_round (st_rs (s k)) = size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (k ∈ st_obs (s j)) then λ r' : nat, if decide (r' < state_round (st_rs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else if decide (state_round (st_rs (s j)) <= r' < size (st_obs (s j)) - 1) then (mkSt (st_obs (s j)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s j)) - 1) then (mkSt (st_obs (s j)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else λ r' : nat, if decide (r' ≤ state_round (st_rs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else if decide (state_round (st_rs (s j)) < r' < size (st_obs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (VLSM.l item))) > state_round_inc (s (projT1 (VLSM.l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ state_round_inc (destination {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} (projT1 (VLSM.l {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |}))) > state_round_inc (s (projT1 (VLSM.l {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)) → input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)) → input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hc: MC_constraint (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hc: let (x, l0) := VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)) in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hc: valid_message_prop MC_composite_vlsm (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))input_valid_transition_item MC_composite_vlsm s {| l := existT j receive; input := Some (mkMsg k (state_round (st_rs (s k))) undecided); destination := state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)) (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) > state_round (st_rs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))state_round (st_rs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) <= message_round (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) < size (st_obs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)) (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) > state_round (st_rs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > state_round_inc (s j)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
r: RoundStatus
H10: st_rs (s j) = Some r
Hjundecided: state_status (Some r) = undecided
Hjinvariant: state_round (Some r) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (Some r) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s k)) + 1 > state_round (Some r)
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))S (state_round (st_rs (s k)) + 1) > S (rs_round r)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
roundk: nat
Heqroundk: roundk = state_round (st_rs (s k))
Hkinvariant: roundk < size (st_obs (s k))
r: RoundStatus
H10: st_rs (s j) = Some r
Hjundecided: state_status (Some r) = undecided
Hjinvariant: state_round (Some r) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (Some r) < roundk
Hneqobsminus1: roundk ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: roundk + 1 > state_round (Some r)
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k roundk undecided))S (roundk + 1) > S (rs_round r)by cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
roundk: nat
Heqroundk: roundk = state_round (st_rs (s k))
Hkinvariant: roundk < size (st_obs (s k))
r: RoundStatus
H10: st_rs (s j) = Some r
Hjundecided: state_status (Some r) = undecided
Hjinvariant: state_round (Some r) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (Some r) < roundk
Hneqobsminus1: roundk ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: roundk + 1 > from_option rs_round 0 (Some r)
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (roundk + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k roundk undecided))S (roundk + 1) > S (rs_round r)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hv: valid (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) (s, input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))
Hroundk: state MC_composite_vlsm
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))state_round (st_rs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) <= message_round (input (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))) < size (st_obs (s (projT1 (VLSM.l (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided)))))) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hv: MC_valid receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hroundk: composite_state MCVLSM
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))state_round (st_rs (s j)) <= state_round (st_rs (s k)) < size (st_obs (s j)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hv: MC_valid receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hroundk: composite_state MCVLSM
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))state_round (st_rs (s k)) < size (st_obs (s j)) - 1by rewrite Hjk_eq_size_muddy; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hv: MC_valid receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided))
Hroundk: composite_state MCVLSM
Hk: valid_state_message_prop MC_composite_vlsm Hroundk (Some (mkMsg k (state_round (st_rs (s k))) undecided))state_round (st_rs (s k)) < size (st_obs (s j)) - 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))by unfold MC_no_equivocation; repeat case_match; [left | cbn in *; lia].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hjundecided': state_status (st_rs (s j)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
st_rs0: option RoundStatus
Hsj: s j = mkSt st_obs0 st_rs0
Hjundecided': state_status (st_rs (mkSt st_obs0 st_rs0)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
r: RoundStatus
Hsj: s j = mkSt st_obs0 (Some r)
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some r))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s k j undecided (mkRS (state_round (st_rs (s k)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided(let (si', om') := MC_transition j receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition j receive (s j) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition j receive (mkSt st_obs0 (Some (mkRS rj undecided))) (Some (mkMsg k (state_round (st_rs (s k))) undecided)) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition_clause_5 j st_obs0 rj k (decide (k ∈ st_obs0)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (k ∈ st_obs0) then λ r' : nat, if decide (r' < rj) then (mkSt st_obs0 (Some (mkRS rj undecided)), None) else if decide (rj <= r' < size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt st_obs0 (Some (mkRS rj undecided)), None) else λ r' : nat, if decide (r' ≤ rj) then (mkSt st_obs0 (Some (mkRS rj undecided)), None) else if decide (rj < r' < size st_obs0) then (mkSt st_obs0 (Some (mkRS r' undecided)), None) else if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else (mkSt st_obs0 (Some (mkRS rj undecided)), None)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rj undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)by rewrite decide_True, decide_False, decide_True; [| lia.. |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
st_obs0: indexSet
rj: nat
stj: ChildStatus
Hjundecided: stj = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
l: state_round (st_rs (s j)) < state_round (st_rs (s k))
Hneqobsminus1: state_round (st_rs (s k)) ≠ size (st_obs (s j)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg k (state_round (st_rs (s k))) undecided)
Hsj: s j = mkSt st_obs0 (Some (mkRS rj stj))
Hjundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rj stj)))) = undecided
Hjundecided'': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (k ∈ st_obs (s j)) then λ r' : nat, if decide (r' < state_round (st_rs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else if decide (state_round (st_rs (s j)) <= r' < size (st_obs (s j)) - 1) then (mkSt (st_obs (s j)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s j)) - 1) then (mkSt (st_obs (s j)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else λ r' : nat, if decide (r' ≤ state_round (st_rs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None) else if decide (state_round (st_rs (s j)) < r' < size (st_obs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s j))) then (mkSt (st_obs (s j)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)), None)) (state_round (st_rs (s k))) in (state_update MCVLSM s j si', om')) = (state_update MCVLSM s j (mkSt (st_obs (mkSt (st_obs (s j)) (Some (mkRS (state_round (st_rs (s j))) undecided)))) (Some (mkRS (state_round (st_rs (s k)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ state_round_inc (destination {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} (projT1 (l {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |}))) > state_round_inc (s (projT1 (l {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)) → input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hc: MC_constraint (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hc: let (x, l0) := l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)) in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hc: valid_message_prop MC_composite_vlsm (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)) (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))))) > state_round (st_rs (s (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s j)) + 1 > state_round (st_rs (s k))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)by case_match; cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s j)) + 1 > state_round (st_rs (s k))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > match st_rs (s k) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))by unfold MC_no_equivocation; repeat case_match; cbn in *; [left | contradict n; eexists; rewrite H10].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hkundecided': state_status (st_rs (s k)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
st_rs0: option RoundStatus
Hsk: s k = mkSt st_obs0 st_rs0
Hkundecided': state_status (st_rs (mkSt st_obs0 st_rs0)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
r: RoundStatus
Hsk: s k = mkSt st_obs0 (Some r)
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some r))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) muddy))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedoption_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedMC_valid receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)by eapply MC_valid_noequiv_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedoption_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) undecided))by rewrite Hsk; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedMC_valid receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (mkSt st_obs0 (Some (mkRS rk undecided))) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition_clause_5 k st_obs0 rk j (decide (j ∈ st_obs0)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' < rk) then (mkSt st_obs0 (Some (mkRS rk undecided)), None) else if decide (rk <= r' < size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt st_obs0 (Some (mkRS rk undecided)), None) else λ r' : nat, if decide (r' ≤ rk) then (mkSt st_obs0 (Some (mkRS rk undecided)), None) else if decide (rk < r' < size st_obs0) then (mkSt st_obs0 (Some (mkRS r' undecided)), None) else if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else (mkSt st_obs0 (Some (mkRS rk undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := (if decide (j ∈ st_obs (s k)) then λ r' : nat, if decide (r' < rk) then (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else if decide (rk <= r' < size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else λ r' : nat, if decide (r' ≤ rk) then (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else if decide (rk < r' < size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt (st_obs (s k)) (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)by rewrite decide_True, decide_False, decide_False, decide_True; [| | lia.. |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
e: state_round (st_rs (s j)) = size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided', Hkundecided'': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := (if decide (j ∈ st_obs (s k)) then λ r' : nat, if decide (r' < state_round (st_rs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else if decide (state_round (st_rs (s k)) <= r' < size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else λ r' : nat, if decide (r' ≤ state_round (st_rs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else if decide (state_round (st_rs (s k)) < r' < size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1length (enum index) > 1 → ∃ item : composite_transition_item MCVLSM, input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ state_round_inc (destination {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} (projT1 (l {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |}))) > state_round_inc (s (projT1 (l {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |})))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)) → input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)) → input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hc: MC_constraint (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hc: let (x, l0) := l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)) in match l0 with | init => consistent s | emit => True | receive => match (let (_, input, _, _) := MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided) in input) with | Some m => MC_no_equivocation s m | None => True end endinput_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hc: valid_message_prop MC_composite_vlsm (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))input_valid_transition_item MC_composite_vlsm s {| l := existT k receive; input := Some (mkMsg j (state_round (st_rs (s j))) undecided); destination := state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))); output := None |} ∧ S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)) (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))))) > state_round (st_rs (s (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > state_round_inc (s k)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (destination (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)) (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))))) > state_round (st_rs (s (projT1 (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))))))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > match st_rs (s k) with | Some rs => S (rs_round rs) | None => 0 endby case_match; cbn in Ht; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Ht: state_round (st_rs (s j)) + 1 > state_round (st_rs (s k))
Hm: option_valid_message_prop MC_composite_vlsm (input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hv: valid (l (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))) (s, input (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided)))
Hroundj: state MC_composite_vlsm
Hj: valid_state_message_prop MC_composite_vlsm Hroundj (Some (mkMsg j (state_round (st_rs (s j))) undecided))S (state_round (st_rs (s j)) + 1) > match st_rs (s k) with | Some rs => S (rs_round rs) | None => 0 endindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
st_obs0: indexSet
st_rs0: option RoundStatus
r: RoundStatus
rs_round0: nat
rs_status0: ChildStatus
H12: r = mkRS rs_round0 rs_status0
H11: st_rs0 = Some (mkRS rs_round0 rs_status0)
H10: s j = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
Hkobs: k ∈ st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = undecided
Hjinvariant: state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) < size (st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))))
Hjk_eq_size_muddy: size (st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ rs_round0 < state_round (st_rs (s k))
n2: state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1undecided = rs_status0 ∧ state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = rs_round0 ∨ undecided = undecided ∧ state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) < rs_round0index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
st_obs0: indexSet
st_rs0: option RoundStatus
H11: st_rs0 = None
H10: s j = mkSt st_obs0 None
Hkobs: k ∈ st_obs (mkSt st_obs0 None)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (mkSt st_obs0 None)) = undecided
Hjinvariant: state_round (st_rs (mkSt st_obs0 None)) < size (st_obs (mkSt st_obs0 None))
Hjk_eq_size_muddy: size (st_obs (mkSt st_obs0 None)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ 0 < state_round (st_rs (s k))
n2: state_round (st_rs (mkSt st_obs0 None)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1Falseby rewrite <- H10; left; rewrite H10.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
st_obs0: indexSet
st_rs0: option RoundStatus
r: RoundStatus
rs_round0: nat
rs_status0: ChildStatus
H12: r = mkRS rs_round0 rs_status0
H11: st_rs0 = Some (mkRS rs_round0 rs_status0)
H10: s j = mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))
Hkobs: k ∈ st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = undecided
Hjinvariant: state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) < size (st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0))))
Hjk_eq_size_muddy: size (st_obs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ rs_round0 < state_round (st_rs (s k))
n2: state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1undecided = rs_status0 ∧ state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) = rs_round0 ∨ undecided = undecided ∧ state_round (st_rs (mkSt st_obs0 (Some (mkRS rs_round0 rs_status0)))) < rs_round0by contradict n; eexists; rewrite H10.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
st_obs0: indexSet
st_rs0: option RoundStatus
H11: st_rs0 = None
H10: s j = mkSt st_obs0 None
Hkobs: k ∈ st_obs (mkSt st_obs0 None)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (mkSt st_obs0 None)) = undecided
Hjinvariant: state_round (st_rs (mkSt st_obs0 None)) < size (st_obs (mkSt st_obs0 None))
Hjk_eq_size_muddy: size (st_obs (mkSt st_obs0 None)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ 0 < state_round (st_rs (s k))
n2: state_round (st_rs (mkSt st_obs0 None)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)input_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hkundecided': state_status (st_rs (s k)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
st_rs0: option RoundStatus
Hsk: s k = mkSt st_obs0 st_rs0
Hkundecided': state_status (st_rs (mkSt st_obs0 st_rs0)) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
r: RoundStatus
Hsk: s k = mkSt st_obs0 (Some r)
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some r))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedinput_valid_transition_item MC_composite_vlsm s (MC_transition_item_update s j k undecided (mkRS (state_round (st_rs (s j)) + 1) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedoption_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedMC_valid receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)by eapply MC_valid_noequiv_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedoption_valid_message_prop MC_composite_vlsm (Some (mkMsg j (state_round (st_rs (s j))) undecided))by rewrite Hsk; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecidedMC_valid receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
Hkundecided: state_status (st_rs (s k)) = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition k receive (s k) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition k receive (mkSt st_obs0 (Some (mkRS rk undecided))) (Some (mkMsg j (state_round (st_rs (s j))) undecided)) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := MC_transition_clause_5 k st_obs0 rk j (decide (j ∈ st_obs0)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (j ∈ st_obs0) then λ r' : nat, if decide (r' < rk) then (mkSt st_obs0 (Some (mkRS rk undecided)), None) else if decide (rk <= r' < size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size st_obs0 - 1) then (mkSt st_obs0 (Some (mkRS (r' + 1) muddy)), None) else (mkSt st_obs0 (Some (mkRS rk undecided)), None) else λ r' : nat, if decide (r' ≤ rk) then (mkSt st_obs0 (Some (mkRS rk undecided)), None) else if decide (rk < r' < size st_obs0) then (mkSt st_obs0 (Some (mkRS r' undecided)), None) else if decide (r' = size st_obs0) then (mkSt st_obs0 (Some (mkRS r' muddy)), None) else (mkSt st_obs0 (Some (mkRS rk undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt st_obs0 (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (j ∈ st_obs (s k)) then λ r' : nat, if decide (r' < rk) then (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else if decide (rk <= r' < size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else λ r' : nat, if decide (r' ≤ rk) then (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None) else if decide (rk < r' < size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS rk undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt (st_obs (s k)) (Some (mkRS rk undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)by rewrite decide_True, decide_False, decide_True; [| lia.. |]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
Hnoninit: ¬ composite_initial_state_prop MCVLSM s
i: index
Hi: ¬ state_status (st_rs (s i)) ≠ undecided
n: ¬ (∃ i : index, MC_initial_state_prop (s i))
Hundecided: state_status (st_rs (s i)) = undecided
Hmuddy: MuddyUnion s ≢ ∅
Hobs: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
n0: ¬ (∃ j : index, state_status (st_rs (s j)) = muddy)
j: index
Hjmuddyunion: j ∈ MuddyUnion s
Hvalid': MC_composite_invariant_inductive s
k: index
Hkobs: k ∈ st_obs (s j)
st_obs0: indexSet
rk: nat
stk: ChildStatus
Hkundecided: stk = undecided
Hkinvariant: state_round (st_rs (s k)) < size (st_obs (s k))
Hjundecided: state_status (st_rs (s j)) = undecided
Hjinvariant: state_round (st_rs (s j)) < size (st_obs (s j))
Hjk_eq_size_muddy: size (st_obs (s j)) = size (st_obs (s k))
Hjobs: j ∈ st_obs (s k)
n1: ¬ state_round (st_rs (s j)) < state_round (st_rs (s k))
n2: state_round (st_rs (s j)) ≠ size (st_obs (s k)) - 1
H9: length (enum index) > 1
Hnoequiv: MC_no_equivocation s (mkMsg j (state_round (st_rs (s j))) undecided)
Hsk: s k = mkSt st_obs0 (Some (mkRS rk stk))
Hkundecided': state_status (st_rs (mkSt st_obs0 (Some (mkRS rk stk)))) = undecided
Hjundecided': state_status (st_rs (s j)) = undecided(let (si', om') := (if decide (j ∈ st_obs (s k)) then λ r' : nat, if decide (r' < state_round (st_rs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else if decide (state_round (st_rs (s k)) <= r' < size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) undecided)), None) else if decide (r' = size (st_obs (s k)) - 1) then (mkSt (st_obs (s k)) (Some (mkRS (r' + 1) muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else λ r' : nat, if decide (r' ≤ state_round (st_rs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None) else if decide (state_round (st_rs (s k)) < r' < size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' undecided)), None) else if decide (r' = size (st_obs (s k))) then (mkSt (st_obs (s k)) (Some (mkRS r' muddy)), None) else (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)), None)) (state_round (st_rs (s j))) in (state_update MCVLSM s k si', om')) = (state_update MCVLSM s k (mkSt (st_obs (mkSt (st_obs (s k)) (Some (mkRS (state_round (st_rs (s k))) undecided)))) (Some (mkRS (state_round (st_rs (s j)) + 1) undecided))), None)
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSMvalid_state_prop MC_composite_vlsm s → ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: valid_state_prop MC_composite_vlsm s
i: indexstate_round_inc (s i) ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: MC_composite_invariant s
i: indexstate_round_inc (s i) ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: indexstate_round_inc (s i) ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: indexmatch st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinit: initial_state_prop (s i)match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinv: MC_component_invariant s imatch st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinit: initial_state_prop (s i)match st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinit: st_rs (s i) = Nonematch st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1by lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinit: st_rs (s i) = None0 ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinv: MC_component_invariant s imatch st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
Hinv: match state_status (st_rs (s i)) with | undecided => state_round (st_rs (s i)) < size (st_obs (s i)) | muddy => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i)) | clean => state_round (st_rs (s i)) = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i)) endmatch st_rs (s i) with | Some rs => S (rs_round rs) | None => 0 end ≤ size (st_obs (s i)) + 1by destruct status; lia. Qed. Definition steps_until_final_component (s : State) : nat := size (st_obs s) + 1 - state_round_inc s. Definition steps_until_final_composite (s : composite_state MCVLSM) : nat := sum_list_with (fun i => steps_until_final_component (s i)) (enum index).index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hvalid: ∀ i : index, MC_initial_state_prop (s i) ∨ MC_component_invariant s i
i: index
round: nat
status: ChildStatus
Hinv: match status with | undecided => round < size (st_obs (s i)) | muddy => round = size (st_obs (s i)) ∧ size (MuddyUnion s) - 1 = size (st_obs (s i)) | clean => round = size (st_obs (s i)) ∧ size (MuddyUnion s) = size (st_obs (s i)) endS round ≤ size (st_obs (s i)) + 1
The main result states that, from any non-initial valid state,
there is a trace to a final state.
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSMMC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSMMC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
nr_steps: nat
Heqnr_steps: nr_steps = steps_until_final_composite sMC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat∀ s : composite_state MCVLSM, nr_steps = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf∀ s : composite_state MCVLSM, nr_steps = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
Hprogress: ∃ item : composite_transition_item MCVLSM, let i := projT1 (l item) in input_valid_transition_item MC_composite_vlsm s item ∧ state_round_inc (destination item i) > state_round_inc (s i)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1MC_non_initial_valid_state (destination item)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: MC_non_initial_valid_state (destination item)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1MC_non_initial_valid_state (destination item)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1MC_non_initial_valid_state destinationindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1¬ composite_initial_state_prop MCVLSM destinationindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hinit: composite_initial_state_prop MCVLSM destinationFalseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hinit: initial_state_prop (destination x)Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hinit: st_rs (destination x) = NoneFalseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
st_obs0: indexSet
st_rs0: option RoundStatus
Hround: state_round_inc (mkSt st_obs0 st_rs0) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hinit: st_rs0 = NoneFalseby lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
Hind: ∀ y : nat, y < steps_until_final_composite s → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
st_obs0: indexSet
Hround: 0 > match st_rs (s x) with | Some rs => S (rs_round rs) | None => 0 end
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1Falseindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: MC_non_initial_valid_state (destination item)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm (destination item) sf tr ∧ MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: MC_non_initial_valid_state (destination item)steps_until_final_composite (destination item) < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm (destination item) sf tr ∧ MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfby exists (item :: tr), sf; split; [destruct item; constructor |].index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (destination item) sf tr
Hsf: MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
item: composite_transition_item MCVLSM
Hitem: input_valid_transition_item MC_composite_vlsm s item
Hround: state_round_inc (destination item (projT1 (l item))) > state_round_inc (s (projT1 (l item)))
Hdest: valid_state_prop MC_composite_vlsm (destination item)
Hrounddest: ∀ i : index, state_round_inc (destination item i) ≤ size (st_obs (destination item i)) + 1
Hdestvalid: MC_non_initial_valid_state (destination item)steps_until_final_composite (destination item) < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hitem: input_valid_transition_item MC_composite_vlsm s {| l := existT x l; input := input; destination := destination; output := output |}
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hdestvalid: MC_non_initial_valid_state destinationsteps_until_final_composite destination < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hv: input_valid MC_composite_vlsm (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: transition (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hdestvalid: MC_non_initial_valid_state destinationsteps_until_final_composite destination < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hv: input_valid MC_composite_vlsm (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |})
Ht: (let (initial_state_prop, initial_state, _, initial_message_prop, initial_message, transition, _) := vlsm_machine MC_composite_vlsm in transition) (VLSM.l {| l := existT x l; input := input; destination := destination; output := output |}) (s, VLSM.input {| l := existT x l; input := input; destination := destination; output := output |}) = (VLSM.destination {| l := existT x l; input := input; destination := destination; output := output |}, VLSM.output {| l := existT x l; input := input; destination := destination; output := output |})
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hdestvalid: MC_non_initial_valid_state destinationsteps_until_final_composite destination < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hv: valid_state_prop MC_composite_vlsm s ∧ option_valid_message_prop MC_composite_vlsm input ∧ MC_valid l (s x) input ∧ match l with | init => consistent s | emit => True | receive => match input with | Some m => MC_no_equivocation s m | None => True end end
Ht: (let (si', om') := MC_transition x l (s x) input in (state_update MCVLSM s x si', om')) = (destination, output)
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hdestvalid: MC_non_initial_valid_state destination
Hobs: MC_obs_equivalence s destinationsteps_until_final_composite destination < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
nr_steps: nat
Hind: ∀ y : nat, y < nr_steps → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
s: composite_state MCVLSM
Heqnr_steps: nr_steps = steps_until_final_composite s
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input: option Message
destination: composite_state MCVLSM
output: option Message
Hv: valid_state_prop MC_composite_vlsm s ∧ option_valid_message_prop MC_composite_vlsm input ∧ MC_valid l (s x) input ∧ match l with | init => consistent s | emit => True | receive => match input with | Some m => MC_no_equivocation s m | None => True end end
s0: State
o: option Message
Htx: MC_transition x l (s x) input = (s0, o)
Ht: (state_update MCVLSM s x s0, o) = (destination, output)
Hround: state_round_inc (destination x) > state_round_inc (s x)
Hdest: valid_state_prop MC_composite_vlsm destination
Hrounddest: ∀ i : index, state_round_inc (destination i) ≤ size (st_obs (destination i)) + 1
Hdestvalid: MC_non_initial_valid_state destination
Hobs: MC_obs_equivalence s destinationsteps_until_final_composite destination < nr_stepsindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
Hind: ∀ y : nat, y < steps_until_final_composite s → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input, output: option Message
Hv: valid_state_prop MC_composite_vlsm s ∧ option_valid_message_prop MC_composite_vlsm input ∧ MC_valid l (s x) input ∧ match l with | init => consistent s | emit => True | receive => match input with | Some m => MC_no_equivocation s m | None => True end end
s0: State
Htx: MC_transition x l (s x) input = (s0, output)
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hdestvalid: MC_non_initial_valid_state (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hdest: valid_state_prop MC_composite_vlsm (state_update MCVLSM s x s0)
Hround: state_round_inc (state_update MCVLSM s x s0 x) > state_round_inc (s x)
Ht: (state_update MCVLSM s x s0, output) = (state_update MCVLSM s x s0, output)steps_until_final_composite (state_update MCVLSM s x s0) < steps_until_final_composite sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
Hind: ∀ y : nat, y < steps_until_final_composite s → ∀ s : composite_state MCVLSM, y = steps_until_final_composite s → MC_non_initial_valid_state s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm s sf tr ∧ MC_final_state sf
Hnoninitvalid: MC_non_initial_valid_state s
n: ¬ MC_final_state s
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
l: Label
input, output: option Message
Hv: valid_state_prop MC_composite_vlsm s ∧ option_valid_message_prop MC_composite_vlsm input ∧ MC_valid l (s x) input ∧ match l with | init => consistent s | emit => True | receive => match input with | Some m => MC_no_equivocation s m | None => True end end
s0: State
Htx: MC_transition x l (s x) input = (s0, output)
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hdestvalid: MC_non_initial_valid_state (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hdest: valid_state_prop MC_composite_vlsm (state_update MCVLSM s x s0)
Hround: state_round_inc s0 > state_round_inc (s x)
Ht: (state_update MCVLSM s x s0, output) = (state_update MCVLSM s x s0, output)steps_until_final_composite (state_update MCVLSM s x s0) < steps_until_final_composite sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)steps_until_final_composite (state_update MCVLSM s x s0) < steps_until_final_composite sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) (enum index) < sum_list_with (λ i : index, steps_until_final_component (s i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
Hxelem: x ∈ enum indexsum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) (enum index) < sum_list_with (λ i : index, steps_until_final_component (s i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
Hxelem: x ∈ enum index
Hnodup: NoDup (enum index)sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) (enum index) < sum_list_with (λ i : index, steps_until_final_component (s i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)x ∈ enum index → NoDup (enum index) → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) (enum index) < sum_list_with (λ i : index, steps_until_final_component (s i)) (enum index)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)∀ is : list index, x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) isx ∈ a :: is → NoDup (a :: is) → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) (a :: is) < sum_list_with (λ i : index, steps_until_final_component (s i)) (a :: is)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) isx = a ∨ x ∈ is → (a ∉ is) ∧ NoDup is → steps_until_final_component (state_update MCVLSM s x s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Hx: x ∈ is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s x s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
s0: State
a: index
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s a s0 i) ≤ size (st_obs (state_update MCVLSM s a s0 i)) + 1
Hobs: MC_obs_equivalence s (state_update MCVLSM s a s0)
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s a s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Hx: x ∈ is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s x s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: x ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Hx: x ∈ is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isby lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
x: index
s0: State
Hobs: MC_obs_equivalence s (state_update MCVLSM s x s0)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s x s0 i) ≤ size (st_obs (state_update MCVLSM s x s0 i)) + 1
Hround: state_round_inc s0 > state_round_inc (s x)
a: index
is: list index
IHis: sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Hx: x ∈ is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s x s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
s0: State
a: index
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s a s0 i) ≤ size (st_obs (state_update MCVLSM s a s0 i)) + 1
Hobs: MC_obs_equivalence s (state_update MCVLSM s a s0)
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s a s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
Hrounds: ∀ i : index, state_round_inc (s i) ≤ size (st_obs (s i)) + 1
s0: State
a: index
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s a s0 i) ≤ size (st_obs (state_update MCVLSM s a s0 i)) + 1
Hobs: st_obs (s a) ≡ st_obs (state_update MCVLSM s a s0 a)
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s a s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: ∀ i : index, state_round_inc (state_update MCVLSM s a s0 i) ≤ size (st_obs (state_update MCVLSM s a s0 i)) + 1
Hobs: st_obs (s a) ≡ st_obs (state_update MCVLSM s a s0 a)
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s a s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc (state_update MCVLSM s a s0 a) ≤ size (st_obs (state_update MCVLSM s a s0 a)) + 1
Hobs: st_obs (s a) ≡ st_obs (state_update MCVLSM s a s0 a)
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component (state_update MCVLSM s a s0 a) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component s0 + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) is → steps_until_final_component s0 + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) is → steps_until_final_component s0 + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issteps_until_final_component s0 + sum_list_with (λ i : index, steps_until_final_component (s i)) is < steps_until_final_component (s a) + sum_list_with (λ i : index, steps_until_final_component (s i)) isby destruct s0 as (obs0 & [[round0 status0] |]), (s a) as (obsa & [[rounda statusa] |]); cbn in *; rewrite Hobs; lia.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, size (st_obs (state_update MCVLSM s a s0 i)) + 1 - state_round_inc (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, size (st_obs (s i)) + 1 - state_round_inc (s i)) is
Ha: a ∉ is
Hnodup: NoDup issize (st_obs s0) + 1 - state_round_inc s0 + sum_list_with (λ i : index, size (st_obs (s i)) + 1 - state_round_inc (s i)) is < size (st_obs (s a)) + 1 - state_round_inc (s a) + sum_list_with (λ i : index, size (st_obs (s i)) + 1 - state_round_inc (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
IHis: a ∈ is → NoDup is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is < sum_list_with (λ i : index, steps_until_final_component (s i)) is
Ha: a ∉ is
Hnodup: NoDup issum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list index
Ha: a ∉ issum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
is: list indexa ∉ is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
a0: index
is: list index
IHis: a ∉ is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isa ∉ a0 :: is → steps_until_final_component (state_update MCVLSM s a s0 a0) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = steps_until_final_component (s a0) + sum_list_with (λ i : index, steps_until_final_component (s i)) isindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
a0: index
is: list index
IHis: a ∉ is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) isa ≠ a0 ∧ a ∉ is → steps_until_final_component (state_update MCVLSM s a s0 a0) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = steps_until_final_component (s a0) + sum_list_with (λ i : index, steps_until_final_component (s i)) isby rewrite IHis; [state_update_simpl |]. Qed.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
s: composite_state MCVLSM
a: index
Hrounds: state_round_inc (s a) ≤ size (st_obs (s a)) + 1
s0: State
Hround: state_round_inc s0 > state_round_inc (s a)
Hrounddest: state_round_inc s0 ≤ size (st_obs s0) + 1
Hobs: st_obs (s a) ≡ st_obs s0
a0: index
is: list index
IHis: a ∉ is → sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = sum_list_with (λ i : index, steps_until_final_component (s i)) is
Hdiff: a ≠ a0
Hnotelem: a ∉ issteps_until_final_component (state_update MCVLSM s a s0 a0) + sum_list_with (λ i : index, steps_until_final_component (state_update MCVLSM s a s0 i)) is = steps_until_final_component (s a0) + sum_list_with (λ i : index, steps_until_final_component (s i)) is
This corollary states that final states are reachable from initial consistent states.
This is weaker then the MC_safety result, and is proved using it.
index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: indexcomposite_initial_state_prop MCVLSM s → consistent s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: indexcomposite_initial_state_prop MCVLSM s → consistent s → ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: composite_initial_state_prop MCVLSM s
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}s i = mkSt (st_obs (s i)) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}s i = mkSt (st_obs (s i)) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
st_obs0: indexSet
st_rs0: option RoundStatus
Hsi: s i = mkSt st_obs0 st_rs0mkSt st_obs0 st_rs0 = mkSt (st_obs (mkSt st_obs0 st_rs0)) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, st_rs (s n) = None
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
st_obs0: indexSet
st_rs0: option RoundStatus
Hsi: s i = mkSt st_obs0 st_rs0mkSt st_obs0 st_rs0 = mkSt st_obs0 Noneby rewrite Hinit.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, st_rs (s n) = None
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
st_obs0: indexSet
st_rs0: option RoundStatus
Hsi: s i = mkSt st_obs0 st_rs0mkSt st_obs0 (st_rs (s i)) = mkSt st_obs0 Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}valid_state_prop MC_composite_vlsm sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)by apply initial_state_is_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}valid_state_prop MC_composite_vlsm sby apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}option_valid_message_prop MC_composite_vlsm Noneby rewrite Hsi; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}(let (si', om') := MC_transition_clause_1 i (st_obs (s i)) (size (st_obs (s i))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) None)) (Some (mkRS 0 muddy))), None)by rewrite e.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}(let (si', om') := match size (st_obs (s i)) with | 0 => (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
Hvalids': MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)¬ composite_initial_state_prop MCVLSM (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))by eapply input_valid_transition_destination.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)¬ composite_initial_state_prop MCVLSM (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))by state_update_simpl; cbn.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)st_rs (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))) i) ≠ Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
Hvalids': MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))))∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
Hvalids': ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr ∧ MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr
Hsf: MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_init_to MC_composite_vlsm s sf ([item] ++ tr) ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s sf ([item] ++ tr)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) [item]by apply finite_valid_trace_from_to_singleton.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
e: size (st_obs (s i)) = 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)))) [{| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 muddy))); output := None |}]index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}valid_state_prop MC_composite_vlsm sindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}option_valid_message_prop MC_composite_vlsm Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)by apply initial_state_is_valid.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}valid_state_prop MC_composite_vlsm sby apply option_valid_message_None.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}option_valid_message_prop MC_composite_vlsm Noneby rewrite Hsi; constructor.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}MC_valid init (s i) Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}(let (si', om') := MC_transition i init (s i) None in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}(let (si', om') := MC_transition_clause_1 i (st_obs (s i)) (size (st_obs (s i))) in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (mkSt (st_obs (s i)) None)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}(let (si', om') := match size (st_obs (s i)) with | 0 => (mkSt (st_obs (s i)) (Some (mkRS 0 muddy)), None) | S _ => (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)), None) end in (state_update MCVLSM s i si', om')) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)by inversion H9.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n0: nat
H10: size (st_obs (s i)) = S n0
n: S n0 ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
s0: State
o: option Message
H9: (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)), None) = (s0, o)(state_update MCVLSM s i s0, o) = (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
Hvalids': MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)¬ composite_initial_state_prop MCVLSM (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))by eapply input_valid_transition_destination.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)valid_state_prop MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)¬ composite_initial_state_prop MCVLSM (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))by state_update_simpl; cbn.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)st_rs (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))) i) ≠ Noneindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
Hvalids': MC_non_initial_valid_state (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))))∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
Hvalids': ∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr ∧ MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr
Hsf: MC_final_state sf∃ (tr : list (composite_transition_item MCVLSM)) (sf : composite_state MCVLSM), finite_valid_trace_init_to MC_composite_vlsm s sf tr ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_init_to MC_composite_vlsm s sf ([item] ++ tr) ∧ MC_final_state sfindex: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s sf ([item] ++ tr)index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) [item]by apply finite_valid_trace_from_to_singleton. Qed. End sec_muddy.index: Type
EqDecision0: EqDecision index
H: finite.Finite index
H0: Inhabited index
indexSet: Type
H1: ElemOf index indexSet
H2: Empty indexSet
H3: Singleton index indexSet
H4: Union indexSet
H5: Intersection indexSet
H6: Difference indexSet
H7: Elements index indexSet
EqDecision1: EqDecision index
H8: FinSet index indexSet
Hindex: length (enum index) > 1
s: composite_state MCVLSM
i: index
Hinit: ∀ n : index, MC_initial_state_prop (s n)
Hnempty: MuddyUnion s ≢ ∅
Hcons: ∀ n : index, st_obs (s n) ≡ MuddyUnion s ∖ {[n]}
Hsi: s i = mkSt (st_obs (s i)) None
n: size (st_obs (s i)) ≠ 0
item: transition_item
Heqitem: item = {| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}
Hvalidtr: input_valid_transition MC_composite_vlsm (existT i init) (s, None) (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))), None)
tr: list (composite_transition_item MCVLSM)
sf: composite_state MCVLSM
Htr: finite_valid_trace_from_to MC_composite_vlsm (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) sf tr
Hsf: MC_final_state sffinite_valid_trace_from_to MC_composite_vlsm s (state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided)))) [{| l := existT i init; input := None; destination := state_update MCVLSM s i (mkSt (st_obs (s i)) (Some (mkRS 0 undecided))); output := None |}]