From stdpp Require Import prelude finite.From VLSM.Lib Require Import Preamble ListSetExtras. From VLSM.Core Require Import VLSM PreloadedVLSM Composition VLSMEmbedding.
Core: Traceable VLSMs
A class with a measure of size for states which is monotonic for valid transitions.
Class TransitionMonotoneVLSM `(X : VLSM message) (state_size : state X -> nat) : Prop := { transition_monotonicity : forall s1 s2 : state X, valid_transition_next X s1 s2 -> state_size s1 < state_size s2 }. #[global] Hint Mode TransitionMonotoneVLSM - ! - : typeclass_instances.message: Type
X: VLSM message
state_size: state X → nat
TransitionMonotoneVLSM0: TransitionMonotoneVLSM X state_sizeTransitionMonotoneVLSM (preloaded_with_all_messages_vlsm X) state_sizemessage: Type
X: VLSM message
state_size: state X → nat
TransitionMonotoneVLSM0: TransitionMonotoneVLSM X state_sizeTransitionMonotoneVLSM (preloaded_with_all_messages_vlsm X) state_sizeby apply transition_monotonicity, valid_transition_next_preloaded_iff. Qed.message: Type
X: VLSM message
state_size: state X → nat
TransitionMonotoneVLSM0: TransitionMonotoneVLSM X state_size
s1, s2: state (preloaded_with_all_messages_vlsm X)
Ht: valid_transition_next (preloaded_with_all_messages_vlsm X) s1 s2state_size s1 < state_size s2message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s, sf: state X
Hfutures: in_futures X s sfstate_size s ≤ state_size sfmessage: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s, sf: state X
Hfutures: in_futures X s sfstate_size s ≤ state_size sfmessage: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s, sf: state X
tr: list transition_item
Htr: finite_valid_trace_from_to X s sf trstate_size s ≤ state_size sfby apply input_valid_transition_forget_input, transition_next, transition_monotonicity in Ht; lia. Qed.message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s, f: state X
tl: list transition_item
Htr: finite_valid_trace_from_to X s f tl
s': state X
iom, oom: option message
l: label X
Ht: input_valid_transition X l (s', iom) (s, oom)
IHHtr: state_size s ≤ state_size fstate_size s' ≤ state_size fmessage: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size∀ (s : state X) (tr : list transition_item), finite_valid_trace_from_to X s s tr → tr = []message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size∀ (s : state X) (tr : list transition_item), finite_valid_trace_from_to X s s tr → tr = []message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s: state X
tr: list transition_item
f: state X
Heqf: f = s
Htr: finite_valid_trace_from_to X s f trtr = []message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s', s: state X
tl: list transition_item
Htr: finite_valid_trace_from_to X s s' tl
iom, oom: option message
l: label X
Ht: input_valid_transition X l (s', iom) (s, oom)
IHHtr: s' = s → tl = []{| l := l; input := iom; destination := s; output := oom |} :: tl = []by apply input_valid_transition_forget_input, transition_next, transition_monotonicity in Ht; lia. Qed.message: Type
X: VLSM message
state_size: state X → nat
H: TransitionMonotoneVLSM X state_size
s', s: state X
tl: list transition_item
Htr: finite_valid_trace_from_to X s s' tl
iom, oom: option message
l: label X
Ht: input_valid_transition X l (s', iom) (s, oom)
IHHtr: s' = s → tl = []
H0: state_size s ≤ state_size s'{| l := l; input := iom; destination := s; output := oom |} :: tl = []
A class characterizing VLSMs with reversible transitions. A VLSM is traceable
when given a constrained state, one can compute a set of constrained transitions
leading to it.
We assume that such machines are TransitionMonotoneVLSMs and that the
set of constrained transitions associated to a constrained state is empty
iff the state is initial.
#[global] Hint Mode TraceableVLSM - ! - - : typeclass_instances.
Section sec_traceable_vlsm_props. Context `(X : VLSM message) (state_size : state X -> nat) (state_destructor : state X -> list (transition_item X * state X)) `{!TraceableVLSM X state_destructor state_size} .message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ s' : state X, constrained_state_prop X s' → ∀ (s : state X) (item : transition_item), (item, s) ∈ state_destructor s' → state_size s < state_size s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ s' : state X, constrained_state_prop X s' → ∀ (s : state X) (item : transition_item), (item, s) ∈ state_destructor s' → state_size s < state_size s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
H: constrained_state_prop X s'
s: state X
item: transition_item
H0: (item, s) ∈ state_destructor s'state_size s < state_size s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
H: constrained_state_prop X s'
s: state X
item: transition_item
H0: (item, s) ∈ state_destructor s'valid_transition_next X s s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
H: constrained_state_prop X s'
s: state X
item: transition_item
H0: (item, s) ∈ state_destructor s'valid_transition_next X s (destination item)by eapply valid_transition_preloaded_iff, input_valid_transition_forget_input, tv_state_destructor_transition. Qed.message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
H: constrained_state_prop X s'
s: state X
item: transition_item
H0: (item, s) ∈ state_destructor s'valid_transition X ?l s ?iom (destination item) ?oom
Given any constrained state we can extract a trace leading to it by recursively
following the transitions leading to it.
Equations state_to_trace (s' : state X) (Hs' : constrained_state_prop X s') : state X * list (transition_item X) by wf (state_size s') lt := state_to_trace s' Hs' with inspect (state_destructor s') := | [] eq: _ => (s', []) | ((item, s) :: _) eq: Hdestruct => let (is, tr) := state_to_trace s _ in (is, tr ++ [item]).message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)), (λ y : list (transition_item * state X), state_destructor s' = y) ((item, s) :: l0) → constrained_state_prop X s' → (∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item) → constrained_state_prop X smessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)), (λ y : list (transition_item * state X), state_destructor s' = y) ((item, s) :: l0) → constrained_state_prop X s' → (∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item) → constrained_state_prop X smessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
state_to_trace: ∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_itemconstrained_state_prop X sby rewrite Hdestruct; left. Qed.message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
state_to_trace: ∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item(?item, s) ∈ state_destructor s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)), (λ y : list (transition_item * state X), state_destructor s' = y) ((item, s) :: l0) → constrained_state_prop X s' → (∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item) → state_size s < state_size s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)), (λ y : list (transition_item * state X), state_destructor s' = y) ((item, s) :: l0) → constrained_state_prop X s' → (∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item) → state_size s < state_size s'message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
state_to_trace: ∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_itemstate_size s < state_size s'by rewrite Hdestruct; left. Qed.message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
state_to_trace: ∀ x : state X, constrained_state_prop X x → state_size x < state_size s' → state X * list transition_item(?item, s) ∈ state_destructor s'
Traces extracted using state_to_trace are constrained traces.
message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s : state X) (Hs : constrained_state_prop X s) (is : state X) (tr : list transition_item), state_to_trace s Hs = (is, tr) → finite_constrained_trace_init_to X is s trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s : state X) (Hs : constrained_state_prop X s) (is : state X) (tr : list transition_item), state_to_trace s Hs = (is, tr) → finite_constrained_trace_init_to X is s trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s: state X
Hs: constrained_state_prop X s∀ (is : state X) (tr : list transition_item), state_to_trace s Hs = (is, tr) → finite_constrained_trace_init_to X is s trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (e : state_destructor s' = []), constrained_state_prop X s' → inspect (state_destructor s') = [] ↾ e → ∀ (is : state X) (tr : list transition_item), (s', []) = (is, tr) → finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)) (Hdestruct : state_destructor s' = (item, s) :: l0) (Hs' : constrained_state_prop X s'), (∀ (is : state X) (tr : list transition_item), state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') = ( is, tr) → finite_constrained_trace_init_to X is s tr) → inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestruct → ∀ (is : state X) (tr : list transition_item), (let (is0, tr0) := state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') in (is0, tr0 ++ [item])) = ( is, tr) → finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (e : state_destructor s' = []), constrained_state_prop X s' → inspect (state_destructor s') = [] ↾ e → ∀ (is : state X) (tr : list transition_item), (s', []) = (is, tr) → finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
Hdestruct: state_destructor s' = []
Hs': constrained_state_prop X s'
Heq: inspect (state_destructor s') = [] ↾ Hdestruct
is: state X
tr: list transition_item
Heqis_tr: (s', []) = (is, tr)finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
is: state X
Hdestruct: state_destructor is = []
Heq: inspect (state_destructor is) = [] ↾ Hdestruct
Hs': constrained_state_prop X is
Heqis_tr: (is, []) = (is, [])finite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) is is []message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
is: state X
Hdestruct: state_destructor is = []
Heq: inspect (state_destructor is) = [] ↾ Hdestruct
Hs': constrained_state_prop X is
Heqis_tr: (is, []) = (is, [])initial_state_prop isby rapply @finite_valid_trace_from_to_empty.message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
is: state X
Hdestruct: state_destructor is = []
Heq: inspect (state_destructor is) = [] ↾ Hdestruct
Hs': constrained_state_prop X is
Heqis_tr: (is, []) = (is, [])finite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) is is []by eapply @tv_state_destructor_initial with (X := X).message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
is: state X
Hdestruct: state_destructor is = []
Heq: inspect (state_destructor is) = [] ↾ Hdestruct
Hs': constrained_state_prop X is
Heqis_tr: (is, []) = (is, [])initial_state_prop ismessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size∀ (s' : state X) (item : transition_item) (s : state X) (l0 : list (transition_item * state X)) (Hdestruct : state_destructor s' = (item, s) :: l0) (Hs' : constrained_state_prop X s'), (∀ (is : state X) (tr : list transition_item), state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') = (is, tr) → finite_constrained_trace_init_to X is s tr) → inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestruct → ∀ (is : state X) (tr : list transition_item), (let (is0, tr0) := state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') in (is0, tr0 ++ [item])) = (is, tr) → finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
Hind: ∀ (is : state X) (tr : list transition_item), state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') = ( is, tr) → finite_constrained_trace_init_to X is s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestruct
is: state X
tr: list transition_item
Heqis_tr: (let (is, tr) := state_to_trace s (state_to_trace_obligations_obligation_1 s' item s l0 Hdestruct Hs') in (is, tr ++ [item])) = ( is, tr)finite_constrained_trace_init_to X is s' trmessage: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructfinite_constrained_trace_init_to X is s' (_tr ++ [item])message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructfinite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) is s' (_tr ++ [item])message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructfinite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) is (destination item) (_tr ++ [item])message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructfinite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) s (destination item) [item]message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructinput_constrained_transition_item X s item → finite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) s (destination item) [item]message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructinput_constrained_transition_item X s itemby destruct item; cbn; rapply (@finite_valid_trace_from_to_singleton).message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructinput_constrained_transition_item X s item → finite_valid_trace_from_to (preloaded_with_all_messages_vlsm X) s (destination item) [item]message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestructinput_constrained_transition_item X s itemby rewrite Hdestruct; left. Qed. End sec_traceable_vlsm_props.message: Type
X: VLSM message
state_size: state X → nat
state_destructor: state X → list (transition_item * state X)
TraceableVLSM0: TraceableVLSM X state_destructor state_size
s': state X
item: transition_item
s: state X
l0: list (transition_item * state X)
Hdestruct: state_destructor s' = (item, s) :: l0
Hs': constrained_state_prop X s'
_tr: list transition_item
is: state X
Hind: ∀ (is0 : state X) (tr : list transition_item), (is, _tr) = (is0, tr) → finite_constrained_trace_init_to X is0 s tr
Heq: inspect (state_destructor s') = ((item, s) :: l0) ↾ Hdestruct(item, s) ∈ state_destructor s'
Section sec_traceable_vlsm_composition. Context {message : Type} `{finite.Finite index} (IM : index -> VLSM message) (state_destructor : forall i, state (IM i) -> list (transition_item (IM i) * state (IM i))) (state_size : forall i, state (IM i) -> nat) `{forall i, TraceableVLSM (IM i) (state_destructor i) (state_size i)} (Free := free_composite_vlsm IM) .
The following definitions and lemmas lift the corresponding notions and
results from TraceableVLSMs to their composition.
Definition lift_to_composite_transition_item_state (s : composite_state IM) (i : index) (item_s : transition_item (IM i) * state (IM i)) : composite_transition_item IM * composite_state IM := (lift_to_composite_transition_item IM s i item_s.1, lift_to_composite_state IM s i item_s.2). Definition composite_state_destructor (s' : composite_state IM) (i : index) : list (composite_transition_item IM * composite_state IM) := map (lift_to_composite_transition_item_state s' i) (state_destructor i (s' i)). Definition composite_state_size (s : composite_state IM) := foldr Nat.add 0 (map (fun i => state_size i (s i)) (enum index)).message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s : composite_state IM) (i : index) (si : state (IM i)), state_size i (s i) < state_size i si → composite_state_size s < composite_state_size (lift_to_composite_state IM s i si)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s : composite_state IM) (i : index) (si : state (IM i)), state_size i (s i) < state_size i si → composite_state_size s < composite_state_size (lift_to_composite_state IM s i si)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i sicomposite_state_size s < composite_state_size (lift_to_composite_state IM s i si)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i sifoldr Nat.add 0 (map (λ i : index, state_size i (s i)) (enum index)) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) (enum index))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i si∀ l : list index, i ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
a: index
si: state (IM a)
Hlt: state_size a (s a) < state_size a si
l: list index
H1: a ∈ a :: l
IHl: a ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)NoDup (a :: l) → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a :: l)) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) (a :: l))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i si
a: index
l: list index
IHl: i ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)
H1: i ∈ a :: l
H2: i ∈ lNoDup (a :: l) → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a :: l)) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) (a :: l))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
a: index
si: state (IM a)
Hlt: state_size a (s a) < state_size a si
l: list index
H1: a ∈ a :: l
IHl: a ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)NoDup (a :: l) → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a :: l)) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) (a :: l))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
a: index
si: state (IM a)
Hlt: state_size a (s a) < state_size a si
l: list index
H1: a ∈ a :: l
IHl: a ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)
H2: NoDup (a :: l)
Hna: a ∉ lstate_size a (s a) + foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < state_size a (lift_to_composite_state IM s a si a) + foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
a: index
si: state (IM a)
Hlt: state_size a (s a) < state_size a si
l: list index
H1: a ∈ a :: l
IHl: a ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)
H2: NoDup (a :: l)
Hna: a ∉ lstate_size a (s a) + foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < state_size a si + foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
a: index
si: state (IM a)
Hlt: state_size a (s a) < state_size a si
l: list index
H1: a ∈ a :: l
IHl: a ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l)
H2: NoDup (a :: l)
Hna: a ∉ lfoldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l) = foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l)message, index: Type
EqDecision0: EqDecision index
IM: index → VLSM message
state_size: ∀ i : index, state (IM i) → nat
s: composite_state IM
a: index
si: state (IM a)
a0: index
l: list index
IHl: a ∉ l → foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l) = foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l)a ∉ a0 :: l → foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) (a0 :: l)) = foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a0 :: l))by cbn; rewrite state_update_neq, IHl.message, index: Type
EqDecision0: EqDecision index
IM: index → VLSM message
state_size: ∀ i : index, state (IM i) → nat
s: composite_state IM
a: index
si: state (IM a)
a0: index
l: list index
IHl: a ∉ l → foldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) l) = foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l)
H: a ≠ a0
H0: a ∉ lfoldr Nat.add 0 (map (λ i : index, state_size i (lift_to_composite_state IM s a si i)) (a0 :: l)) = foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a0 :: l))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i si
a: index
l: list index
IHl: i ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)
H1: i ∈ a :: l
H2: i ∈ lNoDup (a :: l) → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a :: l)) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) (a :: l))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i si
a: index
l: list index
IHl: i ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)
H1: i ∈ a :: l
H2: i ∈ l
H3: NoDup (a :: l)
H6: a ∉ l
H7: NoDup lfoldr Nat.add 0 (map (λ i : index, state_size i (s i)) (a :: l)) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) (a :: l))by apply Nat.add_lt_mono_l, IHl. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
i: index
si: state (IM i)
Hlt: state_size i (s i) < state_size i si
a: index
l: list index
IHl: i ∈ l → NoDup l → foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)
H1: i ∈ a :: l
H2: i ∈ l
H3: NoDup (a :: l)
H6: a ∉ l
H7: NoDup lstate_size a (s a) + foldr Nat.add 0 (map (λ i : index, state_size i (s i)) l) < state_size a (s a) + foldr Nat.add 0 (map (λ i0 : index, state_size i0 (lift_to_composite_state IM s i si i0)) l)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s' s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → destination item = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s' s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → destination item = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s', s: composite_state IM
item: composite_transition_item IM
i: index(∃ y : transition_item * state (IM i), (item, s) = lift_to_composite_transition_item_state s' i y ∧ y ∈ state_destructor i (s' i)) → destination item = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)destination (lift_to_composite_transition_item IM s' i itemi) = s'by erewrite tv_state_destructor_destination, state_update_id. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)state_update IM s' i (destination itemi) = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
s: composite_state IM
item: composite_transition_item IM
i: index(item, s) ∈ composite_state_destructor s' i → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
s: composite_state IM
item: composite_transition_item IM
i: index(∃ y : transition_item * state (IM i), (item, s) = lift_to_composite_transition_item_state s' i y ∧ y ∈ state_destructor i (s' i)) → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)input_constrained_transition_item Free (lift_to_composite_state IM s' i si) (lift_to_composite_transition_item IM s' i itemi)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)input_valid_transition (preloaded_with_all_messages_vlsm (IM i)) (l itemi) (si, input (lift_to_composite_transition_item IM s' i itemi)) (destination itemi, output (lift_to_composite_transition_item IM s' i itemi))by eapply composite_constrained_state_project. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)constrained_state_prop (IM i) (s' i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s' s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → projT1 (l item) = imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (s' s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → projT1 (l item) = imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s', s: composite_state IM
item: composite_transition_item IM
i: index(item, s) ∈ composite_state_destructor s' i → projT1 (l item) = iby intros ((itemi, si) & [= ->] & Hin). Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s', s: composite_state IM
item: composite_transition_item IM
i: index(∃ y : transition_item * state (IM i), (item, s) = lift_to_composite_transition_item_state s' i y ∧ y ∈ state_destructor i (s' i)) → projT1 (l item) = imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → s' = state_update IM s i (destination item i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → s' = state_update IM s i (destination item i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
s: composite_state IM
item: composite_transition_item IM
i: index
Hin: (item, s) ∈ composite_state_destructor s' is' = state_update IM s i (destination item i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
item: composite_transition_item IM
Hs': constrained_state_prop Free (destination item)
s: composite_state IM
i: index
Hin: (item, s) ∈ composite_state_destructor (destination item) idestination item = state_update IM s i (destination item i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
item: composite_transition_item IM
Hs': constrained_state_prop Free (destination item)
s: composite_state IM
i: index
Hin: (item, s) ∈ composite_state_destructor (destination item) i
Hl: projT1 (l item) = idestination item = state_update IM s i (destination item i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
item: composite_transition_item IM
Hs': constrained_state_prop Free (destination item)
s: composite_state IM
i: index
Hl: projT1 (l item) = i
Ht: transition (l item) (s, input item) = (destination item, output item)destination item = state_update IM s i (destination item i)by state_update_simpl. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
i: index
l: label (IM i)
input, output: option message
s: composite_state IM
s0: state (IM i)
Hs': constrained_state_prop Free (state_update IM s i s0)
Ht: (state_update IM s i s0, output) = (state_update IM s i s0, output)state_update IM s i s0 = state_update IM s i (state_update IM s i s0 i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s : composite_state IM, constrained_state_prop Free s → ∀ i : index, initial_state_prop (s i) ↔ composite_state_destructor s i = []message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s : composite_state IM, constrained_state_prop Free s → ∀ i : index, initial_state_prop (s i) ↔ composite_state_destructor s i = []message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: initial_state_prop (s i)map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []initial_state_prop (s i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: initial_state_prop (s i)map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: initial_state_prop (s i)[] = state_destructor i (s i)by eapply composite_constrained_state_project.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: initial_state_prop (s i)constrained_state_prop (IM i) (s i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []initial_state_prop (s i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []constrained_state_prop (IM i) (s i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []state_destructor i (s i) = []by eapply composite_constrained_state_project.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []constrained_state_prop (IM i) (s i)by eapply fmap_nil_inv. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
H1: constrained_state_prop Free s
i: index
Hinit: map (lift_to_composite_transition_item_state s i) (state_destructor i (s i)) = []state_destructor i (s i) = []message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (s : composite_state IM) (item : composite_transition_item IM), (item, s) ∈ composite_state_destructor s' i → ∀ j : index, initial_state_prop (s' j) → s j = s' jmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (s : composite_state IM) (item : composite_transition_item IM), (item, s) ∈ composite_state_destructor s' i → ∀ j : index, initial_state_prop (s' j) → s j = s' jmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
s: composite_state IM
item: composite_transition_item IM
Hdestruct: (item, s) ∈ composite_state_destructor s' i∀ j : index, initial_state_prop (s' j) → s j = s' jmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
s: composite_state IM
item: composite_transition_item IM
Hdestruct: (item, s) ∈ composite_state_destructor s' i
Heqs': s' = state_update IM s i (destination item i)∀ j : index, initial_state_prop (s' j) → s j = s' jmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
item: composite_transition_item IM
j: index
Hinit: initial_state_prop (state_update IM s j (destination item j) j)
Hs': constrained_state_prop Free (state_update IM s j (destination item j))
Hdestruct: (item, s) ∈ composite_state_destructor (state_update IM s j (destination item j)) js j = state_update IM s j (destination item j) jby rewrite Hinit in Hdestruct; inversion Hdestruct. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s: composite_state IM
item: composite_transition_item IM
j: index
Hinit: composite_state_destructor (state_update IM s j (destination item j)) j = []
Hs': constrained_state_prop Free (state_update IM s j (destination item j))
Hdestruct: (item, s) ∈ composite_state_destructor (state_update IM s j (destination item j)) js j = state_update IM s j (destination item j) jmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → composite_state_size s < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (s : composite_state IM) (item : composite_transition_item IM) (i : index), (item, s) ∈ composite_state_destructor s' i → composite_state_size s < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
s: composite_state IM
item: composite_transition_item IM
i: index(item, s) ∈ composite_state_destructor s' i → composite_state_size s < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
s: composite_state IM
item: composite_transition_item IM
i: index(∃ y : transition_item * state (IM i), (item, s) = lift_to_composite_transition_item_state s' i y ∧ y ∈ state_destructor i (s' i)) → composite_state_size s < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)composite_state_size (lift_to_composite_state IM s' i si) < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)composite_state_size (lift_to_composite_state IM s' i si) < composite_state_size (state_update IM (state_update IM s' i si) i (s' i))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)state_update IM (state_update IM s' i si) i (s' i) = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)composite_state_size (lift_to_composite_state IM s' i si) < composite_state_size (state_update IM (state_update IM s' i si) i (s' i))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)state_size i (state_update IM s' i si i) < state_size i (s' i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)state_size i si < state_size i (s' i)by eapply composite_constrained_state_project.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)constrained_state_prop (IM i) (s' i)by rewrite state_update_twice, state_update_id. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
itemi: transition_item
si: state (IM i)
Hin: (itemi, si) ∈ state_destructor i (s' i)state_update IM (state_update IM s' i si) i (s' i) = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (n : nat) (item : transition_item) (s : state (preloaded_with_all_messages_vlsm Free)), composite_state_destructor s' i !! n = Some (item, s) → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (n : nat) (item : transition_item) (s : state (preloaded_with_all_messages_vlsm Free)), composite_state_destructor s' i !! n = Some (item, s) → input_constrained_transition_item Free s itemby eapply composite_tv_state_destructor_transition, elem_of_list_lookup_2. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
n: nat
item: transition_item
s: state (preloaded_with_all_messages_vlsm Free)
Hdestruct: composite_state_destructor s' i !! n = Some (item, s)input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (item : composite_transition_item IM) (s : composite_state IM), head (composite_state_destructor s' i) = Some (item, s) → input_constrained_transition_item Free s itemmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (item : composite_transition_item IM) (s : composite_state IM), head (composite_state_destructor s' i) = Some (item, s) → input_constrained_transition_item Free s itemby eapply composite_tv_state_destructor_transition, head_Some_elem_of. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
item: composite_transition_item IM
s: composite_state IM
Hdestruct: head (composite_state_destructor s' i) = Some (item, s)input_constrained_transition_item Free s item
Given a composite_state for a composition of TraceableVLSMs,
composite_state_destructor will return a set of possible transitions leading
to that state. Therefore, when trying to reconstruct a particular trace leading
to the given state, we must choose one among the possible transitions.
Let us define the type of such a choice function, which takes a composite
constrained state and a list of indices as arguments and returns one of the
indices and a position in the list of possible transitions to the state for
that particular index.
Definition choice_function : Type :=
forall s' : composite_state IM, constrained_state_prop Free s' -> list index -> index * nat.
Given a choice_function and a particular instance of its arguments,
we say that the choice function is choosing well, provided that:
- if the set of indices is non-empty, then the returned index must belong to it
- if the component state corresponding to the returned index in not initial, then
the returned position must identify a transition leading to the given state
- the choice does not depend on the particular proof for the composite constrained state
Record ChoosingWell (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index) : Prop := { cw_proof_independent : forall Hs'' : constrained_state_prop Free s', choose s' Hs' indices = choose s' Hs'' indices; cw_chosen_index_in_indices : indices <> [] -> (choose s' Hs' indices).1 ∈ indices; cw_chosen_position_exists : forall (i : index) (n : nat), choose s' Hs' indices = (i, n) -> ~ initial_state_prop (IM i) (s' i) -> is_Some (composite_state_destructor s' i !! n); }.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index), ChoosingWell choose s' Hs' indices → let i_n := choose s' Hs' indices in composite_state_destructor s' i_n.1 ≠ [] → composite_state_destructor s' i_n.1 !! i_n.2 ≠ Nonemessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index), ChoosingWell choose s' Hs' indices → let i_n := choose s' Hs' indices in composite_state_destructor s' i_n.1 ≠ [] → composite_state_destructor s' i_n.1 !! i_n.2 ≠ Nonemessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list index
Hchoose: ChoosingWell choose s' Hs' indices
i_n:= choose s' Hs' indices: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 ≠ []composite_state_destructor s' i_n.1 !! i_n.2 ≠ Nonemessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list index
Hchoose: ChoosingWell choose s' Hs' indices
i_n:= choose s' Hs' indices: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 ≠ []choose s' Hs' indices = (i_n.1, i_n.2)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list index
Hchoose: ChoosingWell choose s' Hs' indices
i_n:= choose s' Hs' indices: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 ≠ []¬ initial_state_prop (s' i_n.1)by destruct (choose _ _ _).message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list index
Hchoose: ChoosingWell choose s' Hs' indices
i_n:= choose s' Hs' indices: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 ≠ []choose s' Hs' indices = (i_n.1, i_n.2)by contradict Hdestruct; apply composite_tv_state_destructor_initial. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list index
Hchoose: ChoosingWell choose s' Hs' indices
i_n:= choose s' Hs' indices: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 ≠ []¬ initial_state_prop (s' i_n.1)
This is a property of the set of indices argument to a choice_function,
guaranteeing that the component states for the non-included indices are initial.
Definition not_in_indices_initial_prop
(s' : composite_state IM)
(indices : list index)
: Prop :=
forall i, i ∉ indices -> initial_state_prop (IM i) (s' i).
The constrained transitions leading to a composite constrained state reflect
the not_in_indices_initial_property, or, alternately, the
composite_state_destructor function reflects it.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (n : nat) (s : composite_state IM) (item : composite_transition_item IM), composite_state_destructor s' i !! n = Some (item, s) → ∀ indices : list index, not_in_indices_initial_prop s' indices → not_in_indices_initial_prop s indicesmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ (i : index) (n : nat) (s : composite_state IM) (item : composite_transition_item IM), composite_state_destructor s' i !! n = Some (item, s) → ∀ indices : list index, not_in_indices_initial_prop s' indices → not_in_indices_initial_prop s indicesby erewrite composite_tv_state_destructor_reflects_initiality; [apply Hinits' | | eapply elem_of_list_lookup_2 | apply Hinits']. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
n: nat
s: composite_state IM
item: composite_transition_item IM
Heq: composite_state_destructor s' i !! n = Some (item, s)
indices: list index
Hinits': not_in_indices_initial_prop s' indices
j: index
Hj: j ∉ indicesinitial_state_prop (s j)
Removing an index whose associated set of transitions is empty
(so that the corresponding component state is initial) preserves
not_in_indices_initial_prop.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ i : index, composite_state_destructor s' i = [] → ∀ indices : list index, not_in_indices_initial_prop s' indices → not_in_indices_initial_prop s' (StdppListSet.set_remove i indices)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ s' : composite_state IM, constrained_state_prop Free s' → ∀ i : index, composite_state_destructor s' i = [] → ∀ indices : list index, not_in_indices_initial_prop s' indices → not_in_indices_initial_prop s' (StdppListSet.set_remove i indices)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
Hdestruct: composite_state_destructor s' i = []
indices: list index
Hinit: not_in_indices_initial_prop s' indices
j: index
Hj: j ∉ StdppListSet.set_remove i indicesinitial_state_prop (s' j)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
Hdestruct: composite_state_destructor s' i = []
indices: list index
Hinit: not_in_indices_initial_prop s' indices
j: index
Hj: j ∉ StdppListSet.set_remove i indices
e: j ∈ indicesinitial_state_prop (s' j)by contradict Hj; subst; apply StdppListSet.set_remove_3. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
Hdestruct: composite_state_destructor s' i = []
indices: list index
Hinit: not_in_indices_initial_prop s' indices
j: index
Hj: j ∉ StdppListSet.set_remove i indices
e: j ∈ indices
n: j ≠ iinitial_state_prop (s' j)
A choice_function is choosing well if it is ChoosingWell for any
instance of its arguments (a constrained state and a set of indices)
satisfying the not_in_indices_initial_property.
Definition choosing_well (choose : choice_function) : Prop :=
forall (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index),
NoDup indices -> not_in_indices_initial_prop s' indices ->
ChoosingWell choose s' Hs' indices.
Given a composite constrained state, a choice_function and an initial
set of indices, we can extract a trace leading to that state by following
backwards the transitions yielded by the choice function until the states
corresponding to any of the given indices become initial states.
Equations indexed_composite_state_to_trace (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index) : composite_state IM * list (composite_transition_item IM) by wf (length indices + composite_state_size s') lt := | choose, s', Hs', [] => (s', []) | choose, s', Hs', indices with inspect (choose s' Hs' indices) => | i_n eq: Hchoice with inspect (composite_state_destructor s' i_n.1 !! i_n.2) => | (Some (item, s)) eq: Hdestruct with indexed_composite_state_to_trace choose s _ indices => { | (is, tr) => (is, tr ++ [item]) } | None eq: Hdestruct with inspect (decide (i_n.1 ∈ indices)) => | (right _) eq: Hni => (s', []) | (left _) eq: Hi with inspect (decide (composite_state_destructor s' i_n.1 = [])) => | (right _) eq: Hn => (s', []) | (left _) eq: Hnn => indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 indices).message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM), (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) (Some (item, s)) → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in constrained_state_prop Free sby cbn; intros; eapply input_valid_transition_origin, composite_state_destructor_lookup_reachable. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM), (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) (Some (item, s)) → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in constrained_state_prop Free smessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM), (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) (Some (item, s)) → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in length indices + composite_state_size s < length (i :: l0) + composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM), (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) (Some (item, s)) → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in length indices + composite_state_size s < length (i :: l0) + composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i :: l0) = i_n
indexed_composite_state_to_trace: choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < S (length l0 + composite_state_size s') → composite_state IM * list (composite_transition_item IM)S (length l0 + composite_state_size s) < S (length l0 + composite_state_size s')by eapply composite_tv_state_destructor_size, elem_of_list_lookup_2. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i :: l0) = i_n
indexed_composite_state_to_trace: choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < S (length l0 + composite_state_size s') → composite_state IM * list (composite_transition_item IM)composite_state_size s < composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []), (λ y : {composite_state_destructor s' i_n.1 = []} + {composite_state_destructor s' i_n.1 ≠ []}, decide (composite_state_destructor s' i_n.1 = []) = y) (left e1) → ∀ e0 : i_n.1 ∈ i :: l0, (λ y : {i_n.1 ∈ i :: l0} + {i_n.1 ∉ i :: l0}, decide (i_n.1 ∈ i :: l0) = y) (left e0) → (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) None → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in length (StdppListSet.set_remove i_n.1 indices) + composite_state_size s' < length (i :: l0) + composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []), (λ y : {composite_state_destructor s' i_n.1 = []} + {composite_state_destructor s' i_n.1 ≠ []}, decide (composite_state_destructor s' i_n.1 = []) = y) (left e1) → ∀ e0 : i_n.1 ∈ i :: l0, (λ y : {i_n.1 ∈ i :: l0} + {i_n.1 ∉ i :: l0}, decide (i_n.1 ∈ i :: l0) = y) (left e0) → (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) None → (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n → (choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)) → let indices := i :: l0 in length (StdppListSet.set_remove i_n.1 indices) + composite_state_size s' < length (i :: l0) + composite_state_size s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
l0: list index
i_n: (index * nat)%type
e1: composite_state_destructor s' i_n.1 = []
Hnn: (λ y : {composite_state_destructor s' i_n.1 = []} + {composite_state_destructor s' i_n.1 ≠ []}, decide (composite_state_destructor s' i_n.1 = []) = y) (left e1)
e0: i_n.1 ∈ i :: l0
Hi: (λ y : {i_n.1 ∈ i :: l0} + {i_n.1 ∉ i :: l0}, decide (i_n.1 ∈ i :: l0) = y) (left e0)
Hdestruct: (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) None
Hchoice: (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n
indexed_composite_state_to_trace: choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)
indices:= i :: l0: list indexlength (StdppListSet.set_remove i_n.1 indices) + composite_state_size s' < length (i :: l0) + composite_state_size s'by rewrite <- set_remove_length; [lia |]. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i: index
l0: list index
i_n: (index * nat)%type
e1: composite_state_destructor s' i_n.1 = []
Hnn: (λ y : {composite_state_destructor s' i_n.1 = []} + {composite_state_destructor s' i_n.1 ≠ []}, decide (composite_state_destructor s' i_n.1 = []) = y) (left e1)
e0: i_n.1 ∈ i :: l0
Hi: (λ y : {i_n.1 ∈ i :: l0} + {i_n.1 ∉ i :: l0}, decide (i_n.1 ∈ i :: l0) = y) (left e0)
Hdestruct: (λ y : option (composite_transition_item IM * composite_state IM), composite_state_destructor s' i_n.1 !! i_n.2 = y) None
Hchoice: (λ y : index * nat, choose s' Hs' (i :: l0) = y) i_n
indexed_composite_state_to_trace: choice_function → ∀ x : composite_state IM, constrained_state_prop Free x → ∀ x0 : list index, length x0 + composite_state_size x < length (i :: l0) + composite_state_size s' → composite_state IM * list (composite_transition_item IM)
indices:= i :: l0: list indexS (length (StdppListSet.set_remove i_n.1 indices)) ≤ length indices
For any composite constrained state
s
and for any index i
for which the
component state s i
has the initial_state_property, the component state of
index i
of the origin state of the trace extracted from s
is equal
to s i
.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index) (i : index), initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → is i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index) (i : index), initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → is i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → is i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (i :: l0) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is, tr) → (∀ i0 : index, initial_state_prop (s i0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = ( is0, tr0) → is0 i0 = s i0) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → is0 i0 = s' i0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (n : i_n.1 ∉ i :: l0) (Hni : decide (i_n.1 ∈ i :: l0) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (i_n.1 ∈ i :: l0)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i0 = s' i0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), (∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → is i0 = s' i0) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = ( is, tr) → is i0 = s' i0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i0 = s' i0by inversion 2.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (i :: l0) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is, tr) → (∀ i0 : index, initial_state_prop (s i0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is0, tr0) → is0 i0 = s i0) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → is0 i0 = s' i0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i0: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i0 :: l0) = i_n
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → is i = s i
i: index
Hi: initial_state_prop (s' i)is' i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i0: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i0 :: l0) = i_n
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → is i = s i
i: index
Hi: initial_state_prop (s i)is' i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i0: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i0 :: l0) = i_n
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → is i = s i
i: index
Hi: initial_state_prop (s' i)s i = s' iby eapply Hind.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i0: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i0 :: l0) = i_n
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → is i = s i
i: index
Hi: initial_state_prop (s i)is' i = s iby eapply composite_tv_state_destructor_reflects_initiality; [| eapply elem_of_list_lookup_2 |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
i0: index
l0: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
Hchoice: choose s' Hs' (i0 :: l0) = i_n
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → is i = s i
i: index
Hi: initial_state_prop (s' i)s i = s' iby inversion 5.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (n : i_n.1 ∉ i :: l0) (Hni : decide (i_n.1 ∈ i :: l0) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (i_n.1 ∈ i :: l0)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i0 = s' i0by intros idx indices [j nj] Hdestruct _ Hj _ _ -> Hind _ _ _ _ i Hi is tr Heqis_tr; eapply Hind.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), (∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → is i0 = s' i0) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → is i0 = s' i0by inversion 6. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (i : index) (l0 : list index) (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → ∀ i0 : index, initial_state_prop (s' i0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → is i0 = s' i0
The components of an origin of a trace extracted through indexed_composite_state_to_trace
are initial for any index from the set of indices provided as argument.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ i : index, i ∈ indices → initial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ i : index, i ∈ indices → initial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list indexchoosing_well choose → NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ i : index, i ∈ indices → initial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM), constrained_state_prop Free s' → choosing_well choose → NoDup [] → not_in_indices_initial_prop s' [] → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i : index, i ∈ [] → initial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (i :: l0) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is, tr) → (choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s (i :: l0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = ( is0, tr0) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is0 i0)) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is0 i0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (n : i_n.1 ∉ i :: l0) (Hni : decide (i_n.1 ∈ i :: l0) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (i_n.1 ∈ i :: l0)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), (choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (i :: l0)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (i :: l0)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → ∀ i0 : index, i0 ∈ StdppListSet.set_remove i_n.1 (i :: l0) → initial_state_prop (is i0)) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = ( is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)by intros ? ? ? ? ? ? ? ? ? ? Hi; inversion Hi.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM), constrained_state_prop Free s' → choosing_well choose → NoDup [] → not_in_indices_initial_prop s' [] → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i : index, i ∈ [] → initial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (i :: l0) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is, tr) → (choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s (i :: l0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (i :: l0) = (is0, tr0) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is0 i0)) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is0 i0)by eapply Hind; [| | eapply composite_tv_state_destructor_preserves_not_in_indices_initial | |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → ∀ i : index, i ∈ idx :: indices → initial_state_prop (is i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
i: index
Hi: i ∈ idx :: indicesinitial_state_prop (is' i)by intros; subst; elim n; apply cw_chosen_index_in_indices; [apply Hchoose |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (n : i_n.1 ∉ i :: l0) (Hni : decide (i_n.1 ∈ i :: l0) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (i_n.1 ∈ i :: l0)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), (choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (i :: l0)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (i :: l0)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → ∀ i0 : index, i0 ∈ StdppListSet.set_remove i_n.1 (i :: l0) → initial_state_prop (is i0)) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (i :: l0)) = (is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = (is, tr) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indicesinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indicesinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: i ∈ idx :: indicesinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ removedinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: i ∈ idx :: indices
n: i ∉ removedinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ removedinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices)NoDup (StdppListSet.set_remove (j, nj).1 (idx :: indices))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices)not_in_indices_initial_prop s' (StdppListSet.set_remove (j, nj).1 (idx :: indices))by apply StdppListSet.set_remove_nodup.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices)NoDup (StdppListSet.set_remove (j, nj).1 (idx :: indices))by apply set_remove_preserves_not_in_indices_initial.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices) → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr)
i: index
Hi: i ∈ idx :: indices
e: i ∈ StdppListSet.set_remove ( j, nj).1 (idx :: indices)not_in_indices_initial_prop s' (StdppListSet.set_remove (j, nj).1 (idx :: indices))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: i ∈ idx :: indices
n: i ∉ removedinitial_state_prop (is i)by erewrite indexed_composite_state_to_trace_reflects_initiality_1; [.. | done]; apply composite_tv_state_destructor_initial.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' (j, nj).1 = []
Hj: (j, nj).1 ∈ idx :: indices
is: composite_state IM
tr: list (composite_transition_item IM)
removed: StdppListSet.set index
Heqremoved: removed = StdppListSet.set_remove ( j, nj).1 (idx :: indices)
Hind: choosing_well choose → NoDup removed → not_in_indices_initial_prop s' removed → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ i : index, i ∈ removed → initial_state_prop (is0 i)
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' removed = (is, tr)
i: index
Hi: j ∈ idx :: indices
n: j ∉ removedinitial_state_prop (is j)by intros; subst; clear Heq1; contradict Hdestruct; eapply choosing_well_position_exists; [apply Hchoose |]. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (i : index) (l0 : list index) (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ i :: l0) (Hi : decide (i_n.1 ∈ i :: l0) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (i :: l0) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ i :: l0)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (i :: l0)) = i_n ↾ Hchoice → choosing_well choose → NoDup (i :: l0) → not_in_indices_initial_prop s' (i :: l0) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ i0 : index, i0 ∈ i :: l0 → initial_state_prop (is i0)
Assuming not_in_indices_initial_prop, the origin state of the
trace given by indexed_composite_state_to_trace is initial.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → composite_initial_state_prop IM ismessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → composite_initial_state_prop IM ismessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: indexinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: index
e: i ∈ indicesinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: index
n: i ∉ indicesinitial_state_prop (is i)by eapply indexed_composite_state_to_trace_result_state.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: index
e: i ∈ indicesinitial_state_prop (is i)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: index
n: i ∉ indicesinitial_state_prop (is i)by erewrite indexed_composite_state_to_trace_reflects_initiality_1. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
H1: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index
H2: NoDup indices
Hinitial: not_in_indices_initial_prop s indices
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s Hs indices = (is, tr)
i: index
n: i ∉ indices
H4: initial_state_prop (s i)initial_state_prop (is i)
For any composite constrained state
s
and for any index i
for which the
component state s i
satisfies initial_state_prop, the component state of
index i
of any state of the trace extracted from s
is equal to s i
.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index) (i : index), initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (choose : choice_function) (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index) (i : index), initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list index∀ i : index, initial_state_prop (s i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr) → (∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is0, tr0) → ∀ item0 : transition_item, item0 ∈ tr0 → destination item0 i = s i) → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ item0 : transition_item, item0 ∈ tr0 → destination item0 i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : (j, nj).1 ∉ idx :: indices) (Hni : decide ((j, nj).1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide ((j, nj).1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (e1 : composite_state_destructor s' (j, nj).1 = []) (Hnn : decide (composite_state_destructor s' (j, nj).1 = []) = left e1) (e0 : (j, nj).1 ∈ idx :: indices) (Hi : decide ((j, nj).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), (∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' i) → inspect (decide (composite_state_destructor s' (j, nj).1 = [])) = left e1 ↾ Hnn → inspect (decide ((j, nj).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, nj).1 (idx :: indices)) = ( is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : composite_state_destructor s' (j, nj).1 ≠ []) (Hn : decide (composite_state_destructor s' (j, nj).1 = []) = right n) (e0 : (j, nj).1 ∈ idx :: indices) (Hi : decide ((j, nj).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide (composite_state_destructor s' (j, nj).1 = [])) = right n ↾ Hn → inspect (decide ((j, nj).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr) → (∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is0, tr0) → ∀ item0 : transition_item, item0 ∈ tr0 → destination item0 i = s i) → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ item0 : transition_item, item0 ∈ tr0 → destination item0 i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ tr ++ [item]destination _item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ trdestination _item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ [item]destination _item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ trdestination _item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s i)
_item: transition_item
H_item: _item ∈ trdestination _item i = s imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ trs i = s' iby eapply Hind.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s i)
_item: transition_item
H_item: _item ∈ trdestination _item i = s iby eapply composite_tv_state_destructor_reflects_initiality; [| eapply elem_of_list_lookup_2 |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ trs i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)
_item: transition_item
H_item: _item ∈ [item]destination _item i = s' imessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)destination item i = s' iby eapply elem_of_list_lookup_2.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, nj).1 !! (j, nj).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Hind: ∀ i : index, initial_state_prop (s i) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ item : transition_item, item ∈ tr0 → destination item i = s i
i: index
Hi: initial_state_prop (s' i)(item, ?s) ∈ composite_state_destructor s' ?iby inversion 5; inversion 1.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : (j, nj).1 ∉ idx :: indices) (Hni : decide ((j, nj).1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide ((j, nj).1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' iby eauto.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (e1 : composite_state_destructor s' (j, nj).1 = []) (Hnn : decide (composite_state_destructor s' (j, nj).1 = []) = left e1) (e0 : (j, nj).1 ∈ idx :: indices) (Hi : decide ((j, nj).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), (∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove (j, nj).1 (idx :: indices)) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' i) → inspect (decide (composite_state_destructor s' (j, nj).1 = [])) = left e1 ↾ Hnn → inspect (decide ((j, nj).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove (j, nj).1 (idx :: indices)) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' iby inversion 6; inversion 1. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : composite_state_destructor s' (j, nj).1 ≠ []) (Hn : decide (composite_state_destructor s' (j, nj).1 = []) = right n) (e0 : (j, nj).1 ∈ idx :: indices) (Hi : decide ((j, nj).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, nj).1 !! (j, nj).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide (composite_state_destructor s' (j, nj).1 = [])) = right n ↾ Hn → inspect (decide ((j, nj).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, nj).1 !! (j, nj).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → ∀ i : index, initial_state_prop (s' i) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ item : transition_item, item ∈ tr → destination item i = s' i
The trace extracted from a state using indexed_composite_state_to_trace
ends in that state.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index), NoDup indices → not_in_indices_initial_prop s' indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' indices = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s' : composite_state IM) (Hs' : constrained_state_prop Free s') (indices : list index), NoDup indices → not_in_indices_initial_prop s' indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' indices = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list indexchoosing_well choose → NoDup indices → not_in_indices_initial_prop s' indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' indices = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, jn) item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, jn) item s Hdestruct) (idx :: indices) = (is0, tr0) → finite_trace_last is0 tr0 = s) → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → finite_trace_last is0 tr0 = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (n : (j, jn).1 ∉ idx :: indices) (Hni : decide ((j, jn).1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), inspect (decide ((j, jn).1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (e1 : composite_state_destructor s' (j, jn).1 = []) (Hnn : decide (composite_state_destructor s' (j, jn).1 = []) = left e1) (e0 : (j, jn).1 ∈ idx :: indices) (Hi : decide ((j, jn).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), (choosing_well choose → NoDup (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = ( is, tr) → finite_trace_last is tr = s') → inspect (decide (composite_state_destructor s' (j, jn).1 = [])) = left e1 ↾ Hnn → inspect (decide ((j, jn).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove (j, jn).1 (idx :: indices)) = ( is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (n : composite_state_destructor s' (j, jn).1 ≠ []) (Hn : decide (composite_state_destructor s' (j, jn).1 = []) = right n) (e0 : (j, jn).1 ∈ idx :: indices) (Hi : decide ((j, jn).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), inspect (decide (composite_state_destructor s' (j, jn).1 = [])) = right n ↾ Hn → inspect (decide ((j, jn).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, jn) item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, jn) item s Hdestruct) (idx :: indices) = (is0, tr0) → finite_trace_last is0 tr0 = s) → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → finite_trace_last is0 tr0 = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
tr: list (composite_transition_item IM)
is0: composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is0, tr) = (is, tr0) → finite_trace_last is tr0 = sfinite_trace_last is0 (tr ++ [item]) = s'by eapply composite_tv_state_destructor_destination, elem_of_list_lookup_2.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
tr: list (composite_transition_item IM)
is0: composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is0, tr) = (is, tr0) → finite_trace_last is tr0 = sdestination item = s'by inversion 7.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (n : (j, jn).1 ∉ idx :: indices) (Hni : decide ((j, jn).1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), inspect (decide ((j, jn).1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (e1 : composite_state_destructor s' (j, jn).1 = []) (Hnn : decide (composite_state_destructor s' (j, jn).1 = []) = left e1) (e0 : (j, jn).1 ∈ idx :: indices) (Hi : decide ((j, jn).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), (choosing_well choose → NoDup (StdppListSet.set_remove (j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove (j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove (j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s') → inspect (decide (composite_state_destructor s' (j, jn).1 = [])) = left e1 ↾ Hnn → inspect (decide ((j, jn).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove (j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
e1: composite_state_destructor s' (j, jn).1 = []
Hnn: decide (composite_state_destructor s' (j, jn).1 = []) = left e1
e0: (j, jn).1 ∈ idx :: indices
Hi: decide ((j, jn).1 ∈ idx :: indices) = left e0
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s'
Hwell: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinit: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H1: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = ( is, tr)NoDup (StdppListSet.set_remove (j, jn).1 (idx :: indices))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
e1: composite_state_destructor s' (j, jn).1 = []
Hnn: decide (composite_state_destructor s' (j, jn).1 = []) = left e1
e0: (j, jn).1 ∈ idx :: indices
Hi: decide ((j, jn).1 ∈ idx :: indices) = left e0
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s'
Hwell: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinit: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H1: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = ( is, tr)not_in_indices_initial_prop s' (StdppListSet.set_remove (j, jn).1 (idx :: indices))by apply StdppListSet.set_remove_nodup.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
e1: composite_state_destructor s' (j, jn).1 = []
Hnn: decide (composite_state_destructor s' (j, jn).1 = []) = left e1
e0: (j, jn).1 ∈ idx :: indices
Hi: decide ((j, jn).1 ∈ idx :: indices) = left e0
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s'
Hwell: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinit: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H1: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = ( is, tr)NoDup (StdppListSet.set_remove (j, jn).1 (idx :: indices))by apply set_remove_preserves_not_in_indices_initial.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat
e1: composite_state_destructor s' (j, jn).1 = []
Hnn: decide (composite_state_destructor s' (j, jn).1 = []) = left e1
e0: (j, jn).1 ∈ idx :: indices
Hi: decide ((j, jn).1 ∈ idx :: indices) = left e0
Hdestruct: composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None
Hchoice: choose s' Hs' (idx :: indices) = (j, jn)
Hind: choosing_well choose → NoDup (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = (is, tr) → finite_trace_last is tr = s'
Hwell: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinit: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H1: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove ( j, jn).1 (idx :: indices)) = ( is, tr)not_in_indices_initial_prop s' (StdppListSet.set_remove (j, jn).1 (idx :: indices))by inversion 8. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
jn: nat∀ (n : composite_state_destructor s' (j, jn).1 ≠ []) (Hn : decide (composite_state_destructor s' (j, jn).1 = []) = right n) (e0 : (j, jn).1 ∈ idx :: indices) (Hi : decide ((j, jn).1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' (j, jn).1 !! (j, jn).2 = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, jn)), inspect (decide (composite_state_destructor s' (j, jn).1 = [])) = right n ↾ Hn → inspect (decide ((j, jn).1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' (j, jn).1 !! (j, jn).2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, jn) ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_trace_last is tr = s'
For any given composite constrained state,
indexed_reachable_composite_state_to_trace yields a constrained trace
leading to it.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → finite_constrained_trace_from_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (indices : list index), NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → finite_constrained_trace_from_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s: composite_state IM
Hs: constrained_state_prop Free s
indices: list indexchoosing_well choose → NoDup indices → not_in_indices_initial_prop s indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs indices = (is, tr) → finite_constrained_trace_from_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (idx :: indices) = (is0, tr0) → finite_constrained_trace_from_to Free is0 s tr0) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → finite_constrained_trace_from_to Free is0 s' tr0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (n : i_n.1 ∉ idx :: indices) (Hni : decide (i_n.1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), inspect (decide (i_n.1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_constrained_trace_from_to Free is s' trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ idx :: indices) (Hi : decide (i_n.1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), (choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' tr) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ idx :: indices) (Hi : decide (i_n.1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_constrained_trace_from_to Free is s' trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = i_n) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' i_n item s Hdestruct) (idx :: indices) = (is0, tr0) → finite_constrained_trace_from_to Free is0 s tr0) → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → finite_constrained_trace_from_to Free is0 s' tr0message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s' (idx :: indices)finite_constrained_trace_from_to Free is' s' (tr ++ [item])message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' i_n.1 !! i_n.2 = Some (item, s)
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s (idx :: indices)finite_constrained_trace_from_to Free is' s' (tr ++ [item])message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
item: composite_transition_item IM
s: composite_state IM
Hdestruct: (item, s) ∈ composite_state_destructor s' i_n.1
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s (idx :: indices)finite_constrained_trace_from_to Free is' s' (tr ++ [item])message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
item: composite_transition_item IM
Hs': constrained_state_prop Free (destination item)
idx: index
indices: list index
i_n: (index * nat)%type
s: composite_state IM
Hdestruct: (item, s) ∈ composite_state_destructor (destination item) i_n.1
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s (idx :: indices)finite_constrained_trace_from_to Free is' (destination item) (tr ++ [item])message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
item: composite_transition_item IM
Hs': constrained_state_prop Free (destination item)
idx: index
indices: list index
i_n: (index * nat)%type
s: composite_state IM
Hdestruct: input_constrained_transition_item Free s item
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s (idx :: indices)finite_constrained_trace_from_to Free is' (destination item) (tr ++ [item])by apply extend_right_finite_trace_from_to with s; [apply Hind |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
l: composite_label IM
input: option message
destination: composite_state IM
output: option message
Hs': constrained_state_prop Free destination
idx: index
indices: list index
i_n: (index * nat)%type
s: composite_state IM
Hdestruct: input_constrained_transition_item Free s {| l := l; input := input; destination := destination; output := output |}
tr: list (composite_transition_item IM)
is': composite_state IM
Hind: choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s (idx :: indices) → ∀ (is : composite_state IM) (tr0 : list (composite_transition_item IM)), (is', tr) = (is, tr0) → finite_constrained_trace_from_to Free is s tr0
Hchoose: choosing_well choose
Hnodup: NoDup (idx :: indices)
Hinitial: not_in_indices_initial_prop s (idx :: indices)finite_constrained_trace_from_to Free is' destination (tr ++ [{| l := l; input := input; destination := destination; output := output |}])by intros; contradiction n; subst; apply cw_chosen_index_in_indices; [apply Hchoose |].message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (n : i_n.1 ∉ idx :: indices) (Hni : decide (i_n.1 ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), inspect (decide (i_n.1 ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_constrained_trace_from_to Free is s' trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (e1 : composite_state_destructor s' i_n.1 = []) (Hnn : decide (composite_state_destructor s' i_n.1 = []) = left e1) (e0 : i_n.1 ∈ idx :: indices) (Hi : decide (i_n.1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), (choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr) → finite_constrained_trace_from_to Free is s' tr) → inspect (decide (composite_state_destructor s' i_n.1 = [])) = left e1 ↾ Hnn → inspect (decide (i_n.1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr) → finite_constrained_trace_from_to Free is s' trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 = []
Hj: i_n.1 ∈ idx :: indices
Hind: choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' tr
Hchoose: choosing_well choose
H1: NoDup (idx :: indices)
H2: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr)NoDup (StdppListSet.set_remove i_n.1 (idx :: indices))message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 = []
Hj: i_n.1 ∈ idx :: indices
Hind: choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' tr
Hchoose: choosing_well choose
H1: NoDup (idx :: indices)
H2: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr)not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices))by apply StdppListSet.set_remove_nodup.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 = []
Hj: i_n.1 ∈ idx :: indices
Hind: choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' tr
Hchoose: choosing_well choose
H1: NoDup (idx :: indices)
H2: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr)NoDup (StdppListSet.set_remove i_n.1 (idx :: indices))by apply set_remove_preserves_not_in_indices_initial.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
i_n: (index * nat)%type
Hdestruct: composite_state_destructor s' i_n.1 = []
Hj: i_n.1 ∈ idx :: indices
Hind: choosing_well choose → NoDup (StdppListSet.set_remove i_n.1 (idx :: indices)) → not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices)) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = ( is, tr) → finite_constrained_trace_from_to Free is s' tr
Hchoose: choosing_well choose
H1: NoDup (idx :: indices)
H2: not_in_indices_initial_prop s' (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
H3: indexed_composite_state_to_trace choose s' Hs' (StdppListSet.set_remove i_n.1 (idx :: indices)) = (is, tr)not_in_indices_initial_prop s' (StdppListSet.set_remove i_n.1 (idx :: indices))by intros; clear Heq1; contradict Hdestruct; subst; eapply choosing_well_position_exists; [apply Hchoose |]. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index∀ (i_n : index * nat) (n : composite_state_destructor s' i_n.1 ≠ []) (Hn : decide (composite_state_destructor s' i_n.1 = []) = right n) (e0 : i_n.1 ∈ idx :: indices) (Hi : decide (i_n.1 ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' i_n.1 !! i_n.2 = None) (Hchoice : choose s' Hs' (idx :: indices) = i_n), inspect (decide (composite_state_destructor s' i_n.1 = [])) = right n ↾ Hn → inspect (decide (i_n.1 ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' i_n.1 !! i_n.2) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = i_n ↾ Hchoice → choosing_well choose → NoDup (idx :: indices) → not_in_indices_initial_prop s' (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → finite_constrained_trace_from_to Free is s' tr
Given a composite constrained state, and a choice_function, we can extract a
constrained trace leading to that state by instantiating
indexed_composite_state_to_trace with the set of all indices.
Definition composite_state_to_trace (choose : choice_function) (s : composite_state IM) (Hs : constrained_state_prop Free s) : composite_state IM * list (composite_transition_item IM) := indexed_composite_state_to_trace choose s Hs (enum index).message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (is : composite_state IM) (tr : list (composite_transition_item IM)), composite_state_to_trace choose s Hs = (is, tr) → finite_constrained_trace_init_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ choose : choice_function, choosing_well choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (is : composite_state IM) (tr : list (composite_transition_item IM)), composite_state_to_trace choose s Hs = (is, tr) → finite_constrained_trace_init_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)finite_constrained_trace_init_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)finite_constrained_trace_init_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)
Hinitial: not_in_indices_initial_prop s (enum index)finite_constrained_trace_init_to Free is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)
Hinitial: not_in_indices_initial_prop s (enum index)finite_valid_trace_from_to (preloaded_with_all_messages_vlsm Free) is s trmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)
Hinitial: not_in_indices_initial_prop s (enum index)initial_state_prop isby eapply indexed_reachable_composite_state_to_trace.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)
Hinitial: not_in_indices_initial_prop s (enum index)finite_valid_trace_from_to (preloaded_with_all_messages_vlsm Free) is s trby eapply indexed_composite_state_to_trace_result_state_is_initial. Qed.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
choose: choice_function
Hchoose: choosing_well choose
s: composite_state IM
Hs: constrained_state_prop Free s
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: composite_state_to_trace choose s Hs = (is, tr)
Hnodup: NoDup (enum index)
Hinitial: not_in_indices_initial_prop s (enum index)initial_state_prop is
The property of a choice_function with respect to a predicate on composite
states which states that for any chosen transition, if the predicate holds for
the source of the transition then it must also hold for its target.
See MinimalEquivocationTrace.minimal_equivocation_choice_monotone for an
example of such a function.
Definition chosen_transition_preserves_P
(P : composite_state IM -> Prop) (choose : choice_function) : Prop :=
forall (is : list index), NoDup is ->
forall (s' : composite_state IM) (Hs' : constrained_state_prop Free s'),
not_in_indices_initial_prop s' is ->
forall (i : index) (Hi : i ∈ is) (n : nat),
choose s' Hs' is = (i, n) ->
forall (s : composite_state IM) (item : composite_transition_item IM),
composite_state_destructor s' i !! n = Some (item, s) ->
P s -> P s'.
If a choice_function is choosing well and satisfying
chosen_transition_preserves_P w.r.t. a predicate
See MinimalEquivocationTrace.state_to_minimal_equivocation_trace_equivocation_monotonic
for an example of such a function.
P
then the corresponding
trace obtained through composite_state_to_trace will also preserve P
at every step.
message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (P : composite_state IM → Prop) (choose : choice_function), choosing_well choose → chosen_transition_preserves_P P choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (is : composite_state IM) (tr : list (composite_transition_item IM)), composite_state_to_trace choose s Hs = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (P : composite_state IM → Prop) (choose : choice_function), choosing_well choose → chosen_transition_preserves_P P choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (is : composite_state IM) (tr : list (composite_transition_item IM)), composite_state_to_trace choose s Hs = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message∀ (P : composite_state IM → Prop) (choose : choice_function), choosing_well choose → chosen_transition_preserves_P P choose → ∀ (s : composite_state IM) (Hs : constrained_state_prop Free s) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s Hs (enum index) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
s': composite_state IM
Hs': constrained_state_prop Free s'
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (enum index) = (is, tr)
pre, suf: list (composite_transition_item IM)
item: composite_transition_item IM
Heqtr: tr = pre ++ [item] ++ suf
Heqv: P (finite_trace_last is pre)P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
s': composite_state IM
Hs': constrained_state_prop Free s'∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (enum index) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
s': composite_state IM
Hs': constrained_state_prop Free s'NoDup (enum index) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (enum index) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
s': composite_state IM
Hs': constrained_state_prop Free s'
Hinitial: not_in_indices_initial_prop s' (enum index)NoDup (enum index) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (enum index) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
indices: list indexchoosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' indices → NoDup indices → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' indices = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ [item] ++ suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' j !! nj = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item0 : composite_transition_item IM), tr0 = pre ++ item0 :: suf → P (finite_trace_last is0 pre) → P (destination item0)) → inspect (composite_state_destructor s' j !! nj) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item0 : composite_transition_item IM), tr0 = pre ++ item0 :: suf → P (finite_trace_last is0 pre) → P (destination item0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : j ∉ idx :: indices) (Hni : decide (j ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = ( j, nj)), inspect (decide (j ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (e1 : composite_state_destructor s' j = []) (Hnn : decide (composite_state_destructor s' j = []) = left e1) (e0 : j ∈ idx :: indices) (Hi : decide (j ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = ( j, nj)), (choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → NoDup (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)) → inspect (decide (composite_state_destructor s' j = [])) = left e1 ↾ Hnn → inspect (decide (j ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = ( is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : composite_state_destructor s' j ≠ []) (Hn : decide (composite_state_destructor s' j = []) = right n) (e0 : j ∈ idx :: indices) (Hi : decide (j ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = ( j, nj)), inspect (decide (composite_state_destructor s' j = [])) = right n ↾ Hn → inspect (decide (j ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (item : composite_transition_item IM) (s : composite_state IM) (Hdestruct : composite_state_destructor s' j !! nj = Some (item, s)) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)) (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr) → (choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item0 : composite_transition_item IM), tr0 = pre ++ item0 :: suf → P (finite_trace_last is0 pre) → P (destination item0)) → inspect (composite_state_destructor s' j !! nj) = Some (item, s) ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr ++ [item]) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item0 : composite_transition_item IM), tr0 = pre ++ item0 :: suf → P (finite_trace_last is0 pre) → P (destination item0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr0 = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
pre, suf: list (composite_transition_item IM)
item0: composite_transition_item IM
Heqtr: tr ++ [item] = pre ++ item0 :: suf
Heqv: P (finite_trace_last is pre)P (destination item0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
item: composite_transition_item IM
s: composite_state IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
tr: list (composite_transition_item IM)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item s Hdestruct) (idx :: indices) = (is, tr)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr0 : list (composite_transition_item IM)), (is, tr) = (is0, tr0) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr0 = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
pre, suf: list (composite_transition_item IM)
item0: composite_transition_item IM
Heqtr: tr ++ [item] = pre ++ item0 :: suf
Heqv: P (finite_trace_last is pre)
Hinitials: not_in_indices_initial_prop s (idx :: indices)P (destination item0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Heqv: P (finite_trace_last is pre)
Hinitials: not_in_indices_initial_prop s (idx :: indices)P (destination item0)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Heqv: P (finite_trace_last is pre)
Hinitials: not_in_indices_initial_prop s (idx :: indices)P s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Hinitials: not_in_indices_initial_prop s (idx :: indices)
Heqv: P sP s'message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Hinitials: not_in_indices_initial_prop s (idx :: indices)
Heqv: P sj ∈ idx :: indicesmessage, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Hinitials: not_in_indices_initial_prop s (idx :: indices)
Heqv: P s
n: j ∉ idx :: indicesj ∈ idx :: indicesby clear -Hdestruct n; rewrite n in Hdestruct.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
s: composite_state IM
item0: composite_transition_item IM
Hdestruct: composite_state_destructor s' j !! nj = Some (item0, s)
Hchoice: choose s' Hs' (idx :: indices) = (j, nj)
is: composite_state IM
pre: list (composite_transition_item IM)
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s (idx :: indices) → NoDup (idx :: indices) → ∀ (is0 : composite_state IM) (tr : list (composite_transition_item IM)), (is, pre) = (is0, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is0 pre) → P (destination item)
Heq: indexed_composite_state_to_trace choose s (indexed_composite_state_to_trace_obligations_obligation_1 s' Hs' (j, nj) item0 s Hdestruct) (idx :: indices) = (is, pre)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
Hinitials: not_in_indices_initial_prop s (idx :: indices)
Heqv: P s
n: composite_state_destructor s' j = []j ∈ idx :: indicesby intros; inversion Heqis_tr; subst; destruct pre.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : j ∉ idx :: indices) (Hni : decide (j ∈ idx :: indices) = right n) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide (j ∈ idx :: indices)) = right n ↾ Hni → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (e1 : composite_state_destructor s' j = []) (Hnn : decide (composite_state_destructor s' j = []) = left e1) (e0 : j ∈ idx :: indices) (Hi : decide (j ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), (choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → NoDup (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)) → inspect (decide (composite_state_destructor s' j = [])) = left e1 ↾ Hnn → inspect (decide (j ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' j = []
Hj: j ∈ idx :: indices
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → NoDup (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = ( is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hinitial: not_in_indices_initial_prop s' (idx :: indices)
Hnodup: NoDup (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr)
pre, suf: list (composite_transition_item IM)
item: composite_transition_item IM
Heqtr: tr = pre ++ item :: suf
Heqv: P (finite_trace_last is pre)P (destination item)message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' j = []
Hj: j ∈ idx :: indices
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → NoDup (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = ( is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hnodup: NoDup (idx :: indices)
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr)
pre, suf: list (composite_transition_item IM)
item: composite_transition_item IM
Heqtr: tr = pre ++ item :: suf
Heqv: P (finite_trace_last is pre)
Hinitial: not_in_indices_initial_prop s' (StdppListSet.set_remove j (idx :: indices))P (destination item)by eapply Hind.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat
Hdestruct: composite_state_destructor s' j = []
Hj: j ∈ idx :: indices
Hind: choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → NoDup (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = ( is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)
Hwell: choosing_well choose
HchooseP: chosen_transition_preserves_P P choose
Hnodup: NoDup (StdppListSet.set_remove j (idx :: indices))
is: composite_state IM
tr: list (composite_transition_item IM)
Heqis_tr: indexed_composite_state_to_trace choose s' Hs' (if decide (j = idx) then indices else idx :: StdppListSet.set_remove j indices) = (is, tr)
pre, suf: list (composite_transition_item IM)
item: composite_transition_item IM
Heqtr: tr = pre ++ item :: suf
Heqv: P (finite_trace_last is pre)
Hinitial: not_in_indices_initial_prop s' (StdppListSet.set_remove j (idx :: indices))P (destination item)by intros; inversion Heqis_tr; subst; destruct pre. Qed. End sec_traceable_vlsm_composition.message, index: Type
EqDecision0: EqDecision index
H: finite.Finite index
IM: index → VLSM message
state_destructor: ∀ i : index, state (IM i) → list (transition_item * state (IM i))
state_size: ∀ i : index, state (IM i) → nat
H0: ∀ i : index, TraceableVLSM (IM i) (state_destructor i) (state_size i)
Free:= free_composite_vlsm IM: VLSM message
P: composite_state IM → Prop
choose: choice_function
s': composite_state IM
Hs': constrained_state_prop Free s'
idx: index
indices: list index
j: index
nj: nat∀ (n : composite_state_destructor s' j ≠ []) (Hn : decide (composite_state_destructor s' j = []) = right n) (e0 : j ∈ idx :: indices) (Hi : decide (j ∈ idx :: indices) = left e0) (Hdestruct : composite_state_destructor s' j !! nj = None) (Hchoice : choose s' Hs' (idx :: indices) = (j, nj)), inspect (decide (composite_state_destructor s' j = [])) = right n ↾ Hn → inspect (decide (j ∈ idx :: indices)) = left e0 ↾ Hi → inspect (composite_state_destructor s' j !! nj) = None ↾ Hdestruct → inspect (choose s' Hs' (idx :: indices)) = (j, nj) ↾ Hchoice → choosing_well choose → chosen_transition_preserves_P P choose → not_in_indices_initial_prop s' (idx :: indices) → NoDup (idx :: indices) → ∀ (is : composite_state IM) (tr : list (composite_transition_item IM)), (s', []) = (is, tr) → ∀ (pre suf : list (composite_transition_item IM)) (item : composite_transition_item IM), tr = pre ++ item :: suf → P (finite_trace_last is pre) → P (destination item)