Built with Alectryon, running Coq+SerAPI v8.18.0+0.18.1. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl.
[Loading ML file ring_plugin.cmxs (using legacy method) ... done]
[Loading ML file zify_plugin.cmxs (using legacy method) ... done]
[Loading ML file micromega_plugin.cmxs (using legacy method) ... done]
[Loading ML file btauto_plugin.cmxs (using legacy method) ... done]
[Loading ML file coq-itauto.plugin ... done]
From stdpp Require Import prelude finite fin_sets fin_maps nmap. From Coq Require Import Streams FunctionalExtensionality Eqdep_dec. From VLSM.Lib Require Import Preamble ListExtras FinSetExtras. From VLSM.Core Require Import VLSM Plans VLSMProjections Composition. From VLSM.Examples Require Import Consensus.

Paxos: Specification of Consensus by Voting

A specification of a consensus algorithm where a set of nodes, the acceptors, agree on a value by voting.
The algorithm uses numbered rounds of voting. Each acceptor casts at most one vote per round. A value is chosen when a sufficient set of acceptors, a quorum, have voted for the same value in the same round.
This specification is closer to something implementable than the Consensus module, but still has some restrictions on the allowed transitions of an acceptor that can only be checked with access to the states of all acceptors. We model this as a composite VLSM with a composition constraint.
Ballot numbers
Definition Ballot := N.
Ballot numbers extended with a minimal element "-1" that comes before others
Definition Ballot' := option Ballot.

Definition Ballot'_to_Z (b : Ballot') : Z := from_option Z.of_N (-1)%Z b.
Definition Ballot_to_Z (b : Ballot) : Z := Z.of_N b.

Coercion Ballot_to_Z : Ballot >-> Z.
Coercion Ballot'_to_Z : Ballot' >-> Z.

Ltac ballot_lia := unfold Ballot_to_Z in *; lia.


wf (λ x y : Ballot, (x < y)%Z)

wf (λ x y : Ballot, (x < y)%Z)

wf N.lt → wf (λ x y : Ballot, (x < y)%Z)
by apply (wf_projected N.lt id), N2Z.inj_lt. Qed. Definition Ballot_peano_ind : forall (P : Ballot -> Prop), P 0%N -> (forall (b : Ballot), P b -> P (N.succ b)) -> forall b, P b := N.peano_ind.
ballot-indexed-map
Definition Bmap : Type -> Type := @Nmap.

#[export] Instance Bmap_eq_dec :  {A : Type}, EqDecision A -> EqDecision (Bmap A) := _.

#[export] Instance Bfmap : FMap Bmap := _.
#[export] Instance Blookup :  {A : Type}, Lookup Ballot A (Bmap A) := _.
#[export] Instance Bempty :  {A : Type}, Empty (Bmap A) := _.
#[export] Instance Bpartial_alter :  {A : Type}, PartialAlter Ballot A (Bmap A) := _.
#[export] Instance Bomap : OMap Bmap := _.
#[export] Instance Bmerge : Merge Bmap := _.
#[export] Instance FinMap_Ballot_Bmap : FinMap Ballot Bmap := _.

Notation Bset := (mapset.mapset Bmap).

#[export] Instance Bmap_dom {A : Type} : Dom (Bmap A) Bset := _.
#[export] Instance FinMapDom_Bmap_Bset : fin_map_dom.FinMapDom Ballot Bmap Bset := _.

b: Ballot'
x: Ballot

(b : Z) = (x : Z) ↔ b = Some x
b: Ballot'
x: Ballot

(b : Z) = (x : Z) ↔ b = Some x
by destruct b; cbn; [rewrite (inj_iff Z.of_N) |]; split; (congruence || ballot_lia). Qed. Section sec_voting_spec. Context (Value : Type) (VSet : Type) `{FinSet Value VSet} {VSDec : RelDecision (∈@{VSet})} . Section sec_acceptor_vlsm. Record acceptor_state : Type := { votes : Bmap VSet; maxBal : Ballot'; }.
Defining several predicates over single voting states.
Section sec_acceptor_state_predicates.

Context
  (AState : acceptor_state)
  .

Definition voted_for (b : Ballot) (v : Value) := v ∈ votes AState !!! b.

#[export] Instance voted_for_dec (b : Ballot) (v : Value) : Decision (voted_for b v) :=
  VSDec _ _.

Definition did_not_vote_in (b : Ballot) : Prop :=
  forall v : Value, ~ voted_for b v.

End sec_acceptor_state_predicates.
Acceptor labels:
Inductive acceptor_label : Type :=
| SkipToRound (b : Ballot)
| Vote (b : Ballot) (v : Value).

#[export] Program Instance acceptor_label_dec : EqDecision acceptor_label :=
  fun x y =>
    match x, y with
    | SkipToRound x_b, SkipToRound y_b =>
        if decide (x_b = y_b) then left _ else right _
    | Vote x_b x_v, Vote y_b y_v =>
        if decide (x_b = y_b) then
          if decide (x_v = y_v) then left _
          else right _
        else right _
    | SkipToRound _, Vote _ _ => right _
    | Vote _ _, SkipToRound _ => right _
    end.
Solve Obligations with congruence.

Definition acceptor_type : VLSMType False :=
{|
  state := acceptor_state;
  label := acceptor_label;
|}.

Definition acceptor_initial : acceptor_state :=
{|
  votes := ∅;
  maxBal := None;
|}.

Definition acceptor_initial_state_prop (s : acceptor_state) : Prop :=
  s = acceptor_initial.

Definition acceptor_valid : acceptor_label -> acceptor_state * option False -> Prop :=
  fun l '(s, _) =>
    match l with
    | SkipToRound b => (maxBal s < b)%Z
    | Vote b v => (maxBal s <= b)%Z /\ did_not_vote_in s b
    end.

Definition acceptor_transition
  : acceptor_label -> acceptor_state * option False -> acceptor_state * option False :=
  fun l '(s, _) =>
    let s_votes := votes s in
    match l with
    | SkipToRound b => ({| votes := s_votes; maxBal := Some b; |}, None)
    | Vote b v => ({| votes := mmap_insert b v s_votes; maxBal := Some b; |}, None)
    end.

Definition acceptor_machine : VLSMMachine acceptor_type :=
{|
  initial_state_prop := (.= acceptor_initial);
  s0 := populate (exist _ acceptor_initial eq_refl);
  initial_message_prop := (fun _ => False);
  transition := acceptor_transition;
  valid := acceptor_valid;
|}.

Definition acceptor_vlsm : VLSM False := mk_vlsm acceptor_machine.

End sec_acceptor_vlsm.

(* Little predicates *)
Definition voted_none_but (sa : acceptor_state) (b : Ballot) (v : Value) :=
  forall w : Value, voted_for sa b w -> w = v.

Definition vote_committed (sa : acceptor_state) (b : Ballot) := (maxBal sa > b)%Z.

Context
  (Acceptor : Type) `{finite.Finite Acceptor}
  (ASet : Type) `{FinSet Acceptor ASet}
  .

Section sec_voting_vlsm.

Context
  (Quorum : ASet -> Prop)
  (IM := fun (_ : Acceptor) => acceptor_vlsm)
  .

#[local] Notation composition_constraint :=
  (composite_label IM -> composite_state IM * option False -> Prop).

Definition allQuorum (Q : sig Quorum) (P : Acceptor -> Prop) : Prop :=
  forall a : Acceptor, a ∈ `Q -> P a.

Definition consensus_blocking_quorum
  (s : composite_state IM) (b : Ballot) (v : Value) (Q : sig Quorum) : Prop :=
    allQuorum Q (fun a : Acceptor => vote_committed (s a) b)
    /\ allQuorum Q (fun a : Acceptor => voted_none_but (s a) b v).

Definition SafeAt (s : composite_state IM) (v : Value) (b : Ballot) : Prop :=
  forall (d : Ballot), (d < b)%Z -> exists Q : sig Quorum, consensus_blocking_quorum s d v Q.
SafeAt as defined in the paper
Inductive SafeAt_alt (s : composite_state IM) (v : Value) : Ballot -> Prop :=
| SafeStart : SafeAt_alt s v 0%N
| SafeQuorum : forall (b : Ballot) (Q : sig Quorum) (c' : Ballot'),
    (c' < b)%Z ->
    allQuorum Q (fun a : Acceptor => (maxBal (s a) >= b)%Z) ->
    (forall (d : Ballot), (c' < d)%Z -> (d < b)%Z ->
      allQuorum Q (fun a => did_not_vote_in (s a) d)) ->
    (c' <> None ->
     SafeAt_alt s v (default 0%N c')
       /\ allQuorum Q (fun a : Acceptor => voted_none_but (s a) (default 0%N c') v)) ->
    SafeAt_alt s v b.

(*
  I think SafeAt is equivalent to saying that at each earlier step c there is a quorum of
  nodes that have passed round c and only voted for v or for nothing at round c.
*)

Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt_alt s v b → SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt_alt s v b → SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot

SafeAt_alt s v b → SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y

SafeAt_alt s v b → SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Hsafe: SafeAt_alt s v b

SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_d_lt_c: (c' > d)%Z
Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

allQuorum Q (λ a : Acceptor, vote_committed (s a) d)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z
allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

allQuorum Q (λ a : Acceptor, vote_committed (s a) d)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z
a: Acceptor
Ha: a ∈ `Q

(maxBal (s a) >= b)%Z → vote_committed (s a) d
by unfold vote_committed; lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_use_Q: (c' ≤ d)%Z

allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H16: (c' < d)%Z

allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
d: Ballot
Hc': (Some d < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hc_votes: Some d ≠ None → SafeAt_alt s v (default 0%N (Some d)) ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N (Some d)) v)
Hno_vote: d0 : Ballot, (Some d < d0)%Z → (d0 < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d0)
Hd: (d < b)%Z
allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H16: (c' < d)%Z

allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H16: (c' < d)%Z
a: Acceptor
Ha: a ∈ `Q
w: Value
Hw: voted_for (s a) d w

w = v
by apply Hno_vote in Hw.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
d: Ballot
Hc': (Some d < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hc_votes: Some d ≠ None → SafeAt_alt s v (default 0%N (Some d)) ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N (Some d)) v)
Hno_vote: d0 : Ballot, (Some d < d0)%Z → (d0 < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d0)
Hd: (d < b)%Z

allQuorum Q (λ a : Acceptor, voted_none_but (s a) d v)
by destruct Hc_votes as [_ Hc_votes].
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c': Ballot'
Hc': (c' < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (c' < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: c' ≠ None → SafeAt_alt s v (default 0%N c') ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) (default 0%N c') v)
d: Ballot
Hd: (d < b)%Z
H_d_lt_c: (c' > d)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c: Ballot
Hc': (Z.of_N c < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, (Z.of_N c < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: Some c ≠ None → SafeAt_alt s v c ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) c v)
d: Ballot
Hd: (d < b)%Z
H_d_lt_c: (Z.of_N c > d)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: y : Ballot, (y < b)%Z → SafeAt_alt s v y → SafeAt s v y
Q: {x : ASet | Quorum x}
c: Ballot
Hc': ((c : Z) < b)%Z
H_maxBal: allQuorum Q (λ a : Acceptor, (maxBal (s a) >= b)%Z)
Hno_vote: d : Ballot, ((c : Z) < d)%Z → (d < b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Hc_votes: Some c ≠ None → SafeAt_alt s v c ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) c v)
d: Ballot
Hd: (d < b)%Z
H_d_lt_c: ((c : Z) > d)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum s d v Q
by apply (IHb c Hc'); [apply Hc_votes | lia]. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt s v b → SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt s v b → SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot

SafeAt s v b → SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)

SafeAt_alt s v (N.succ b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

SafeAt_alt s v (N.succ b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

(Z.of_N b < N.succ b)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
allQuorum Q (λ a : Acceptor, (maxBal (s a) >= N.succ b)%Z)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
d : Ballot, (Z.of_N b < d)%Z → (d < N.succ b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
Some b ≠ None → SafeAt_alt s v b ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

(Z.of_N b < N.succ b)%Z
by ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

allQuorum Q (λ a : Acceptor, (maxBal (s a) >= N.succ b)%Z)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
a: Acceptor
Ha: a ∈ `Q

vote_committed (s a) b → (maxBal (s a) >= N.succ b)%Z
by unfold vote_committed; destruct (maxBal (s a)); simpl; ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

d : Ballot, (Z.of_N b < d)%Z → (d < N.succ b)%Z → allQuorum Q (λ a : Acceptor, did_not_vote_in (s a) d)
by ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

Some b ≠ None → SafeAt_alt s v b ∧ allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)

SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
IHb: SafeAt s v b → SafeAt_alt s v b
H_Sb: SafeAt s v (N.succ b)
Q: {x : ASet | Quorum x}
H_Q_maxBal: allQuorum Q (λ a : Acceptor, vote_committed (s a) b)
H_Q_votes: allQuorum Q (λ a : Acceptor, voted_none_but (s a) b v)
d: Ballot
Hd: (d < b)%Z

(d < N.succ b)%Z
by ballot_lia. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt s v b ↔ SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (v : Value) (b : Ballot), SafeAt s v b ↔ SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot

SafeAt s v b → SafeAt_alt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot
SafeAt_alt s v b → SafeAt s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot

SafeAt s v b → SafeAt_alt s v b
by apply safeAt_alt_rev.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
v: Value
b: Ballot

SafeAt_alt s v b → SafeAt s v b
by apply safeAt_alt_fwd. Qed. Definition voting_composition_constraint : composite_label IM -> composite_state IM * option False -> Prop := fun '(existT i l) '(s, _) => match l with | SkipToRound _ => True (* Local checks suffice here *) | Vote b v => (* Voting needs to check that value v is consistent with all other acceptors. *) (forall j, j <> i -> voted_none_but (s j) b v) /\ SafeAt s v b end. Definition voting_vlsm : VLSM False := composite_vlsm IM voting_composition_constraint. End sec_voting_vlsm.
How to recognize consensus, also for the refinement mapping.
Context
  (Quorum : ASet -> Prop)
  (QDec : forall Q, Decision (Quorum Q))
  (* The key quorum assumption is that any two quorums have a nonempty intersection. *)
  (QA : forall (Q1 Q2 : sig Quorum), exists a : Acceptor, a ∈ `Q1 ∩ `Q2)
  (* For convenience we also assume that any superset of a quorum is a quorum. *)
  (QClosed : Proper (subseteq ==> impl) Quorum)
  (IM : Acceptor -> VLSM False := fun _ => acceptor_vlsm).

Definition ballot_chose (s : composite_state IM) (b : Ballot) (v : Value) : Prop :=
  exists (Q : sig Quorum), allQuorum Quorum Q (fun a : Acceptor => voted_for (s a) b v).

Definition ballots_with_votes (s : composite_state IM) : Nset :=
  union_list ((fun a : Acceptor => dom (votes (s a))) <$> enum Acceptor).

Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ballot_chose s b v → b ∈ ballots_with_votes s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ballot_chose s b v → b ∈ ballots_with_votes s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)) → b ∈ ballots_with_votes s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)

b ∈ ballots_with_votes s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

b ∈ ballots_with_votes s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

X : Nset, X ∈ (λ a : Acceptor, dom (votes (s a))) <$> enum Acceptor ∧ b ∈ X
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

b ∈ dom (votes (s a))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

is_Some (votes (s a) !! b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

voted_for (s a) b v → is_Some (votes (s a) !! b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

v ∈ votes (s a) !!! b → is_Some (votes (s a) !! b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

v ∈ default inhabitant (votes (s a) !! b) → is_Some (votes (s a) !! b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

v ∈ ∅ → is_Some None
by rewrite elem_of_empty. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), Decision (ballot_chose s b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), Decision (ballot_chose s b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet

Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet

Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q

allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v) ↔ Q ⊆ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q

allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v) → Q ⊆ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
Q ⊆ voters → allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q

allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v) → Q ⊆ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
HallQ: allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ Q

a ∈ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
a: Acceptor
HallQ: (λ a : Acceptor, voted_for (s a) b v) a
Ha: a ∈ Q

a ∈ voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
a: Acceptor
HallQ: (λ a : Acceptor, voted_for (s a) b v) a
Ha: a ∈ Q

voted_for (s a) b v ∧ a ∈ enum Acceptor
by split; [| apply elem_of_enum].
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q

Q ⊆ voters → allQuorum Quorum (Q ↾ HQ) (λ a : Acceptor, voted_for (s a) b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
Hsub: Q ⊆ voters
a: Acceptor
Ha: a ∈ `(Q ↾ HQ)

voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
a: Acceptor
Hsub: a ∈ voters
Ha: a ∈ `(Q ↾ HQ)

voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Q: ASet
HQ: Quorum Q
a: Acceptor
Hsub: voted_for (s a) b v ∧ a ∈ enum Acceptor
Ha: a ∈ `(Q ↾ HQ)

voted_for (s a) b v
by apply Hsub.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
yes: Quorum voters

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
no: ¬ Quorum voters
Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
yes: Quorum voters

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
yes: Quorum voters

allQuorum Quorum (voters ↾ yes) (λ a : Acceptor, voted_for (s a) b v)
by apply Hvoters.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
no: ¬ Quorum voters

Decision ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
no: ¬ Quorum voters

¬ ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
no: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)

Quorum voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Q': {x : ASet | Quorum x}
HQ': allQuorum Quorum Q' (λ a : Acceptor, voted_for (s a) b v)

Quorum voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Q': {x : ASet | Quorum x}
HQ': `Q' ⊆ voters

Quorum voters
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
voters:= list_to_set (filter (λ a : Acceptor, voted_for (s a) b v) (enum Acceptor)) : ASet: ASet
Hvoters: Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v) ↔ `Q ⊆ voters
Q': {x : ASet | Quorum x}
HQ': `Q' ⊆ voters

Quorum (`Q') → Quorum voters
by eapply QClosed. Qed. Definition ballot_proposals (s : composite_state IM) (b : Ballot) : VSet := union_list ((fun a => votes (s a) !!! b) <$> enum Acceptor).
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ballot_chose s b v → v ∈ ballot_proposals s b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ballot_chose s b v → v ∈ ballot_proposals s b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state IM) (b : Ballot) (v : Value), ( Q : {x : ASet | Quorum x}, allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)) → v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)

v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: voted_for (s a) b v
Ha: a ∈ `Q

v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: v ∈ votes (s a) !!! b
Ha: a ∈ `Q

v ∈ ⋃ ((λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: v ∈ votes (s a) !!! b
Ha: a ∈ `Q

X : VSet, X ∈ (λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor ∧ v ∈ X
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: v ∈ votes (s a) !!! b
Ha: a ∈ `Q

votes (s a) !!! b ∈ (λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor ∧ v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: v ∈ votes (s a) !!! b
Ha: a ∈ `Q

votes (s a) !!! b ∈ (λ a : Acceptor, votes (s a) !!! b) <$> enum Acceptor
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state IM
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
a: Acceptor
HQ: v ∈ votes (s a) !!! b
Ha: a ∈ `Q

a ∈ enum Acceptor
by apply elem_of_enum. Qed. Definition chosen_by_ballot (s : state (voting_vlsm Quorum)) (b : Ballot) : VSet := set_filter (fun v => ballot_chose s b v) _ (ballot_proposals s b). Definition chosen (s : state (voting_vlsm Quorum)) : VSet := union_list (chosen_by_ballot s <$> elements (ballots_with_votes s)).
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : state (voting_vlsm Quorum)) (v : Value), v ∈ chosen s ↔ ( b : Ballot, ballot_chose s b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : state (voting_vlsm Quorum)) (v : Value), v ∈ chosen s ↔ ( b : Ballot, ballot_chose s b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value

v ∈ chosen s ↔ ( b : Ballot, ballot_chose s b v)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value

v ∈ chosen s → b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
( b : Ballot, ballot_chose s b v) → v ∈ chosen s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value

v ∈ chosen s → b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
Hv: v ∈ chosen s

b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
Hv: v ∈ ⋃ (chosen_by_ballot s <$> elements (ballots_with_votes s))

b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
C_by_b: VSet
HC: C_by_b ∈ chosen_by_ballot s <$> elements (ballots_with_votes s)
Hb: v ∈ C_by_b

b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: v ∈ chosen_by_ballot s b

b : Ballot, ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: v ∈ chosen_by_ballot s b

ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: v ∈ set_filter (λ v : Value, ballot_chose s b v) (ballot_chose_dec s b) (ballot_proposals s b)

ballot_chose s b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v ∧ v ∈ ballot_proposals s b

ballot_chose s b v
by apply Hb.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value

( b : Ballot, ballot_chose s b v) → v ∈ chosen s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v

v ∈ chosen s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

v ∈ chosen s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

v ∈ ⋃ (chosen_by_ballot s <$> elements (ballots_with_votes s))
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

X : VSet, X ∈ chosen_by_ballot s <$> elements (ballots_with_votes s) ∧ v ∈ X
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

chosen_by_ballot s b ∈ chosen_by_ballot s <$> elements (ballots_with_votes s) ∧ v ∈ chosen_by_ballot s b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

v ∈ chosen_by_ballot s b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

v ∈ set_filter (λ v : Value, ballot_chose s b v) (ballot_chose_dec s b) (ballot_proposals s b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
v: Value
b: Ballot
Hb: ballot_chose s b v
Hb_dom: b ∈ ballots_with_votes s

ballot_chose s b v ∧ v ∈ ballot_proposals s b
by apply ballot_chose_in_proposals in Hb as Hv. Qed.
Each acceptor can vote for at most one value in any given ballot
(* VInv1 *)
Definition acceptor_no_conflict_prop (s : state (voting_vlsm Quorum)) : Prop :=
  forall (a : Acceptor) (b : Ballot) (v w : Value),
    voted_for (s a) b v -> voted_for (s a) b w -> v = w.

(* VInv2 *)
Every vote is for a safe value
Definition acceptor_votes_safe_prop (s : state (voting_vlsm Quorum)) : Prop :=
  forall (a : Acceptor) (b : Ballot) (v : Value),
    voted_for (s a) b v -> SafeAt Quorum s v b.

Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → let (i, l) := li in s' = state_update IM s i (acceptor_transition l (s i, om)).1match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → let (i, l) := li in s' = state_update IM s i (acceptor_transition l (s i, om)).1match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i l) (s, om)
Hcv: voting_composition_constraint Quorum (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')

s' = state_update IM s i (acceptor_transition l (s i, om)).1match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i l) (s, om)
Hcv: voting_composition_constraint Quorum (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')

s' = state_update IM s i (acceptor_transition l (s i, om)).1
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i l) (s, om)
Hcv: voting_composition_constraint Quorum (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')
match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i l) (s, om)
Hcv: voting_composition_constraint Quorum (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')

s' = state_update IM s i (acceptor_transition l (s i, om)).1
by destruct l; inversion Htrans.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i l) (s, om)
Hcv: voting_composition_constraint Quorum (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')

match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b: Ballot
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i (SkipToRound b)) (s, om)
Hcv: voting_composition_constraint Quorum (existT i (SkipToRound b)) ( s, om)
Htrans: transition (existT i (SkipToRound b)) (s, om) = (s', om')

(maxBal (s i) < b)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b: Ballot
v: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i (Vote b v)) (s, om)
Hcv: voting_composition_constraint Quorum (existT i (Vote b v)) ( s, om)
Htrans: transition (existT i (Vote b v)) (s, om) = (s', om')
(maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b: Ballot
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i (SkipToRound b)) (s, om)
Hcv: voting_composition_constraint Quorum (existT i (SkipToRound b)) ( s, om)
Htrans: transition (existT i (SkipToRound b)) (s, om) = (s', om')

(maxBal (s i) < b)%Z
by apply Hlv.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b: Ballot
v: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hlv: valid (existT i (Vote b v)) (s, om)
Hcv: voting_composition_constraint Quorum (existT i (Vote b v)) ( s, om)
Htrans: transition (existT i (Vote b v)) (s, om) = (s', om')

(maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b
by split_and!; [apply Hlv.. | apply Hcv]. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → let (i, l) := li in match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → let (i, l) := li in match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hvalid: valid (existT i l) (s, om)
Htrans: transition (existT i l) (s, om) = (s', om')

match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hvalid: valid (existT i l) (s, om)
Htrans: ((transition (existT i l) (s, om)).1, (transition (existT i l) (s, om)).2) = (s', om')

match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hvalid: valid (existT i l) (s, om)
Htrans: ((transition (existT i l) (s, om)).1, (transition (existT i l) (s, om)).2).1 = (s', om').1

match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hvalid: valid (existT i l) (s, om)
Htrans: (transition (existT i l) (s, om)).1 = s'

match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om, om': option False
Hvalid: valid (existT i l) (s, om)

match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ (transition (existT i l) (s, om)).1 = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ (transition (existT i l) (s, om)).1 = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
by destruct l; simpl; repeat split; apply Hvalid. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → a : Acceptor, let (i, l) := li in if decide (a = i) then s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → a : Acceptor, let (i, l) := li in if decide (a = i) then s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: input_valid_transition (voting_vlsm Quorum) (existT i l) (s, om) ( s', om')
a: Acceptor

if decide (a = i) then s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor

if decide (a = i) then state_update IM s i (acceptor_transition l (s i, om)).1 a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else state_update IM s i (acceptor_transition l (s i, om)).1 a = s a
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
e: a = i

state_update IM s i (acceptor_transition l (s i, om)).1 a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor

(acceptor_transition l (s a, om)).1 = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
mb_a: Bmap VSet
votes_a: Ballot'

(acceptor_transition l ({| votes := mb_a; maxBal := votes_a |}, om)).1 = {| votes := match l with | SkipToRound _ => votes {| votes := mb_a; maxBal := votes_a |} | Vote b v => mmap_insert b v (votes {| votes := mb_a; maxBal := votes_a |}) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om': option False
a: Acceptor
mb_a: Bmap VSet
votes_a: Ballot'

(acceptor_transition l ({| votes := mb_a; maxBal := votes_a |}, None)).1 = {| votes := match l with | SkipToRound _ => votes {| votes := mb_a; maxBal := votes_a |} | Vote b v => mmap_insert b v (votes {| votes := mb_a; maxBal := votes_a |}) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |}
by destruct l; simpl. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), voted_for (s' a) b v → let (i, l) := li in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(li : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) li (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), voted_for (s' a) b v → let (i, l) := li in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Hs': s' = state_update IM s i (acceptor_transition l (s i, om)).1
Ht: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
a: Acceptor
b: Ballot
v: Value

voted_for (s' a) b v → match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
a: Acceptor
b: Ballot
v: Value

voted_for (state_update IM s i (acceptor_transition l (s i, om)).1 a) b v → match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
a: Acceptor
b: Ballot
v: Value
e: a = i

voted_for (state_update IM s i (acceptor_transition l (s i, om)).1 a) b v → match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label acceptor_vlsm
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
Ht: match l with | SkipToRound b' => (maxBal (s a) < b')%Z | Vote b v => (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v b end
b: Ballot
v: Value

voted_for (acceptor_transition l (s a, om)).1 b v → match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = a ∧ b = bv ∧ v = vv ∨ (a ≠ a ∨ b ≠ bv) ∧ voted_for (s a) b v end
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
Ht: (maxBal (s a) ≤ b0)%Z ∧ did_not_vote_in (s a) b0 ∧ SafeAt Quorum s v0 b0
b: Ballot
v: Value

voted_for {| votes := mmap_insert b0 v0 (votes (s a)); maxBal := Some b0 |} b v → a = a ∧ b = b0 ∧ v = v0 ∨ (a ≠ a ∨ b ≠ b0) ∧ voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
Ht: (maxBal (s a) ≤ b0)%Z ∧ did_not_vote_in (s a) b0 ∧ SafeAt Quorum s v0 b0
b: Ballot
v: Value

v ∈ <[b0:={[v0]} ∪ votes (s a) !!! b0]> (votes (s a)) !!! b → a = a ∧ b = b0 ∧ v = v0 ∨ (a ≠ a ∨ b ≠ b0) ∧ v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
Ht: (maxBal (s a) ≤ b0)%Z ∧ did_not_vote_in (s a) b0 ∧ SafeAt Quorum s v0 b0
b: Ballot
v: Value
e: b0 = b

v ∈ <[b0:={[v0]} ∪ votes (s a) !!! b0]> (votes (s a)) !!! b → a = a ∧ b = b0 ∧ v = v0 ∨ (a ≠ a ∨ b ≠ b0) ∧ v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
v0: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
b: Ballot
Ht: (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v0 b
v: Value

v ∈ {[v0]} ∪ votes (s a) !!! b → a = a ∧ b = b ∧ v = v0 ∨ (a ≠ a ∨ b ≠ b) ∧ v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
v0: Value
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
a: Acceptor
b: Ballot
Ht: (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v0 b
v: Value

v ∈ {[v0]} ∪ ∅ → a = a ∧ b = b ∧ v = v0 ∨ (a ≠ a ∨ b ≠ b) ∧ v ∈ ∅
by left; set_solver. Qed.
A stronger claim which is that two votes on the same ballot cannot be for different values, even if they are from different acceptors. This is not needed to prove safety, only for liveness.
(* VInv3 *)
Definition ballot_no_conflict_prop (s : state (voting_vlsm Quorum)) : Prop :=
  forall (a1 a2 : Acceptor) (b : Ballot) (v1 v2 : Value),
    voted_for (s a1) b v1 -> voted_for (s a2) b v2 -> v1 = v2.

Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)

ballot_no_conflict_prop s → acceptor_no_conflict_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)

ballot_no_conflict_prop s → acceptor_no_conflict_prop s
by unfold acceptor_no_conflict_prop, ballot_no_conflict_prop; eauto. Qed.
invariance of VInv1
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_no_conflict_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_no_conflict_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

acceptor_no_conflict_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v, w: Value

voted_for (s a) b v → voted_for (s a) b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_no_conflict_prop s
a: Acceptor
b: Ballot
v, w: Value
voted_for (s' a) b v → voted_for (s' a) b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v, w: Value

voted_for (s a) b v → voted_for (s a) b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b: Ballot
v, w: Value

voted_for acceptor_initial b v → voted_for acceptor_initial b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b: Ballot
v, w: Value

¬ voted_for acceptor_initial b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b: Ballot
v, w: Value

v ∉ ∅ !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b: Ballot
v, w: Value

v ∉ inhabitant
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b: Ballot
v, w: Value

v ∉ ∅
by apply not_elem_of_empty.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_no_conflict_prop s
a: Acceptor
b: Ballot
v, w: Value

voted_for (s' a) b v → voted_for (s' a) b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_no_conflict_prop s
a: Acceptor
b: Ballot
v, w: Value
Hv: voted_for (s' a) b v
Hw: voted_for (s' a) b w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_no_conflict_prop s
a: Acceptor
b: Ballot
v, w: Value
Hv: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Hw: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b w | Vote bv vv => a = i ∧ b = bv ∧ w = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b w end

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
a: Acceptor
b: Ballot
v, w: Value
IHHs: voted_for (s a) b v → voted_for (s a) b w → v = w
Hv: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Hw: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b w | Vote bv vv => a = i ∧ b = bv ∧ w = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b w end

v = w
by destruct l as [i []]; [auto | itauto congruence]. Qed.
VInv4 in the paper
Definition maxBal_limits_votes_prop (s : state (voting_vlsm Quorum)) : Prop :=
  forall a (b : Ballot),
    (maxBal (s a) < b)%Z -> did_not_vote_in (s a) b.

Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → maxBal_limits_votes_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → maxBal_limits_votes_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

(a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot

(maxBal (s a) < b)%Z → did_not_vote_in (s a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot
(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot

(maxBal (s a) < b)%Z → did_not_vote_in (s a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

¬ voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

¬ voted_for acceptor_initial b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

v ∉ ∅ !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

v ∉ ∅
by apply not_elem_of_empty.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) (existT i l) (s, om) ( s', om')
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
Hs': s' = state_update IM s i (acceptor_transition l (s i, om)).1
Hl: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
a: Acceptor
Hl: match l with | SkipToRound b' => (maxBal (s a) < b')%Z | Vote b v => (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v b end
Hs': s' = state_update IM s a (acceptor_transition l (s a, om)).1
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
Hs': s' = state_update IM s i (acceptor_transition l (s i, om)).1
Hl: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot
n: a ≠ i
(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
a: Acceptor
Hl: match l with | SkipToRound b' => (maxBal (s a) < b')%Z | Vote b v => (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v b end
Hs': s' = state_update IM s a (acceptor_transition l (s a, om)).1
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
a: Acceptor
Hl: match l with | SkipToRound b' => (maxBal (s a) < b')%Z | Vote b v => (maxBal (s a) ≤ b)%Z ∧ did_not_vote_in (s a) b ∧ SafeAt Quorum s v b end
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(maxBal (acceptor_transition l (s a, om)).1 < b)%Z → did_not_vote_in (acceptor_transition l (s a, om)).1 b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
b': Ballot
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) < b')%Z
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(Z.of_N b' < b)%Z → did_not_vote_in {| votes := votes (s a); maxBal := Some b' |} b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
(Z.of_N bv < b)%Z → did_not_vote_in {| votes := mmap_insert bv vv (votes (s a)); maxBal := Some bv |} b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
b': Ballot
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) < b')%Z
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(Z.of_N b' < b)%Z → did_not_vote_in {| votes := votes (s a); maxBal := Some b' |} b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
b': Ballot
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) < b')%Z
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb': (Z.of_N b' < b)%Z

did_not_vote_in {| votes := votes (s a); maxBal := Some b' |} b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
b': Ballot
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) < b')%Z
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb': (Z.of_N b' < b)%Z

(maxBal (s a) < b)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
b': Ballot
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb': (Z.of_N b' < b)%Z

(maxBal (s a) < b')%Z → (maxBal (s a) < b)%Z
by destruct (maxBal (s a)); simpl; ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot

(Z.of_N bv < b)%Z → did_not_vote_in {| votes := mmap_insert bv vv (votes (s a)); maxBal := Some bv |} b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb: (Z.of_N bv < b)%Z
v: Value

¬ voted_for {| votes := mmap_insert bv vv (votes (s a)); maxBal := Some bv |} b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb: (Z.of_N bv < b)%Z
v: Value

v ∉ <[bv:={[vv]} ∪ votes (s a) !!! bv]> (votes (s a)) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb: (Z.of_N bv < b)%Z
v: Value

v ∉ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb: (Z.of_N bv < b)%Z
v: Value

(maxBal (s a) < b)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': composite_state (λ _ : Acceptor, acceptor_vlsm)
bv: Ballot
vv: Value
om, om': option False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
a: Acceptor
Hl: (maxBal (s a) ≤ bv)%Z ∧ did_not_vote_in (s a) bv ∧ SafeAt Quorum s vv bv
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
b: Ballot
Hb: (Z.of_N bv < b)%Z
v: Value

(maxBal (s a) ≤ bv)%Z → (maxBal (s a) < b)%Z
by destruct (maxBal (s a)); simpl; ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
Hs': s' = state_update IM s i (acceptor_transition l (s i, om)).1
Hl: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot
n: a ≠ i

(maxBal (s' a) < b)%Z → did_not_vote_in (s' a) b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
l: label acceptor_vlsm
om, om': option False
s: state (voting_vlsm Quorum)
Hs': s' = state_update IM s i (acceptor_transition l (s i, om)).1
Hl: match l with | SkipToRound b' => (maxBal (s i) < b')%Z | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ SafeAt Quorum s v b end
IHHs: (a : Acceptor) (b : Ballot), (maxBal (s a) < b)%Z → did_not_vote_in (s a) b
a: Acceptor
b: Ballot
n: a ≠ i

(maxBal (s a) < b)%Z → did_not_vote_in (s a) b
by apply IHHs. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state (λ _ : Acceptor, acceptor_vlsm)) (b : Ballot) (v w : Value), ( Q : {x : ASet | Quorum x}, consensus_blocking_quorum Quorum s b v Q) → ballot_chose s b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(s : composite_state (λ _ : Acceptor, acceptor_vlsm)) (b : Ballot) (v w : Value), ( Q : {x : ASet | Quorum x}, consensus_blocking_quorum Quorum s b v Q) → ballot_chose s b w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
b: Ballot
v, w: Value
Q: {x : ASet | Quorum x}
Hblock: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) b v)
QQ: {x : ASet | Quorum x}
Hchose: allQuorum Quorum QQ (λ a : Acceptor, voted_for (s a) b w)

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
b: Ballot
v, w: Value
Q: {x : ASet | Quorum x}
Hblock: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) b v)
QQ: {x : ASet | Quorum x}
Hchose: allQuorum Quorum QQ (λ a : Acceptor, voted_for (s a) b w)
a: Acceptor
Ha: a ∈ `Q ∩ `QQ

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
b: Ballot
v, w: Value
Q: {x : ASet | Quorum x}
Hblock: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) b v)
QQ: {x : ASet | Quorum x}
Hchose: allQuorum Quorum QQ (λ a : Acceptor, voted_for (s a) b w)
a: Acceptor
Ha_Q: a ∈ `Q
Ha_QQ: a ∈ `QQ

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: composite_state (λ _ : Acceptor, acceptor_vlsm)
b: Ballot
v, w: Value
Q: {x : ASet | Quorum x}
Hblock: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) b v)
QQ: {x : ASet | Quorum x}
Hchose: allQuorum Quorum QQ (λ a : Acceptor, voted_for (s a) b w)
a: Acceptor
Ha_Q: a ∈ `Q
Ha_QQ: a ∈ `QQ

w = v
by eapply Hblock; [| apply Hchose]. Qed. (* does not need invariants yet *)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → (v w : Value) (b c : Ballot), (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → (v w : Value) (b c : Ballot), (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: Ballot

c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: Ballot
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop

c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: Ballot
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop

P b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b

P (N.succ b)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
c: Ballot
Hc: (N.succ b > c)%N
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
c: Ballot
Hc: (b > c)%N
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s b w
v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
c: Ballot
Hc: (b > c)%N
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
c: Ballot
Hc: (b > c)%N
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s c w

SafeAt Quorum s v b
by intros d Hd; apply Hsafe; ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s b w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v, w: Value
b: N
P:= λ b : N, c : Ballot, (b > c)%N → SafeAt Quorum s v b → ballot_chose s c w → v = w: N → Prop
IHb: P b
Hsafe: SafeAt Quorum s v (N.succ b)
Hchosen: ballot_chose s b w

Q : {x : ASet | Quorum x}, consensus_blocking_quorum Quorum s b v Q
by apply Hsafe; ballot_lia. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

Q : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

Q : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
Q: {x : ASet | Quorum x}

a : Acceptor, a ∈ `Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
Q: {x : ASet | Quorum x}
a: Acceptor
Ha: a ∈ `Q ∩ `Q

a : Acceptor, a ∈ `Q
by exists a; apply intersection_idemp. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_no_conflict_prop s → acceptor_votes_safe_prop s → v w : Value, v ∈ chosen s → w ∈ chosen s → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_no_conflict_prop s → acceptor_votes_safe_prop s → v w : Value, v ∈ chosen s → w ∈ chosen s → v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
Hv: v ∈ chosen s
Hw: w ∈ chosen s

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hw: w ∈ chosen s

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: av ∈ `Q

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: av ∈ `Q
aw: Acceptor
Haw: aw ∈ `R

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: aw ∈ `R

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_b_lt_c: (b < c)%N

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_bc: b = c
v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_c_lt_b: (c < b)%N
v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_b_lt_c: (b < c)%N

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_b_lt_c: (c > b)%N

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_b_lt_c: (c > b)%N

w = v
by eapply VT0.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_bc: b = c

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_bc: b = c
a: Acceptor
Ha: a ∈ `Q ∩ `R

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_bc: b = c
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hchose_w: ballot_chose s b w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) b w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) b w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w b
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hchose_w: ballot_chose s b w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) b w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) b w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w b
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R

voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hchose_w: ballot_chose s b w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) b w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) b w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w b
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R
voted_for (s a) b w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hchose_w: ballot_chose s b w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) b w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) b w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w b
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R

voted_for (s a) b v
by apply Hv.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
Hchose_w: ballot_chose s b w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) b w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) b w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w b
a: Acceptor
Ha_Q: a ∈ `Q
Ha_R: a ∈ `R

voted_for (s a) b w
by apply Hw.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_c_lt_b: (c < b)%N

v = w
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hinv1: acceptor_no_conflict_prop s
Hinv2: acceptor_votes_safe_prop s
v, w: Value
b: Ballot
Hchose_v: ballot_chose s b v
c: Ballot
Hchose_w: ballot_chose s c w
Q: {x : ASet | Quorum x}
Hv: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
R: {x : ASet | Quorum x}
Hw: allQuorum Quorum R (λ a : Acceptor, voted_for (s a) c w)
av: Acceptor
Hav: voted_for (s av) b v
aw: Acceptor
Haw: voted_for (s aw) c w
H16: SafeAt Quorum s v b
H17: SafeAt Quorum s w c
H_c_lt_b: (b > c)%N

v = w
by eapply VT0. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s s' : state (voting_vlsm Quorum)) (om om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), vote_committed (s a) b → voted_for (s' a) b v → voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s s' : state (voting_vlsm Quorum)) (om om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), vote_committed (s a) b → voted_for (s' a) b v → voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
a: Acceptor
b: Ballot
v: Value

vote_committed (s a) b → voted_for (s' a) b v → voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: let (i, l) := l in match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
a: Acceptor
b: Ballot
v: Value

vote_committed (s a) b → voted_for (s' a) b v → voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: let (i, l) := l in match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (s' a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: (maxBal (s i) < b0)%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b0 |}
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (s' a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
v0: Value
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: (maxBal (s i) ≤ b0)%Z ∧ did_not_vote_in (s i) b0 ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b0 v0) ∧ SafeAt Quorum s v0 b0 ∧ s' = state_update IM s i {| votes := mmap_insert b0 v0 (votes (s i)); maxBal := Some b0 |}
a: Acceptor
b: Ballot
v: Value
(maxBal (s a) > b)%Z → v ∈ votes (s' a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: (maxBal (s i) < b0)%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b0 |}
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (s' a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
s: state (voting_vlsm Quorum)
om, om': option False
Hb: (maxBal (s i) < b0)%Z
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (state_update IM s i {| votes := votes (s i); maxBal := Some b0 |} a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
s: state (voting_vlsm Quorum)
om, om': option False
Hb: (maxBal (s i) < b0)%Z
a: Acceptor
b: Ballot
v: Value
e: a = i

(maxBal (s a) > b)%Z → v ∈ votes (state_update IM s i {| votes := votes (s i); maxBal := Some b0 |} a) !!! b → v ∈ votes (s a) !!! b
by subst i; rewrite state_update_eq.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
v0: Value
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: (maxBal (s i) ≤ b0)%Z ∧ did_not_vote_in (s i) b0 ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b0 v0) ∧ SafeAt Quorum s v0 b0 ∧ s' = state_update IM s i {| votes := mmap_insert b0 v0 (votes (s i)); maxBal := Some b0 |}
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (s' a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
Hbal: (maxBal (s i) ≤ b0)%Z
Hnv: did_not_vote_in (s i) b0
a: Acceptor
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes (state_update IM s i {| votes := mmap_insert b0 v0 (votes (s i)); maxBal := Some b0 |} a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
Hbal: (maxBal (s i) ≤ b0)%Z
Hnv: did_not_vote_in (s i) b0
a: Acceptor
b: Ballot
v: Value
e: a = i

(maxBal (s a) > b)%Z → v ∈ votes (state_update IM s i {| votes := mmap_insert b0 v0 (votes (s i)); maxBal := Some b0 |} a) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
Hbal: (maxBal (s a) ≤ b0)%Z
b: Ballot
v: Value

(maxBal (s a) > b)%Z → v ∈ votes {| votes := mmap_insert b0 v0 (votes (s a)); maxBal := Some b0 |} !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
mb_a: Ballot
Hbal: (Some mb_a ≤ b0)%Z
b: Ballot
v: Value

(Some mb_a > b)%Z → v ∈ votes {| votes := mmap_insert b0 v0 (votes (s a)); maxBal := Some b0 |} !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
mb_a: Ballot
Hbal: (Z.of_N mb_a ≤ b0)%Z
b: Ballot
v: Value

(Z.of_N mb_a > b)%Z → v ∈ mmap_insert b0 v0 (votes (s a)) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
mb_a: Ballot
Hbal: (Z.of_N mb_a ≤ b0)%Z
b: Ballot
v: Value
Hb: (Z.of_N mb_a > b)%Z

v ∈ mmap_insert b0 v0 (votes (s a)) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
mb_a: Ballot
Hbal: (Z.of_N mb_a ≤ b0)%Z
b: Ballot
v: Value
Hb: (Z.of_N mb_a > b)%Z
H16: b ≠ b0

v ∈ mmap_insert b0 v0 (votes (s a)) !!! b → v ∈ votes (s a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
b0: Ballot
v0: Value
s: state (voting_vlsm Quorum)
om, om': option False
a: Acceptor
Hnv: did_not_vote_in (s a) b0
mb_a: Ballot
Hbal: (Z.of_N mb_a ≤ b0)%Z
b: Ballot
v: Value
Hb: (Z.of_N mb_a > b)%Z
H16: b ≠ b0

v0 = v ∧ b0 = b ∨ v ∈ votes (s a) !!! b → v ∈ votes (s a) !!! b
by itauto. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s s' : state (voting_vlsm Quorum)) (om om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (b : Value) (v : Ballot), SafeAt Quorum s b v → SafeAt Quorum s' b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s s' : state (voting_vlsm Quorum)) (om om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (b : Value) (v : Ballot), SafeAt Quorum s b v → SafeAt Quorum s' b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z

Q : {x : ASet | Quorum x}, consensus_blocking_quorum Quorum s' d b Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
Hbal: allQuorum Quorum Q (λ a : Acceptor, vote_committed (s a) d)
Hvotes: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) d b)

Q : {x : ASet | Quorum x}, consensus_blocking_quorum Quorum s' d b Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
Hbal: allQuorum Quorum Q (λ a : Acceptor, vote_committed (s a) d)
Hvotes: allQuorum Quorum Q (λ a : Acceptor, voted_none_but (s a) d b)

consensus_blocking_quorum Quorum s' d b Q
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

vote_committed (s' a) d
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
voted_none_but (s' a) d b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

vote_committed (s' a) d
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

vote_committed (s a) d → vote_committed (s' a) d
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

(maxBal (s a) > d)%Z → (maxBal (s' a) > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
Hstep: let (i, l) := l in if decide (a = i) then s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a

(maxBal (s a) > d)%Z → (maxBal (s' a) > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) (existT i l) (s, om) ( s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
Hstep: if decide (a = i) then s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a

(maxBal (s a) > d)%Z → (maxBal (s' a) > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) (existT i l) (s, om) ( s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
e: a = i
Hstep: s' a = {| votes := match l with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match l with | SkipToRound b' => b' | Vote b _ => b end |}

(maxBal (s a) > d)%Z → (maxBal (s' a) > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) (existT i l) (s, om) ( s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
e: a = i

(maxBal (s a) > d)%Z → (Z.of_N match l with | SkipToRound b' => b' | Vote b _ => b end > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
e: a = i

(maxBal (s a) > d)%Z → (Z.of_N match l with | SkipToRound b' => b' | Vote b _ => b end > d)%Z
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
i: Acceptor
l: label acceptor_vlsm
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: match l with | SkipToRound b' => (maxBal (s i) < b')%Z ∧ s' = state_update IM s i {| votes := votes (s i); maxBal := Some b' |} | Vote b v => (maxBal (s i) ≤ b)%Z ∧ did_not_vote_in (s i) b ∧ ( j : Acceptor, j ≠ i → voted_none_but (s j) b v) ∧ SafeAt Quorum s v b ∧ s' = state_update IM s i {| votes := mmap_insert b v (votes (s i)); maxBal := Some b |} end
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
Ha: i ∈ `Q
Hvotes: voted_none_but (s i) d b

(maxBal (s i) > d)%Z → (Z.of_N match l with | SkipToRound b' => b' | Vote b _ => b end > d)%Z
by destruct l; simpl in Ht |- *; destruct Ht as [Ht _]; revert Ht; destruct (maxBal (s i)); simpl; ballot_lia.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

voted_none_but (s' a) d b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q

w : Value, voted_for (s' a) d w → w = b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s, s': state (voting_vlsm Quorum)
om, om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Value
v: Ballot
Hsafe: SafeAt Quorum s b v
d: Ballot
Hd: (d < v)%Z
Q: {x : ASet | Quorum x}
a: Acceptor
Hbal: vote_committed (s a) d
Hvotes: voted_none_but (s a) d b
Ha: a ∈ `Q
w: Value
Hw: voted_for (s' a) d w

w = b
by apply Hvotes, (prev_vote_committed Ht). Qed. (* Need invariance of VInv2 VInv3 VInv4 *)
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_votes_safe_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → acceptor_votes_safe_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

acceptor_votes_safe_prop s
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

voted_for (s a) b v → SafeAt Quorum s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_votes_safe_prop s
a: Acceptor
b: Ballot
v: Value
voted_for (s' a) b v → SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

voted_for (s a) b v → SafeAt Quorum s v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

¬ voted_for (s a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

¬ voted_for acceptor_initial b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

v ∉ ∅ !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: initial_state_prop s
a: Acceptor
b: Ballot
v: Value

v ∉ ∅
by apply not_elem_of_empty.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_votes_safe_prop s
a: Acceptor
b: Ballot
v: Value

voted_for (s' a) b v → SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_votes_safe_prop s
a: Acceptor
b: Ballot
v: Value
Hvote: voted_for (s' a) b v

SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_votes_safe_prop s
a: Acceptor
b: Ballot
v: Value
Hvote: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end

SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
IHHs: acceptor_votes_safe_prop s
a: Acceptor
b: Ballot
v: Value
Hvote: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Hpres: SafeAt Quorum s v b → SafeAt Quorum s' v b

SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
l: label (voting_vlsm Quorum)
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
a: Acceptor
b: Ballot
v: Value
IHHs: voted_for (s a) b v → SafeAt Quorum s v b
Hvote: let (i, l) := l in match l with | SkipToRound _ => voted_for (s a) b v | Vote bv vv => a = i ∧ b = bv ∧ v = vv ∨ (a ≠ i ∨ b ≠ bv) ∧ voted_for (s a) b v end
Hpres: SafeAt Quorum s v b → SafeAt Quorum s' v b

SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
b': Ballot
w: Value
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) (existT i (Vote b' w)) ( s, om) (s', om')
a: Acceptor
b: Ballot
v: Value
IHHs: voted_for (s a) b v → SafeAt Quorum s v b
Hvote: a = i ∧ b = b' ∧ v = w ∨ (a ≠ i ∨ b ≠ b') ∧ voted_for (s a) b v
Hpres: SafeAt Quorum s v b → SafeAt Quorum s' v b

SafeAt Quorum s' v b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s': state (voting_vlsm Quorum)
i: Acceptor
b': Ballot
w: Value
om, om': option False
s: state (voting_vlsm Quorum)
Ht: input_valid_transition (voting_vlsm Quorum) (existT i (Vote b' w)) ( s, om) (s', om')
a: Acceptor
b: Ballot
v: Value
IHHs: voted_for (s a) b v → SafeAt Quorum s v b
Hvote: a = i ∧ b = b' ∧ v = w ∨ (a ≠ i ∨ b ≠ b') ∧ voted_for (s a) b v
Hpres: SafeAt Quorum s v b → SafeAt Quorum s' v b
Hw: SafeAt Quorum s w b'

SafeAt Quorum s' v b
by decompose [and or] Hvote; subst; auto. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (a : Acceptor) (b : Ballot) (v : Value), voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
a: Acceptor
b: Ballot
v: Value

voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
i: Acceptor
li: acceptor_label
Ht: input_valid_transition (voting_vlsm Quorum) (sigT_of_prod (i, li)) ( s, om) (s', om')
a: Acceptor
b: Ballot
v: Value

voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
i: Acceptor
li: acceptor_label
a: Acceptor
Ht: if decide (a = i) then s' a = {| votes := match li with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match li with | SkipToRound b' => b' | Vote b _ => b end |} else s' a = s a
b: Ballot
v: Value

voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
i: Acceptor
li: acceptor_label
a: Acceptor
e: a = i
Ht: s' a = {| votes := match li with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match li with | SkipToRound b' => b' | Vote b _ => b end |}
b: Ballot
v: Value

voted_for (s a) b v → voted_for (s' a) b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
i: Acceptor
li: acceptor_label
a: Acceptor
e: a = i
Ht: s' a = {| votes := match li with | SkipToRound _ => votes (s a) | Vote b v => mmap_insert b v (votes (s a)) end; maxBal := Some match li with | SkipToRound b' => b' | Vote b _ => b end |}
b: Ballot
v: Value

v ∈ votes (s a) !!! b → v ∈ votes (s' a) !!! b
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
i: Acceptor
b0: Ballot
v0: Value
a: Acceptor
e: a = i
Ht: s' a = {| votes := mmap_insert b0 v0 (votes (s a)); maxBal := Some b0 |}
b: Ballot
v: Value

v ∈ votes (s a) !!! b → v ∈ mmap_insert b0 v0 (votes (s a)) !!! b
by rewrite elem_of_mmap_insert; itauto. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (b : Ballot) (v : Value), ballot_chose s b v → ballot_chose s' b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

(l : label (voting_vlsm Quorum)) (s : state (voting_vlsm Quorum)) (om : option False) (s' : state (voting_vlsm Quorum)) (om' : option False), input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om') → (b : Ballot) (v : Value), ballot_chose s b v → ballot_chose s' b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)

ballot_chose s' b v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
l: label (voting_vlsm Quorum)
s: state (voting_vlsm Quorum)
om: option False
s': state (voting_vlsm Quorum)
om': option False
Ht: input_valid_transition (voting_vlsm Quorum) l (s, om) (s', om')
b: Ballot
v: Value
Q: {x : ASet | Quorum x}
HQ: allQuorum Quorum Q (λ a : Acceptor, voted_for (s a) b v)
a: Acceptor
Ha: a ∈ `Q

voted_for (s' a) b v
by eapply vote_stable, HQ. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

(b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

( (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), (b1 < b2)%Z → (λ (b3 : Ballot) (v3 : Value) (b4 : Ballot) (v4 : Value), ballot_chose s b3 v3 → ballot_chose s b4 v4 → v3 = v4) b1 v1 b2 v2) → (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
(b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), (b1 < b2)%Z → (λ (b3 : Ballot) (v3 : Value) (b4 : Ballot) (v4 : Value), ballot_chose s b3 v3 → ballot_chose s b4 v4 → v3 = v4) b1 v1 b2 v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

( (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), (b1 < b2)%Z → (λ (b3 : Ballot) (v3 : Value) (b4 : Ballot) (v4 : Value), ballot_chose s b3 v3 → ballot_chose s b4 v4 → v3 = v4) b1 v1 b2 v2) → (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
Hwlog: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), (b1 < b2)%Z → (λ (b3 : Ballot) (v3 : Value) (b4 : Ballot) (v4 : Value), ballot_chose s b3 v3 → ballot_chose s b4 v4 → v3 = v4) b1 v1 b2 v2
b1: Ballot
v1: Value
b2: Ballot
v2: Value
Hv1: ballot_chose s b1 v1
Hv2: ballot_chose s b2 v2

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
Hv1: ballot_chose s b1 v1
Hv2: ballot_chose s b2 v2
Heq: b1 = b2

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1, v2: Value
Hv1: ballot_chose s b1 v1
Hv2: ballot_chose s b1 v2

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1, v2: Value
Q_v1: {x : ASet | Quorum x}
Hv1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
Hv2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b1 v2)

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1, v2: Value
Q_v1: {x : ASet | Quorum x}
Hv1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
Hv2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b1 v2)
a: Acceptor
Ha_Q1: a ∈ `Q_v1
Ha_Q2: a ∈ `Q_v2

v1 = v2
by eapply (inv_acceptor_no_conflict _ Hs); eauto.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s

(b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), (b1 < b2)%Z → (λ (b3 : Ballot) (v3 : Value) (b4 : Ballot) (v4 : Value), ballot_chose s b3 v3 → ballot_chose s b4 v4 → v3 = v4) b1 v1 b2 v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
H_lt: (b1 < b2)%Z
Q_v1: {x : ASet | Quorum x}
H_v1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
H_v2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b2 v2)

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
H_lt: (b1 < b2)%Z
Q_v1: {x : ASet | Quorum x}
H_v1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
H_v2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b2 v2)
a2: Acceptor
Ha2: a2 ∈ `Q_v2

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
H_lt: (b1 < b2)%Z
Q_v1: {x : ASet | Quorum x}
H_v1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
H_v2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b2 v2)
a2: Acceptor
Ha2: a2 ∈ `Q_v2
Q12: {x : ASet | Quorum x}
H_b1_v2: consensus_blocking_quorum Quorum s b1 v2 Q12

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
H_lt: (b1 < b2)%Z
Q_v1: {x : ASet | Quorum x}
H_v1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
H_v2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b2 v2)
a2: Acceptor
Ha2: a2 ∈ `Q_v2
Q12: {x : ASet | Quorum x}
H_b1_v2: consensus_blocking_quorum Quorum s b1 v2 Q12
a: Acceptor
Ha_Q1: a ∈ `Q_v1
Ha_Q12: a ∈ `Q12

v1 = v2
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
b1: Ballot
v1: Value
b2: Ballot
v2: Value
H_lt: (b1 < b2)%Z
Q_v1: {x : ASet | Quorum x}
H_v1: allQuorum Quorum Q_v1 (λ a : Acceptor, voted_for (s a) b1 v1)
Q_v2: {x : ASet | Quorum x}
H_v2: allQuorum Quorum Q_v2 (λ a : Acceptor, voted_for (s a) b2 v2)
a2: Acceptor
Ha2: a2 ∈ `Q_v2
Q12: {x : ASet | Quorum x}
H_b1_v2: allQuorum Quorum Q12 (λ a : Acceptor, voted_none_but (s a) b1 v2)
a: Acceptor
Ha_Q1: a ∈ `Q_v1
Ha_Q12: a ∈ `Q12

v1 = v2
by apply (H_b1_v2 _ Ha_Q12) in H_v1. Qed.
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → v : Value, v ∈ chosen s → chosen s ≡ {[v]}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False

s : state (voting_vlsm Quorum), valid_state_prop (voting_vlsm Quorum) s → v : Value, v ∈ chosen s → chosen s ≡ {[v]}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
Hv: v ∈ chosen s

chosen s ≡ {[v]}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
Hv: v ∈ chosen s
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2

chosen s ≡ {[v]}
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
Hv: v ∈ chosen s
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
x: Value

x ∈ chosen s ↔ x = v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
Hv: v ∈ chosen s
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
x: Value

x ∈ chosen s → x = v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
Hv: v ∈ chosen s
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
x: Value
Hx: x ∈ chosen s

x = v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
bv: Ballot
Hv: ballot_chose s bv v
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
x: Value
Hx: x ∈ chosen s

x = v
Value, VSet: Type
H: ElemOf Value VSet
H0: Empty VSet
H1: Singleton Value VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements Value VSet
EqDecision0: EqDecision Value
H6: FinSet Value VSet
VSDec: RelDecision elem_of
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
IM:= λ _ : Acceptor, acceptor_vlsm: Acceptor → VLSM False
s: state (voting_vlsm Quorum)
Hs: valid_state_prop (voting_vlsm Quorum) s
v: Value
bv: Ballot
Hv: ballot_chose s bv v
Hsafe: (b1 : Ballot) (v1 : Value) (b2 : Ballot) (v2 : Value), ballot_chose s b1 v1 → ballot_chose s b2 v2 → v1 = v2
x: Value
bx: Ballot
Hx: ballot_chose s bx x

x = v
by eapply Hsafe. Qed. End sec_voting_spec. Section sec_voting_simulates_consensus. Context (V : Type) (VSet : Type) `{FinSet V VSet} {VSDec : RelDecision (∈@{VSet})} (CV : VLSM False := consensus_vlsm V VSet) (Acceptor : Type) `{finite.Finite Acceptor} (ASet : Type) `{FinSet Acceptor ASet} (Quorum : ASet -> Prop) (QDec : forall Q, Decision (Quorum Q)) (QA : forall (Q1 Q2 : sig Quorum), exists a, a ∈ `Q1 ∩ `Q2) (QClosed : Proper (subseteq ==> impl) Quorum) (VV := voting_vlsm V VSet Acceptor ASet Quorum) .
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
Q: ASet

{Quorum Q} + { Q' : ASet, Q' ⊆ Q → ¬ Quorum Q'}
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
Q: ASet

{Quorum Q} + { Q' : ASet, Q' ⊆ Q → ¬ Quorum Q'}
by destruct (QDec Q); [left | right; intros Q' ->]. Defined. Definition state_project (s : state VV) : state CV := chosen V VSet Acceptor ASet Quorum QDec QClosed s. Definition label_state_project (s : state VV) (l : label VV) : label CV := let cs := state_project s in if decide (elements cs = []) then let s' := fst (transition VV l (s, None)) in match elements (state_project s') with | v :: _ => Some (Decided _ v) | _ => None end else None. Context (H_VSet_L : LeibnizEquiv VSet) .
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

(l : label (voting_vlsm V VSet Acceptor ASet Quorum)) (s : state (voting_vlsm V VSet Acceptor ASet Quorum)) (om : option False) (s' : state (voting_vlsm V VSet Acceptor ASet Quorum)) (om' : option False), input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, om) (s', om') → valid_state_prop (consensus_vlsm V VSet) (state_project s) → input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, om) (state_project s', om')
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

(l : label (voting_vlsm V VSet Acceptor ASet Quorum)) (s : state (voting_vlsm V VSet Acceptor ASet Quorum)) (om : option False) (s' : state (voting_vlsm V VSet Acceptor ASet Quorum)) (om' : option False), input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, om) (s', om') → valid_state_prop (consensus_vlsm V VSet) (state_project s) → input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, om) (state_project s', om')
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)

input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, None) (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)

(let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)) → input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, None) (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = ( state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)

(let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)) → input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, None) (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
H16: option_valid_message_prop (consensus_vlsm V VSet) None

(let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)) → input_valid_transition (consensus_vlsm V VSet) (label_state_project s l) (state_project s, None) (state_project s', None)
by intros NewGoal; repeat split; [| | apply NewGoal..].
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v

let lc := label_state_project s l in match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV

lc = label_state_project s l → match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV

lc = (if decide (elements (state_project s) = []) then match elements (state_project (transition l (s, None)).1) with | [] => None | v :: _ => Some (Decided V v) end else None) → match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV

lc = (if decide (elements (state_project s) = []) then match elements (state_project (s', None).1) with | [] => None | v :: _ => Some (Decided V v) end else None) → match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True end ∧ consensus_transition V VSet lc (state_project s, None) = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (state_project s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV

lc = (if decide (elements (state_project s) = []) then match elements (state_project s') with | [] => None | v :: _ => Some (Decided V v) end else None) → match lc with | Some (Decided _ _) => state_project s ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ state_project s, None) | None => (state_project s, None) end = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV

lc = (if decide (elements cs = []) then match elements (state_project s') with | [] => None | v :: _ => Some (Decided V v) end else None) → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (state_project s', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'

lc = (if decide (elements cs = []) then match elements cs' with | [] => None | v :: _ => Some (Decided V v) end else None) → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
Hcs': elements cs' = []

lc = None → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
v: V
l0: list V
Hcs': elements cs' = v :: l0
lc = Some (Decided V v) → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs ≠ []
lc = None → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
Hcs': elements cs' = []

lc = None → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
Hcs': elements cs' = []

True ∧ (cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
Hcs': elements cs' = []

(cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
Hcs': elements cs' = []

cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
Hcs': cs' ≡ ∅

cs ≡ cs'
by etransitivity.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
v: V
l0: list V
Hcs': elements cs' = v :: l0

lc = Some (Decided V v) → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs = []
v: V
l0: list V
Hcs': elements cs' = v :: l0

cs ≡ ∅ ∧ ({[v]} ∪ cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

cs ≡ ∅ ∧ ({[v]} ∪ cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

({[v]} ∪ cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

{[v]} ∪ cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

{[v]} ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

v ∈ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0
Hv_cs': v ∈ cs'
{[v]} ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

v ∈ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

v ∈ elements cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0

v ∈ v :: l0
by apply elem_of_list_here.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0
Hv_cs': v ∈ cs'

{[v]} ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0
Hv_cs': v ∈ cs'

cs' ≡ {[v]}
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = chosen V VSet Acceptor ASet Quorum QDec QClosed s'
Hcs: cs ≡ ∅
v: V
l0: list V
Hcs': elements cs' = v :: l0
Hv_cs': v ∈ cs'

cs' ≡ {[v]}
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
Hcs: cs ≡ ∅
v: V
l0: list V
Hv_cs': v ∈ chosen V VSet Acceptor ASet Quorum QDec QClosed s'
Hcs': elements (chosen V VSet Acceptor ASet Quorum QDec QClosed s') = v :: l0

chosen V VSet Acceptor ASet Quorum QDec QClosed s' ≡ {[v]}
by apply chosen_subsingleton.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
lc: label CV
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs ≠ []

lc = None → match lc with | Some (Decided _ _) => cs ≡ ∅ | None => True endmatch lc with | Some (Decided _ v) => ({[v]} ∪ cs, None) | None => (cs, None) end = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs ≠ []

True ∧ (cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs ≠ []

(cs, None) = (cs', None)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: elements cs ≠ []

cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
Hcs: cs ≢ ∅

cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs

cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs

v ∈ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'
cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs

v ∈ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (chosen V VSet Acceptor ASet Quorum QDec QClosed s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
v: V
Hv: v ∈ chosen V VSet Acceptor ASet Quorum QDec QClosed s

v ∈ chosen V VSet Acceptor ASet Quorum QDec QClosed s'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (chosen V VSet Acceptor ASet Quorum QDec QClosed s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
v: V
bv: Ballot
Hv: ballot_chose V VSet Acceptor ASet Quorum s bv v

v ∈ chosen V VSet Acceptor ASet Quorum QDec QClosed s'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (chosen V VSet Acceptor ASet Quorum QDec QClosed s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
v: V
bv: Ballot
Hv: ballot_chose V VSet Acceptor ASet Quorum s' bv v

v ∈ chosen V VSet Acceptor ASet Quorum QDec QClosed s'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
H_proj_s: valid_state_prop (consensus_vlsm V VSet) (chosen V VSet Acceptor ASet Quorum QDec QClosed s)
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
v: V
bv: Ballot
Hv: ballot_chose V VSet Acceptor ASet Quorum s' bv v

b : Ballot, ballot_chose V VSet Acceptor ASet Quorum s' b v
by exists bv.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'

cs ≡ cs'
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'

cs ≡ {[v]}
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'
cs' ≡ {[v]}
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'

cs ≡ {[v]}
by subst cs; apply chosen_subsingleton.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
s, s': state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, None) (s', None)
cs: consensus_state VSet
H_eqn_cs: cs = state_project s
H_proj_s: valid_state_prop (consensus_vlsm V VSet) cs
Hs: valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s
Hs': valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s'
Hs_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s b2 v2 → v1 = v2
Hs'_safe: (b1 : Ballot) (v1 : V) (b2 : Ballot) (v2 : V), ballot_chose V VSet Acceptor ASet Quorum s' b1 v1 → ballot_chose V VSet Acceptor ASet Quorum s' b2 v2 → v1 = v2
Hchoice_stable: (b : Ballot) (v : V), ballot_chose V VSet Acceptor ASet Quorum s b v → ballot_chose V VSet Acceptor ASet Quorum s' b v
cs': consensus_state VSet
H_eqn_cs': cs' = state_project s'
v: V
Hv: v ∈ cs
Hv': v ∈ cs'

cs' ≡ {[v]}
by subst cs'; apply chosen_subsingleton. Qed.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

s : state (voting_vlsm V VSet Acceptor ASet Quorum), initial_state_prop s → initial_state_prop (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

s : state (voting_vlsm V VSet Acceptor ASet Quorum), initial_state_prop s → initial_state_prop (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s

initial_state_prop (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s

state_project s ≡ ∅
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
Hv: v ∈ state_project s

False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Hv: ballot_chose V VSet Acceptor ASet Quorum s bv v

False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Q: {x : ASet | Quorum x}
HvQ: allQuorum Acceptor ASet Quorum Q (λ a : Acceptor, voted_for V VSet (s a) bv v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Q: {x : ASet | Quorum x}
HvQ: allQuorum Acceptor ASet Quorum Q (λ a : Acceptor, voted_for V VSet (s a) bv v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

voted_for V VSet (s a) bv v → False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Q: {x : ASet | Quorum x}
HvQ: allQuorum Acceptor ASet Quorum Q (λ a : Acceptor, voted_for V VSet (s a) bv v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

v ∈ votes VSet (s a) !!! bv → False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Q: {x : ASet | Quorum x}
HvQ: allQuorum Acceptor ASet Quorum Q (λ a : Acceptor, voted_for V VSet (s a) bv v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

v ∈ ∅ !!! bv → False
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s
v: V
bv: Ballot
Q: {x : ASet | Quorum x}
HvQ: allQuorum Acceptor ASet Quorum Q (λ a : Acceptor, voted_for V VSet (s a) bv v)
a: Acceptor
Ha: a ∈ `Q ∩ `Q

v ∈ inhabitant → False
by apply elem_of_empty. Qed.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

s : state (voting_vlsm V VSet Acceptor ASet Quorum), valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s → valid_state_prop (consensus_vlsm V VSet) (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet

s : state (voting_vlsm V VSet Acceptor ASet Quorum), valid_state_prop (voting_vlsm V VSet Acceptor ASet Quorum) s → valid_state_prop (consensus_vlsm V VSet) (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s

valid_state_prop (consensus_vlsm V VSet) (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s': state (voting_vlsm V VSet Acceptor ASet Quorum)
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
om, om': option False
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, om) (s', om')
IHHs: valid_state_prop (consensus_vlsm V VSet) (state_project s)
valid_state_prop (consensus_vlsm V VSet) (state_project s')
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s

valid_state_prop (consensus_vlsm V VSet) (state_project s)
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Hs: initial_state_prop s

initial_state_prop (state_project s)
by apply voting_initial_in_consensus.
V, VSet: Type
H: ElemOf V VSet
H0: Empty VSet
H1: Singleton V VSet
H2: Union VSet
H3: Intersection VSet
H4: Difference VSet
H5: Elements V VSet
EqDecision0: EqDecision V
H6: FinSet V VSet
VSDec: RelDecision elem_of
CV:= consensus_vlsm V VSet: VLSM False
Acceptor: Type
EqDecision1: EqDecision Acceptor
H7: finite.Finite Acceptor
ASet: Type
H8: ElemOf Acceptor ASet
H9: Empty ASet
H10: Singleton Acceptor ASet
H11: Union ASet
H12: Intersection ASet
H13: Difference ASet
H14: Elements Acceptor ASet
EqDecision2: EqDecision Acceptor
H15: FinSet Acceptor ASet
Quorum: ASet → Prop
QDec: Q : ASet, Decision (Quorum Q)
QA: Q1 Q2 : {x : ASet | Quorum x}, a : Acceptor, a ∈ `Q1 ∩ `Q2
QClosed: Proper (subseteq ==> impl) Quorum
VV:= voting_vlsm V VSet Acceptor ASet Quorum: VLSM False
H_VSet_L: LeibnizEquiv VSet
s': state (voting_vlsm V VSet Acceptor ASet Quorum)
l: label (voting_vlsm V VSet Acceptor ASet Quorum)
om, om': option False
s: state (voting_vlsm V VSet Acceptor ASet Quorum)
Ht: input_valid_transition (voting_vlsm V VSet Acceptor ASet Quorum) l (s, om) (s', om')
IHHs: valid_state_prop (consensus_vlsm V VSet) (state_project s)

valid_state_prop (consensus_vlsm V VSet) (state_project s')
by eapply input_valid_transition_destination, voting_step_in_consensus. Qed. End sec_voting_simulates_consensus.