diff --git a/_CoqProject b/_CoqProject index 90967b12..0d4411cd 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,7 +1,5 @@ -Q theories/VLSM VLSM --arg -w -arg -deprecated-hint-without-locality - theories/VLSM/Lib/Preamble.v theories/VLSM/Lib/SsrExport.v theories/VLSM/Lib/StdppExtras.v @@ -26,7 +24,7 @@ theories/VLSM/Core/VLSM.v theories/VLSM/Core/Composition.v theories/VLSM/Core/Plans.v theories/VLSM/Core/VLSMProjections.v -theories/VLSM/Core/Validating.v +theories/VLSM/Core/Validator.v theories/VLSM/Core/Decisions.v theories/VLSM/Core/MessageDependencies.v theories/VLSM/Core/ProjectionTraces.v diff --git a/theories/VLSM/Core/ByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces.v index 98266f4e..630f7671 100644 --- a/theories/VLSM/Core/ByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces.v @@ -1,6 +1,6 @@ From stdpp Require Import prelude. From Coq Require Import FinFun. -From VLSM Require Import Lib.Preamble Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.Validating. +From VLSM Require Import Lib.Preamble Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.Validator. (** * VLSM Byzantine Traces @@ -11,10 +11,10 @@ and equivalent with traces on the corresponding pre-loaded VLSM Note that, contrary to what one might think, the [byzantine_trace_prop]erty does not only capture traces exhibiting byzantine behavior, but also all -[protocol_trace]s (consequence of Lemma [vlsm_incl_pre_loaded_with_all_messages_vlsm]). +[valid_trace]s (consequence of Lemma [vlsm_incl_pre_loaded_with_all_messages_vlsm]). Therefore to avoid confusion we will call _proper byzantine traces_, or _traces exhibiting byzantine behavior_ the collection of traces with -the [byzantine_trace_prop]erty but without the [protocol_trace_prop]erty. +the [byzantine_trace_prop]erty but without the [valid_trace_prop]erty. In the remainder of this section, we fix a (regular) VLSM <> with signature <> and of type <>. @@ -44,7 +44,7 @@ Definition byzantine_trace_prop (tr : vTrace M) := exists (M' : VLSM message) (Proj := binary_free_composition_fst M M'), - protocol_trace_prop Proj tr. + valid_trace_prop Proj tr. (** @@ -56,7 +56,7 @@ Lemma byzantine_pre_loaded_with_all_messages (PreLoaded := pre_loaded_with_all_messages_vlsm M) (tr : vTrace M) (Hbyz : byzantine_trace_prop tr) - : protocol_trace_prop PreLoaded tr. + : valid_trace_prop PreLoaded tr. Proof. destruct Hbyz as [M' Htr]. simpl in Htr. @@ -154,7 +154,7 @@ Definition alternate_byzantine_trace_prop (tr : vTrace M) (Proj := binary_free_composition_fst M emit_any_message_vlsm) := - protocol_trace_prop Proj tr. + valid_trace_prop Proj tr. (** @@ -182,7 +182,7 @@ equivalent to the [byzantine_trace_prop]erty. Since we have already proven that the [alternate_byzantine_trace_prop]erty implies the [byzantine_trace_prop]erty (Lemma [byzantine_alt_byzantine]), and since we know that the traces with the [byzantine_trace_prop]erty -are [protocol_trace]s for the [pre_loaded_with_all_messages_vlsm], to prove the +are [valid_trace]s for the [pre_loaded_with_all_messages_vlsm], to prove the equivalence it is enough to close the circle by proving the [VLSM_incl]usion between the [pre_loaded_with_all_messages_vlsm] and the projection VLSM used in the definition of the [alternate_byzantine_trace_prop]erty. @@ -209,7 +209,7 @@ of <> into <> Lemma alt_pre_loaded_with_all_messages_incl : VLSM_incl Alt1 PreLoaded. Proof. - intros t Hpt. + intros t Hvt. apply byzantine_pre_loaded_with_all_messages. apply byzantine_alt_byzantine. assumption. @@ -219,29 +219,29 @@ of <> into <> To prove the reverse inclusion (between <> and <>) we will use the [basic_VLSM_incl] meta-result about proving inclusions bewteen VLSMs which states that -- if all [valid] messages in the first are [protocol_message]s in the second, and -- if all [protocol_state]s in the first are also [protocol_state]s in the second, -- and if all [protocol_transition]s in the first are also [protocol_transition]s +- if all [valid] messages in the first are [valid_message]s in the second, and +- if all [valid_state]s in the first are also [valid_state]s in the second, +- and if all [input_valid_transition]s in the first are also [input_valid_transition]s in the second, - then the first VLSM is included in the second. We will tackle each of these properties in the sequel. -First note that _all_ messages are [protocol_message]s for <>, as +First note that _all_ messages are [valid_message]s for <>, as [emit_any_message_vlsm] can generate any message without changing state. *) - Lemma alt_option_protocol_message + Lemma alt_option_valid_message (om : option message) - : option_protocol_message_prop Alt om. + : option_valid_message_prop Alt om. Proof. - destruct om as [m|];[|apply option_protocol_message_None]. + destruct om as [m|];[|apply option_valid_message_None]. pose (s :=proj1_sig (vs0 Alt): vstate Alt). unfold vstate in s. exists s. - assert (protocol_prop Alt s None) as Hs - by (apply protocol_initial_state;apply proj2_sig). - eapply (protocol_generated Alt) with s None s None (existT second _) + assert (valid_state_message_prop Alt s None) as Hs + by (apply valid_initial_state;apply proj2_sig). + eapply (valid_generated_state_message Alt) with s None s None (existT second _) ;[assumption|assumption|split; exact I|]. simpl. f_equal. rewrite state_update_id; reflexivity. @@ -251,12 +251,12 @@ First note that _all_ messages are [protocol_message]s for <>, as Using the above, it is straight-forward to show that: *) - Lemma alt_proj_option_protocol_message + Lemma alt_proj_option_valid_message (m : option message) - : option_protocol_message_prop Alt1 m. + : option_valid_message_prop Alt1 m. Proof. - apply protocol_message_projection. - apply alt_option_protocol_message. + apply valid_message_projection. + apply alt_option_valid_message. Qed. (** @@ -271,24 +271,24 @@ by simply setting to <> the corresponding component of the initial (binary_IM M emit_any_message_vlsm) first s. (** -Lifting a [protocol_state] of <> we obtain a [protocol_state] of <>. +Lifting a [valid_state] of <> we obtain a [valid_state] of <>. *) - Lemma preloaded_alt_protocol_state + Lemma preloaded_alt_valid_state (sj : state) (om : option message) - (Hp : protocol_prop PreLoaded sj om) - : protocol_state_prop Alt (lifted_alt_state sj). + (Hp : valid_state_message_prop PreLoaded sj om) + : valid_state_prop Alt (lifted_alt_state sj). Proof. - assert (protocol_state_prop PreLoaded sj) as Hsj + assert (valid_state_prop PreLoaded sj) as Hsj by (exists om;assumption);clear Hp. - induction Hsj using protocol_state_prop_ind. - - apply initial_is_protocol. + induction Hsj using valid_state_prop_ind. + - apply initial_state_is_valid. intros [];[exact Hs|exact I]. - exists om'. - assert (option_protocol_message_prop Alt om0) as Hom0 - by apply alt_option_protocol_message. - cut (protocol_transition Alt (existT first l) (lifted_alt_state s,om0) (lifted_alt_state s', om')) - ;[apply protocol_prop_transition_out|]. + assert (option_valid_message_prop Alt om0) as Hom0 + by apply alt_option_valid_message. + cut (input_valid_transition Alt (existT first l) (lifted_alt_state s,om0) (lifted_alt_state s', om')) + ;[apply input_valid_transition_outputs_valid_state_message|]. split. * split;[assumption|]. split;[assumption|]. @@ -315,13 +315,13 @@ results above to show that <> is included in <>. Proof. apply (basic_VLSM_incl (machine PreLoaded) (machine Alt1)) ; intro; intros; [assumption| | |apply H]. - - apply alt_proj_option_protocol_message. + - apply alt_proj_option_valid_message. - exists (lifted_alt_state s). split; [reflexivity|]. destruct Hv as [[_om Hps] [Hpm Hv]]. repeat split. - + apply preloaded_alt_protocol_state with _om; assumption. - + apply alt_option_protocol_message. + + apply preloaded_alt_valid_state with _om; assumption. + + apply alt_option_valid_message. + assumption. Qed. @@ -359,24 +359,24 @@ End ByzantineTraces. (** ** Byzantine fault tolerance for a single validator -Given that projections of validating VLSMs are equal to their corresponding +Given that projections of composition of validator VLSMs are equal to their corresponding pre-loaded with all messages VLSM ([pre_loaded_with_all_messages_validating_proj_eq]), -we can derive that for validating VLSMs, all their byzantine traces are -included in the [protocol_trace]s of their projection from the composition. +we can derive that for validators, all their byzantine traces are +included in the [valid_trace]s of their projection from the composition. *) -Lemma validating_component_byzantine_fault_tolerance +Lemma validator_component_byzantine_fault_tolerance message `{EqDecision index} (IM : index -> VLSM message) (constraint : composite_label IM -> composite_state IM * option message -> Prop) (i : index) - (Hvalidating: validating_projection_prop IM constraint i) + (Hvalidator: projection_validator_prop IM constraint i) : forall tr, byzantine_trace_prop (IM i) tr -> - protocol_trace_prop (composite_vlsm_constrained_projection IM constraint i) tr. + valid_trace_prop (composite_vlsm_constrained_projection IM constraint i) tr. Proof. intros tr Htr. apply - (VLSM_incl_protocol_trace - (pre_loaded_with_all_messages_validating_proj_incl _ _ _ Hvalidating)). + (VLSM_incl_valid_trace + (pre_loaded_with_all_messages_validator_proj_incl _ _ _ Hvalidator)). revert Htr. apply byzantine_pre_loaded_with_all_messages. Qed. @@ -385,11 +385,11 @@ Qed. (** ** Byzantine fault tolerance for a composition of validators In this section we show that if all components of a composite VLSM <> have -the [validating_projection_prop]erty, then its byzantine traces (that is, +the [projection_validator_prop]erty, then its byzantine traces (that is, traces obtained upon placing this composition in any, possibly adversarial, -context) are [protocol_traces] of <>. +context) are [valid_trace]s of <>. *) -Section composite_validating_byzantine_traces. +Section composite_validator_byzantine_traces. Context {message : Type} {index : Type} @@ -399,7 +399,7 @@ Section composite_validating_byzantine_traces. (X := composite_vlsm IM constraint) (PreLoadedX := pre_loaded_with_all_messages_vlsm X) (FreeX := free_composite_vlsm IM) - (Hvalidating: forall i : index, validating_projection_prop IM constraint i) + (Hvalidator: forall i : index, projection_validator_prop IM constraint i) . (** @@ -412,58 +412,58 @@ of <>, we just need to show that <> is included in <> to prove our main result. *) - Lemma validating_pre_loaded_with_all_messages_incl + Lemma validator_pre_loaded_with_all_messages_incl : VLSM_incl PreLoadedX X. Proof. apply VLSM_incl_finite_traces_characterization. intros. split; [|apply H]. destruct H as [Htr Hs]. - induction Htr using finite_protocol_trace_from_rev_ind. - - apply (finite_ptrace_empty X). - apply initial_is_protocol; assumption. + induction Htr using finite_valid_trace_from_rev_ind. + - apply (finite_valid_trace_from_empty X). + apply initial_state_is_valid; assumption. - specialize (IHHtr Hs) as IHtr; clear IHHtr. apply (extend_right_finite_trace_from X). assumption. destruct Hx as [Hvx Htx]. split; [|assumption]. - apply finite_ptrace_last_pstate in IHtr. + apply finite_valid_trace_last_pstate in IHtr. simpl in *. match type of IHtr with - | protocol_state_prop _ ?s => remember s as lst + | valid_state_prop _ ?s => remember s as lst end. split; [assumption|]. repeat split; [|apply Hvx|apply Hvx]. destruct Hvx as [Hlst [_ [Hv _]]]. destruct l as (i, li). simpl in *. - specialize (protocol_state_project_preloaded_to_preloaded _ IM constraint lst i Hlst) + specialize (valid_state_project_preloaded_to_preloaded _ IM constraint lst i Hlst) as Hlsti. - assert (Hpv : protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (lst i, iom)). + assert (Hiv : input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (lst i, iom)). { split; [assumption|]. split; [|assumption]. - eexists _. apply (pre_loaded_with_all_messages_message_protocol_prop (IM i)). + eexists _. apply (pre_loaded_with_all_messages_message_valid_initial_state_message (IM i)). } - apply Hvalidating in Hpv. - clear -Hpv. - destruct Hpv as [_ [_ [_ [Hinput _]]]]. + apply Hvalidator in Hiv. + clear -Hiv. + destruct Hiv as [_ [_ [_ [Hinput _]]]]. destruct Hinput as [s Hinput]. exists s. revert Hinput. - apply (constraint_subsumption_protocol_prop IM constraint). + apply (constraint_subsumption_valid_state_message_preservation IM constraint). intro. intros. destruct som. apply H. Qed. (** -Finally, we can conclude that a composition of validating components can +Finally, we can conclude that a composition of validator components can resist any kind of external influence: *) - Lemma composite_validating_byzantine_traces_are_not_byzantine + Lemma composite_validator_byzantine_traces_are_not_byzantine (tr : vTrace X) (Hbyz : byzantine_trace_prop X tr) - : protocol_trace_prop X tr. + : valid_trace_prop X tr. Proof. - apply validating_pre_loaded_with_all_messages_incl. + apply validator_pre_loaded_with_all_messages_incl. apply alt_pre_loaded_with_all_messages_incl. apply byzantine_alt_byzantine_iff. assumption. Qed. -End composite_validating_byzantine_traces. +End composite_validator_byzantine_traces. diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index ca426265..13d9a8d5 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -2,7 +2,7 @@ From stdpp Require Import prelude finite. From Coq Require Import FunctionalExtensionality. From VLSM.Lib Require Import Preamble StdppListSet FinFunExtras ListExtras. From VLSM Require Import Core.VLSM Core.MessageDependencies Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.SubProjectionTraces Core.ByzantineTraces. -From VLSM Require Import Core.Validating Core.Equivocation Core.EquivocationProjections Core.Equivocation.NoEquivocation Core.Equivocation.FixedSetEquivocation. +From VLSM Require Import Core.Validator Core.Equivocation Core.EquivocationProjections Core.Equivocation.NoEquivocation Core.Equivocation.FixedSetEquivocation. (** * VLSM Compositions with a fixed set of byzantine nodes @@ -11,13 +11,13 @@ replaced by byzantine nodes, i.e., nodes which can send arbitrary [channel_authenticated_message]s at any time, while the rest are protocol abiding (non-equivocating) nodes. -We will show that, if the non-byzantine nodes are validating for the +We will show that, if the non-byzantine nodes are validators for the composition with that specific fixed-set of nodes as equivocators, then the projections of traces to the non-byzantine nodes are precisely the projections of traces of the composition with that specific fixed-set of nodes as equivocators to the non-equivocating nodes. -That is, validating non-byzantine nodes do not distinguish between byzantine +That is, non-byzantine validator nodes do not distinguish between byzantine nodes and equivocating ones. Therefore, when analyzing the security of a protocol it suffices to consider equivocating nodes. *) @@ -146,7 +146,7 @@ Definition non_byzantine_not_equivocating_constraint (** The first definition of the [fixed_byzantine_trace_prop]erty: Fixed byzantine traces are projections to the subset of protocol-following nodes -of traces which are protocol for the composition in which a fixed set of nodes +of traces which are valid for the composition in which a fixed set of nodes were replaced by byzantine nodes and the rest are protocol-following (i.e., they are not equivocating). *) @@ -155,7 +155,7 @@ Definition fixed_byzantine_trace_prop (tr : list (composite_transition_item (sub_IM fixed_byzantine_IM non_byzantine))) : Prop := exists bis btr, - finite_protocol_trace (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) bis btr /\ + finite_valid_trace (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) bis btr /\ composite_state_sub_projection fixed_byzantine_IM non_byzantine bis = is /\ finite_trace_sub_projection fixed_byzantine_IM non_byzantine btr = tr. @@ -178,7 +178,7 @@ Proof. - intros Hs. exists (lift_sub_state fixed_byzantine_IM non_byzantine s). split. - + apply composite_state_sub_projection_lift. + + apply composite_state_sub_projection_lift_to. + apply (lift_sub_state_initial fixed_byzantine_IM non_byzantine); assumption. Qed. @@ -194,91 +194,40 @@ Proof. - revert H. apply induced_sub_projection_transition_preservation. Qed. -Lemma non_byzantine_lift_full_projection - : VLSM_full_projection fixed_non_byzantine_projection (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) - (lift_sub_label fixed_byzantine_IM non_byzantine) - (lift_sub_state fixed_byzantine_IM non_byzantine). +(** The induced projection from the composition of [fixed_byzantine_IM] under +the [non_byzantine_not_equivocating_constraint] to the non-byzantine nodes has +the [projection_friendly_prop]erty. +*) +Lemma fixed_non_byzantine_projection_friendliness + : projection_friendly_prop + (induced_sub_projection_is_projection fixed_byzantine_IM non_byzantine + non_byzantine_not_equivocating_constraint). Proof. - apply basic_VLSM_full_projection; intro; intros. - - destruct Hv as [_ [_ [lX [sX [Heql [Heqs [HsX [Hom [Hv Hc]]]]]]]]]. - destruct lX as [i li]. cbn in Hv, Hc. - unfold composite_label_sub_projection_option in Heql. - simpl in Heql. - case_decide; [|congruence]. - inversion Heql. subst l. clear Heql. - cbn. unfold constrained_composite_valid. cbn. - unfold lift_sub_state. - rewrite lift_sub_state_to_eq with (Hi := H). - subst. - split; [assumption|]. - destruct om as [m|]; [|exact I]. - destruct (sender m) as [v|]; [|exact I]. - simpl in *. - case_decide; [|exact I]. - rewrite lift_sub_state_to_eq with (Hi := H0). - assumption. - - apply proj2 in H. revert H. cbn. - destruct (vtransition _ _ _) as (si', _om'). - inversion 1. subst. clear H. - f_equal. - extensionality i. - destruct l as (sub_j, lj). - destruct_dec_sig sub_j j Hj Heqsub_j. - subst. - simpl. - destruct (decide (i = j)). - + subst. rewrite state_update_eq. - unfold lift_sub_state. - rewrite lift_sub_state_to_eq with (Hi := Hj). - unfold composite_state_sub_projection. - simpl. - rewrite state_update_eq. - reflexivity. - + rewrite state_update_neq by congruence. - destruct (decide (i ∈ non_byzantine)). - * unfold lift_sub_state. - rewrite! lift_sub_state_to_eq with (Hi := e). - unfold composite_state_sub_projection. simpl. - rewrite state_update_neq by congruence. - rewrite lift_sub_state_to_eq with (Hi := e). - reflexivity. - * unfold lift_sub_state, lift_sub_state_to. - destruct (decide _); [contradiction|reflexivity]. - - apply (lift_sub_state_initial fixed_byzantine_IM). - destruct H as [sX [Hs_pr HsX]]. - subst. - intro sub_i. - destruct_dec_sig sub_i i Hi Heqsub_i. - subst. apply HsX. - - assumption. + apply induced_sub_projection_friendliness. + apply induced_sub_projection_lift. + intros s1 s2 Heq l om. + destruct om as [m|]; [|intuition]. + cbn. + destruct (sender m) as [v|]; [|intuition]. + simpl. + case_decide as HAv; [|intuition]. + replace (s2 (A v)) with (s1 (A v)); [intuition|]. + exact (equal_f_dep Heq (dexist (A v) HAv)). Qed. (** Characterization result for the first definition: the [fixed_byzantine_trace_prop]erty is equivalent to the -[finite_protocol_trace_prop]erty of the [induced_sub_projection] of the +[finite_valid_trace_prop]erty of the [induced_sub_projection] of the the composition in which a fixed set of nodes were replaced by byzantine nodes and the rest are protocol-following to the set of protocol-following nodes. *) Lemma fixed_byzantine_trace_char1 is tr : fixed_byzantine_trace_prop is tr <-> - finite_protocol_trace fixed_non_byzantine_projection is tr. + finite_valid_trace fixed_non_byzantine_projection is tr. Proof. - split. - - intros [bis [btr [Hbtr [His Htr]]]]. - subst. - specialize - (induced_sub_projection_is_projection - fixed_byzantine_IM non_byzantine non_byzantine_not_equivocating_constraint) - as Hproj. - apply (VLSM_projection_finite_protocol_trace Hproj) in Hbtr. - replace (finite_trace_sub_projection _ _ _) with (VLSM_projection_trace_project Hproj btr) - ; [assumption|reflexivity]. - - intro Htr. - eexists _,_. - intuition. - + apply (VLSM_full_projection_finite_protocol_trace non_byzantine_lift_full_projection), Htr. - + apply composite_state_sub_projection_lift. - + apply composite_trace_sub_projection_lift. + apply and_comm. + apply (projection_friendly_trace_char (induced_sub_projection_is_projection _ _ _)). + apply fixed_non_byzantine_projection_friendliness. Qed. End fixed_byzantine_traces_as_projections. @@ -294,7 +243,7 @@ Section fixed_byzantine_traces_as_pre_loaded. Definition fixed_set_signed_message (m : message) : Prop := non_sub_index_authenticated_message non_byzantine A sender m /\ - (exists i, i ∈ non_byzantine /\ exists l s, protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, Some m)). + (exists i, i ∈ non_byzantine /\ exists l s, input_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, Some m)). (** Given that we're only looking at the composition of nodes in the @@ -421,10 +370,10 @@ Proof. + simpl in Hsent. simpl in HsY. apply - (preloaded_composite_sent_protocol (sub_IM fixed_byzantine_IM non_byzantine) + (preloaded_composite_sent_valid (sub_IM fixed_byzantine_IM non_byzantine) (finite_sub_index non_byzantine (listing_from_finite index)) sub_fixed_byzantine_IM_Hbs _ _ _ HsY _ Hsent). - + apply initial_message_is_protocol. cbn. right. assumption. + + apply initial_message_is_valid. cbn. right. assumption. - split. + apply proj2,proj2 in Hv. revert Hv. @@ -464,7 +413,7 @@ Proof. revert Hsent. apply (sub_IM_has_been_sent_iff_by_sender fixed_byzantine_IM non_byzantine A sender fixed_byzantine_IM_sender_safety sub_fixed_byzantine_IM_Hbs). - * apply (VLSM_incl_protocol_state + * apply (VLSM_incl_valid_state (composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages (sub_IM fixed_byzantine_IM non_byzantine) _ _)) in HsX. assumption. @@ -499,12 +448,12 @@ Proof. - destruct Hseeded as [[i [Hi Hsender]] Hvalid]. pose (X := (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint)). pose (s0 := proj1_sig (composite_s0 fixed_byzantine_IM)). - assert (Hs0 : protocol_prop X s0 None). - { apply (protocol_initial X). + assert (Hs0 : valid_state_message_prop X s0 None). + { apply (valid_initial_state_message X). - destruct (composite_s0 fixed_byzantine_IM); assumption. - exact I. } - specialize (protocol_generated X _ _ Hs0 _ _ Hs0) as Hgen. + specialize (valid_generated_state_message X _ _ Hs0 _ _ Hs0) as Hgen. unfold non_byzantine in Hi. rewrite set_diff_iff in Hi. apply not_and_r in Hi as [Hi | Hi]; [elim Hi; apply elem_of_enum|]. @@ -531,8 +480,8 @@ Qed. (** Since the [fixed_non_byzantine_projection] is an [induced_projection] of the composition of [fixed_byzantine_IM] with a [non_byzantine_not_equivocating_constraint], its initial_messages and validity -are derived from protocol messages and protocol validity of the larger -composition; therefore, the following simple result becomes vrey important. +are derived from valid messages and protocol validity of the larger +composition; therefore, the following simple result becomes very important. *) Lemma pre_loaded_fixed_non_byzantine_vlsm_lift : VLSM_full_projection pre_loaded_fixed_non_byzantine_vlsm' (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) @@ -551,18 +500,18 @@ Lemma pre_loaded_fixed_non_byzantine_incl Proof. apply basic_VLSM_incl; intro; intros. - apply fixed_non_byzantine_projection_initial_state_preservation; assumption. - - apply initial_message_is_protocol. + - apply initial_message_is_valid. apply (pre_loaded_fixed_non_byzantine_vlsm_lift_initial_message l s m). + assumption. + apply proj1 in Hv. revert Hv. - apply (VLSM_full_projection_protocol_state pre_loaded_fixed_non_byzantine_vlsm_lift). + apply (VLSM_full_projection_valid_state pre_loaded_fixed_non_byzantine_vlsm_lift). + assumption. - exists (lift_sub_label fixed_byzantine_IM non_byzantine l). exists (lift_sub_state fixed_byzantine_IM non_byzantine s). split; [apply composite_label_sub_projection_option_lift|]. split; [apply composite_state_sub_projection_lift|]. revert Hv. - apply (VLSM_full_projection_protocol_valid pre_loaded_fixed_non_byzantine_vlsm_lift). + apply (VLSM_full_projection_input_valid pre_loaded_fixed_non_byzantine_vlsm_lift). - apply proj2 in H. revert H. apply induced_sub_projection_transition_preservation. @@ -587,11 +536,11 @@ we introduce the following alternate definition to the [fixed_byzantine_trace_prop]erty, this time defined for (not necessarily valid) traces of the full composition, accepting any such trace as long as its projection to the non-byzantine nodes is indistinguishable from a projection -of a valid protocol trace of the composition of [fixed_byzantine_IM] under the +of a valid trace of the composition of [fixed_byzantine_IM] under the [non_byzantine_not_equivocating_constraint]. *) Definition fixed_byzantine_trace_alt_prop is tr : Prop := - finite_protocol_trace pre_loaded_fixed_non_byzantine_vlsm + finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm (composite_state_sub_projection IM non_byzantine is) (finite_trace_sub_projection IM non_byzantine tr). @@ -601,7 +550,7 @@ End fixed_byzantine_traces. In this section we show that while equivocators can always appear as byzantine to the protocol-abiding nodes, the converse is also true if the protocol- -abiding nodes satisfy the [validating_projection_prop]erty, which basically +abiding nodes satisfy the [projection_validator_prop]erty, which basically allows them to locally verify the authenticity of a received message. *) @@ -619,9 +568,10 @@ Context (A : validator -> index) (sender : message -> option validator) {finite_index : finite.Finite index} + (selection_complement := set_diff (enum index) selection) (PreNonByzantine : VLSM message := pre_loaded_fixed_non_byzantine_vlsm IM Hbs selection A sender) (Fixed : VLSM message := fixed_equivocation_vlsm_composition IM Hbs Hbr selection) - (FixedNonEquivocating : VLSM message := induced_sub_projection IM (set_diff (enum index) selection) (fixed_equivocation_constraint IM Hbs Hbr selection)) + (FixedNonEquivocating : VLSM message := induced_sub_projection IM selection_complement (fixed_equivocation_constraint IM Hbs Hbr selection)) (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) (can_emit_signed : channel_authentication_prop IM A sender) (Hsender_safety : sender_safety_alt_prop IM A sender := @@ -647,19 +597,19 @@ Proof. simpl. case_decide; [|exact I]. unfold sub_IM. simpl. - apply (VLSM_incl_protocol_state (constraint_free_incl IM (fixed_equivocation_constraint IM Hbs Hbr selection))) in Hs. - apply (VLSM_incl_protocol_state (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))) in Hs. - assert (Hpre_si : forall i, protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i)). + apply (VLSM_incl_valid_state (constraint_free_incl IM (fixed_equivocation_constraint IM Hbs Hbr selection))) in Hs. + apply (VLSM_incl_valid_state (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))) in Hs. + assert (Hpre_si : forall i, valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i)). { intro i. revert Hs. - apply protocol_state_project_preloaded_to_preloaded. + apply valid_state_project_preloaded_to_preloaded. } cut (@has_been_sent _ _ (Hbs (A v)) (s (A v)) m). { apply has_been_sent_irrelevance; intuition. } apply set_diff_elim2 in H. destruct Hstrong_v as [Hsent | Hemitted]. - destruct Hsent as [i [Hi Hsent]]. - apply protocol_state_has_trace in Hs as [is [tr Htr]]. + apply valid_state_has_trace in Hs as [is [tr Htr]]. apply (has_been_sent_iff_by_sender IM Hbs Hsender_safety Htr Hsender). exists i; assumption. - elim H. @@ -674,7 +624,7 @@ Proof. destruct H as [sX [HeqsX Hinitial]]. subst. apply Hinitial. - - apply (VLSM_incl_protocol_valid fixed_non_equivocating_incl_sub_non_equivocating) in Hv. + - apply (VLSM_incl_input_valid fixed_non_equivocating_incl_sub_non_equivocating) in Hv. destruct Hv as [Hs [_ Hv]]. specialize (sub_IM_no_equivocation_preservation IM (set_diff (enum index) selection) @@ -684,16 +634,16 @@ Proof. + simpl in Hsent. simpl in HsY. apply - (preloaded_composite_sent_protocol (sub_IM IM (set_diff (enum index) selection)) + (preloaded_composite_sent_valid (sub_IM IM (set_diff (enum index) selection)) (finite_sub_index (set_diff (enum index) selection) (listing_from_finite index)) (sub_has_been_sent_capabilities IM (set_diff (enum index) selection) Hbs) _ _ _ HsY _ Hsent). - + apply initial_message_is_protocol. + + apply initial_message_is_valid. right. split; [assumption|]. clear -Hv. apply induced_sub_projection_valid_projection in Hv; assumption. - - apply (VLSM_incl_protocol_valid fixed_non_equivocating_incl_sub_non_equivocating), + - apply (VLSM_incl_input_valid fixed_non_equivocating_incl_sub_non_equivocating), proj2, proj2 in Hv. split. + revert Hv. @@ -714,24 +664,24 @@ Proof. apply induced_sub_projection_transition_preservation. Qed. -(** As a corollary to the above result, we can conclude that valid protocol +(** As a corollary to the above result, we can conclude that valid traces for the composition with <> equivocation have the [fixed_byzantine_trace_alt_prop]erty. *) Corollary fixed_equivocating_traces_are_byzantine is tr - : finite_protocol_trace Fixed is tr -> + : finite_valid_trace Fixed is tr -> fixed_byzantine_trace_alt_prop IM Hbs selection A sender is tr. Proof. intro Htr. - apply (VLSM_incl_finite_protocol_trace fixed_non_equivocating_incl_fixed_non_byzantine). + apply (VLSM_incl_finite_valid_trace fixed_non_equivocating_incl_fixed_non_byzantine). specialize (induced_sub_projection_is_projection IM (set_diff (enum index) selection) (fixed_equivocation_constraint IM Hbs Hbr selection)) as Hproj. - apply (VLSM_projection_finite_protocol_trace Hproj); assumption. + apply (VLSM_projection_finite_valid_trace Hproj); assumption. Qed. -Section validating_fixed_set_byzantine. +Section validator_fixed_set_byzantine. Context (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) @@ -763,7 +713,7 @@ Proof. unfold lift_sub_state. rewrite lift_sub_state_to_eq with (Hi0 := Hi); assumption. + right. - apply can_emit_protocol_iff in HomY. + apply emitted_messages_are_valid_iff in HomY. destruct HomY as [[_v [[_im Him] Heqim]] | Hiom] ; [elim (no_initial_messages_in_IM _v _im); assumption|]. apply (VLSM_incl_can_emit (vlsm_incl_pre_loaded_with_all_messages_vlsm Fixed)) @@ -817,7 +767,7 @@ Context (lift_sub_state IM (set_diff (enum index) selection))). (** Since <> is an [induced_projection] of <>, -its initial_messages and validity are derived from protocol messages and +its initial_messages and validity are derived from valid messages and protocol validity of the larger composition; therefore, the following result becomes very important. *) @@ -838,20 +788,20 @@ Lemma fixed_non_byzantine_incl_fixed_non_equivocating_from_initial Proof. apply basic_VLSM_incl; intro; intros. - exists (lift_sub_state IM (set_diff (enum index) selection) s). - split; [apply composite_state_sub_projection_lift|]. + split; [apply composite_state_sub_projection_lift_to|]. by apply (lift_sub_state_initial IM). - - apply initial_message_is_protocol. + - apply initial_message_is_valid. apply (Hfixed_non_byzantine_vlsm_lift_initial_message l s m). + assumption. + apply proj1 in Hv. revert Hv. - apply (VLSM_full_projection_protocol_state fixed_non_byzantine_vlsm_lift_from_initial). + apply (VLSM_full_projection_valid_state fixed_non_byzantine_vlsm_lift_from_initial). + assumption. - exists (lift_sub_label IM (set_diff (enum index) selection) l). exists (lift_sub_state IM (set_diff (enum index) selection) s). split; [apply composite_label_sub_projection_option_lift|]. split; [apply composite_state_sub_projection_lift|]. revert Hv. - apply (VLSM_full_projection_protocol_valid fixed_non_byzantine_vlsm_lift_from_initial). + apply (VLSM_full_projection_input_valid fixed_non_byzantine_vlsm_lift_from_initial). - apply proj2 in H. revert H. apply induced_sub_projection_transition_preservation. @@ -869,12 +819,12 @@ Qed. End assuming_initial_messages_lift. Context - (Hvalidating: + (Hvalidator: forall i : index, i ∉ selection -> - validating_projection_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) + projection_validator_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) . -Lemma validating_fixed_non_byzantine_vlsm_lift_initial_message +Lemma validator_fixed_non_byzantine_vlsm_lift_initial_message : weak_full_projection_initial_message_preservation PreNonByzantine Fixed (lift_sub_state IM (set_diff (enum index) selection)). Proof. @@ -885,66 +835,55 @@ Proof. subst. unfold sub_IM in Him. simpl in Him. clear -Him. cbn. - apply initial_message_is_protocol. + apply initial_message_is_valid. by exists i, (exist _ m Him). - destruct Hseeded as [Hsigned [i [Hi [li [si Hpre_valid]]]]]. apply set_diff_elim2 in Hi. - specialize (Hvalidating i Hi _ _ Hpre_valid) + specialize (Hvalidator i Hi _ _ Hpre_valid) as [sX [Heqsi [HsX [Hm HvX]]]]. assumption. Qed. +(** The VLSM corresponding to the induced projection from a fixed-set byzantine +composition to the non-byzantine components is trace-equivalent to the VLSM +corresponding to the induced projection from the composition of regular nodes +under a fixed-set equivocation constraint to the same components. +*) +Lemma validator_fixed_non_byzantine_eq_fixed_non_equivocating + : VLSM_eq PreNonByzantine FixedNonEquivocating. +Proof. + apply fixed_non_byzantine_eq_fixed_non_equivocating_from_initial. + apply validator_fixed_non_byzantine_vlsm_lift_initial_message. +Qed. + (** ** The main result -Assuming that the non-byzantine nodes are validating for the +Assuming that the non-byzantine nodes are validators for the [fixed_equivocation_constraint] we give the following characterization of traces with the [fixed_byzantine_trace_alt_prop]erty in terms of equivocation: A trace has the [fixed_byzantine_trace_alt_prop]erty for a <> of -byzantine nodes iff there exists a protocol valid trace for the <> +byzantine nodes iff there exists a valid trace for the <> equivocation composition the projections of the two traces to the non-byzantine nodes, i.e., the nodes in the <> coincide. *) -Lemma validating_fixed_byzantine_traces_equivocation_char bis btr - (selection_complement := set_diff (enum index) selection) +Lemma validator_fixed_byzantine_traces_equivocation_char bis btr : fixed_byzantine_trace_alt_prop IM Hbs selection A sender bis btr <-> exists eis etr, - finite_protocol_trace Fixed eis etr /\ - composite_state_sub_projection IM selection_complement bis = composite_state_sub_projection IM selection_complement eis /\ - finite_trace_sub_projection IM selection_complement btr = finite_trace_sub_projection IM selection_complement etr. + finite_valid_trace Fixed eis etr /\ + composite_state_sub_projection IM selection_complement eis = composite_state_sub_projection IM selection_complement bis /\ + finite_trace_sub_projection IM selection_complement etr = finite_trace_sub_projection IM selection_complement btr. Proof. - split. - - intro Htr. - specialize (fixed_non_byzantine_vlsm_lift_from_initial validating_fixed_non_byzantine_vlsm_lift_initial_message) - as Hproj. - apply (VLSM_full_projection_finite_protocol_trace Hproj) in Htr. - eexists _,_; split; [exact Htr|]. - unfold lift_sub_state. - rewrite composite_state_sub_projection_lift. - split; [reflexivity|]. - symmetry. apply composite_trace_sub_projection_lift. - - intros [eis [etr [Hetr [Heis_pr Hetr_pr]]]]. - apply fixed_equivocating_traces_are_byzantine in Hetr. - revert Hetr. - subst selection_complement. - unfold fixed_byzantine_trace_alt_prop. - rewrite Heis_pr, Hetr_pr. - exact id. -Qed. - -(** A characterization at the level of VLSMs: the VLSM corresponding to -the induced projection from a fixed-set byzantine composition to the -non-byzantine components is trace-equivalent to the VLSM corresponding to -the induced projection from the composition of regular nodes under a fixed-set -equivocation constraint to the same components. -*) -Lemma validating_fixed_non_byzantine_eq_fixed_non_equivocating - : VLSM_eq PreNonByzantine FixedNonEquivocating. -Proof. - apply fixed_non_byzantine_eq_fixed_non_equivocating_from_initial. - apply validating_fixed_non_byzantine_vlsm_lift_initial_message. + unfold fixed_byzantine_trace_alt_prop. + split; intros Htr. + - apply fixed_non_equivocating_traces_char. + apply (VLSM_eq_finite_valid_trace validator_fixed_non_byzantine_eq_fixed_non_equivocating). + assumption. + - apply (VLSM_eq_finite_valid_trace validator_fixed_non_byzantine_eq_fixed_non_equivocating). + apply fixed_non_equivocating_traces_char. + assumption. Qed. -End validating_fixed_set_byzantine. +End validator_fixed_set_byzantine. End fixed_non_equivocating_vs_byzantine. diff --git a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v index 32808c30..d95e448b 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -3,14 +3,14 @@ From Coq Require Import FunctionalExtensionality Reals. From VLSM.Lib Require Import Preamble StdppListSet Measurable FinFunExtras StdppExtras. From VLSM Require Import Core.VLSM Core.MessageDependencies Core.VLSMProjections Core.Composition Core.SubProjectionTraces. From VLSM Require Import Core.ByzantineTraces Core.ByzantineTraces.FixedSetByzantineTraces. -From VLSM Require Import Core.Validating Core.Equivocation Core.Equivocation.FixedSetEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.TraceWiseEquivocation. +From VLSM Require Import Core.Validator Core.Equivocation Core.Equivocation.FixedSetEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.TraceWiseEquivocation. (** * VLSM Compositions with byzantine nodes of limited weight -In this module we define and study protocol traces allowing a (weight-)limited -amount of byzantine faults. +In this module we define and study protocol executions allowing a +(weight-)limited amount of byzantine faults. -We will show that, if the non-byzantine nodes are validating for a +We will show that, if the non-byzantine nodes are validators for a composition constraint allowing only a limited amount of equivocation, then they do not distinguish between byzantine nodes and equivocating ones, that is, projections of traces with byzantine faults to the non-byzantine nodes are @@ -57,7 +57,7 @@ Context {is_equivocating_tracewise_no_has_been_sent_dec : RelDecision (is_equivocating_tracewise_no_has_been_sent IM (fun i => i) sender)} (limited_constraint := limited_equivocation_constraint IM (listing_from_finite index) sender) (Limited : VLSM message := composite_vlsm IM limited_constraint) - (Hvalidating: forall i : index, validating_projection_prop IM limited_constraint i) + (Hvalidator: forall i : index, projection_validator_prop IM limited_constraint i) (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) (can_emit_signed : channel_authentication_prop IM Datatypes.id sender) (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) @@ -85,11 +85,11 @@ Context (tracewise_equivocating_validators := @equivocating_validators _ _ _ _ Htracewise_BasicEquivocation) . -(** When replacing the byzantine components of a composite [protocol_state] with +(** When replacing the byzantine components of a composite [valid_state] with initial states for those machines we obtain a state which is [not_heavy]. *) -Lemma limited_PreNonByzantine_protocol_state_lift_not_heavy s - (Hs : protocol_state_prop PreNonByzantine s) +Lemma limited_PreNonByzantine_valid_state_lift_not_heavy s + (Hs : valid_state_prop PreNonByzantine s) (sX := lift_sub_state IM non_byzantine s) : tracewise_not_heavy sX. Proof. @@ -102,10 +102,10 @@ Proof. - apply NoDup_remove_dups. - intros i Hi. apply elem_of_remove_dups, Hincl, Hi. } - apply protocol_state_has_trace in Hs as [is [tr Htr]]. + apply valid_state_has_trace in Hs as [is [tr Htr]]. specialize (preloaded_non_byzantine_vlsm_lift IM Hbs byzantine (fun i => i) sender) as Hproj. - apply (VLSM_full_projection_finite_protocol_trace_init_to Hproj) in Htr as Hpre_tr. + apply (VLSM_full_projection_finite_valid_trace_init_to Hproj) in Htr as Hpre_tr. intros v Hv. apply equivocating_validators_is_equivocating_tracewise_iff in Hv as Hvs'. specialize (Hvs' _ _ Hpre_tr). @@ -116,11 +116,11 @@ Proof. subst itemX. cbn in Hm0. change (pre ++ item::suf) with (pre ++ [item] ++ suf) in Htr. destruct Htr as [Htr Hinit]. - apply (finite_protocol_trace_from_to_app_split PreNonByzantine) in Htr. + apply (finite_valid_trace_from_to_app_split PreNonByzantine) in Htr. destruct Htr as [Hpre Hitem]. - apply (VLSM_full_projection_finite_protocol_trace_from_to Hproj) in Hpre as Hpre_pre. - apply ptrace_last_pstate in Hpre_pre as Hs_pre. - apply (finite_protocol_trace_from_to_app_split PreNonByzantine), proj1 in Hitem. + apply (VLSM_full_projection_finite_valid_trace_from_to Hproj) in Hpre as Hpre_pre. + apply valid_trace_last_pstate in Hpre_pre as Hs_pre. + apply (finite_valid_trace_from_to_app_split PreNonByzantine), proj1 in Hitem. inversion Hitem; subst; clear Htl Hitem. simpl in Hm0. subst. destruct Ht as [[_ [_ [_ [Hc _]]]] _]. destruct Hc as [[sub_i Hsenti] | Hemit]. @@ -147,7 +147,7 @@ Proof. apply set_diff_intro; [apply finite_index|assumption]. Qed. -(** When replacing the byzantine components of a composite [protocol_state] with +(** When replacing the byzantine components of a composite [valid_state] with initial states for those machines validity of transitions for the non-byzantine components is preserved. *) @@ -166,13 +166,13 @@ Proof. apply (lift_sub_transition IM non_byzantine) in Ht as HtX. simpl in HtX |- *. rewrite HtX. simpl. - apply limited_PreNonByzantine_protocol_state_lift_not_heavy. - eapply protocol_transition_destination; split; eassumption. + apply limited_PreNonByzantine_valid_state_lift_not_heavy. + eapply input_valid_transition_destination; split; eassumption. Qed. -(** By replacing the byzantine components of a composite [protocol_state] with +(** By replacing the byzantine components of a composite [valid_state] with initial states for those machines and ignoring transitions for byzantine nodes -we obtain valid protocol traces for the <> equivocation composition. +we obtain valid traces for the <> equivocation composition. *) Lemma limited_PreNonByzantine_vlsm_lift : VLSM_full_projection PreNonByzantine Limited @@ -188,11 +188,11 @@ Proof. destruct_dec_sig sub_i i Hi Heqsub_i. subst. unfold sub_IM in Him. simpl in Him. clear -Him. - apply initial_message_is_protocol. + apply initial_message_is_valid. exists i, (exist _ m Him); intuition. + destruct Hseeded as [Hsigned [i [Hi [li [si Hpre_valid]]]]]. apply set_diff_elim2 in Hi. - specialize (Hvalidating i _ _ Hpre_valid) + specialize (Hvalidator i _ _ Hpre_valid) as [_ [_ [_ [Hmsg _]]]]. assumption. Qed. @@ -201,40 +201,40 @@ End fixed_limited_selection. (** Given a trace with the [fixed_limited_byzantine_trace_prop]erty for a selection -of <> nodes, there exists protocol valid trace for the <> +of <> nodes, there exists a valid trace for the <> equivocation composition such that the projection of the two traces to the <> nodes coincide. *) -Lemma validating_fixed_limited_non_byzantine_traces_are_limited_non_equivocating s tr byzantine +Lemma validator_fixed_limited_non_byzantine_traces_are_limited_non_equivocating s tr byzantine (not_byzantine : set index := set_diff (enum index) byzantine) : fixed_limited_byzantine_trace_prop s tr byzantine -> exists bs btr, - finite_protocol_trace Limited bs btr /\ + finite_valid_trace Limited bs btr /\ composite_state_sub_projection IM not_byzantine s = composite_state_sub_projection IM not_byzantine bs /\ finite_trace_sub_projection IM not_byzantine tr = finite_trace_sub_projection IM not_byzantine btr. Proof. intros [Hlimit Hfixed]. eexists _, _; split. - - apply (VLSM_full_projection_finite_protocol_trace + - apply (VLSM_full_projection_finite_valid_trace (limited_PreNonByzantine_vlsm_lift byzantine Hlimit)). exact Hfixed. - unfold lift_sub_state. - rewrite composite_state_sub_projection_lift. + rewrite composite_state_sub_projection_lift_to. split; [reflexivity|]. symmetry. apply composite_trace_sub_projection_lift. Qed. (** ** The main result -Given any trace with the [limited_byzantine_trace_prop]erty, there exist -a protocol valid trace for the <> equivocation composition and +Given any trace with the [limited_byzantine_trace_prop]erty, there exists +a valid trace for the <> equivocation composition and a selection of nodes of limited weight such that the projection of the two traces to the nodes not in the selection coincide. *) -Lemma validating_limited_non_byzantine_traces_are_limited_non_equivocating s tr +Lemma validator_limited_non_byzantine_traces_are_limited_non_equivocating s tr : limited_byzantine_trace_prop s tr -> exists bs btr, - finite_protocol_trace Limited bs btr /\ + finite_valid_trace Limited bs btr /\ exists selection (selection_complement := set_diff (enum index) selection), (sum_weights (remove_dups selection) <= `threshold)%R /\ composite_state_sub_projection IM selection_complement s = composite_state_sub_projection IM selection_complement bs /\ @@ -242,7 +242,7 @@ Lemma validating_limited_non_byzantine_traces_are_limited_non_equivocating s tr Proof. intros [byzantine Hlimited]. apply proj1 in Hlimited as Hlimit. - apply validating_fixed_limited_non_byzantine_traces_are_limited_non_equivocating + apply validator_fixed_limited_non_byzantine_traces_are_limited_non_equivocating in Hlimited as [bs [btr [Hlimited [Hs_pr Htr_pr]]]]. exists bs, btr. split; [assumption|]. diff --git a/theories/VLSM/Core/Composition.v b/theories/VLSM/Core/Composition.v index e61da4dc..1be63f00 100644 --- a/theories/VLSM/Core/Composition.v +++ b/theories/VLSM/Core/Composition.v @@ -269,7 +269,7 @@ updating an initial composite state, say [s0], to <> on component <>. End composite_sig. - Section composite_vlsm. + Section sec_composite_vlsm. (** ** Constrained VLSM composition Assume an non-empty <> type, let @@ -342,7 +342,7 @@ the [composite_valid]ity. (l : composite_label) (s s' : composite_state) (om om' : option message) - (Ht : protocol_transition (composite_vlsm constraint) l (s, om) (s', om')) + (Ht : input_valid_transition (composite_vlsm constraint) l (s, om) (s', om')) (i : index) (Hi : i <> projT1 l) : s' i = s i. @@ -357,7 +357,7 @@ the [composite_valid]ity. (l : composite_label) (s s' : composite_state) (om om' : option message) - (Ht : protocol_transition (composite_vlsm constraint) l (s, om) (s', om')) + (Ht : input_valid_transition (composite_vlsm constraint) l (s, om) (s', om')) (il := projT1 l) : s' il = fst (vtransition (IM il) (projT2 l) (s il, om)). Proof. @@ -455,17 +455,17 @@ Thus, the [free_composite_vlsm] is the [composite_vlsm] using the apply basic_VLSM_strong_incl; cbv; intuition. Qed. - Lemma constraint_free_protocol_prop + Lemma constraint_free_valid_state_message_preservation (constraint : composite_label -> composite_state * option message -> Prop) s om - (Hsom : protocol_prop (composite_vlsm constraint) s om) - : protocol_prop free_composite_vlsm s om. + (Hsom : valid_state_message_prop (composite_vlsm constraint) s om) + : valid_state_message_prop free_composite_vlsm s om. Proof. revert Hsom. - apply (VLSM_incl_protocol_prop (constraint_free_incl constraint)); intro; intros; assumption. + apply (VLSM_incl_valid_state_message (constraint_free_incl constraint)); intro; intros; assumption. Qed. - Section protocol_constraint_subsumption. + Section sec_constraint_subsumption. (** ** Constraint subsumption *) (** @@ -479,7 +479,7 @@ than <> for any input. constraint1 l som -> constraint2 l som. (** -A weaker version of [strong_constraint_subsumption] requiring protocol-validity +A weaker version of [strong_constraint_subsumption] requiring [input_valid]ity w.r.t. [pre_loaded_with_all_messages_vlsm] as a precondition for the subsumption property. @@ -488,32 +488,32 @@ pre-loaded with all messages (Lemma [preloaded_constraint_subsumption_incl]). Although there are currently no explicit cases for its usage, it might be more useful than the [strong_constraint_subsumption] property in cases where proving -constraint subsumption relies on the state being protocol and/or the message +constraint subsumption relies on the state being valid and/or the message being valid. *) Definition preloaded_constraint_subsumption (constraint1 constraint2 : composite_label -> composite_state * option message -> Prop) := forall (l : composite_label) (som : state * option message), - protocol_valid (pre_loaded_with_all_messages_vlsm (composite_vlsm constraint1)) l som -> + input_valid (pre_loaded_with_all_messages_vlsm (composite_vlsm constraint1)) l som -> constraint2 l som. -(** -A weaker version of [preloaded_constraint_subsumption] requiring protocol-validity +(** +A weaker version of [preloaded_constraint_subsumption] requiring [input_valid]ity as a precondition for the subsumption property. This definition is usually useful in proving [VLSM_incl]usions between regular [VLSM]s (Lemma [constraint_subsumption_incl]). It is more useful than the [strong_constraint_subsumption] property in cases -where proving constraint subsumption relies on the state/message being protocol +where proving constraint subsumption relies on the state/message being valid and/or the message being valid (e.g., Lemma [Fixed_incl_StrongFixed]). *) - Definition protocol_constraint_subsumption + Definition input_valid_constraint_subsumption (constraint1 constraint2 : composite_label -> composite_state * option message -> Prop) := forall (l : composite_label) (som : composite_state * option message), - protocol_valid (composite_vlsm constraint1) l som -> constraint2 l som. + input_valid (composite_vlsm constraint1) l som -> constraint2 l som. Context (constraint1 constraint2 : composite_label -> composite_state * option message -> Prop) @@ -527,52 +527,52 @@ constraints <> and <>, respectively. Further assume that <> is subsumed by <>. We will show that <> is trace-included into <> by applying -Lemma [basic_VLSM_inclusion] +Lemma [basic_VLSM_incl] *) (* begin hide *) - Lemma constraint_subsumption_protocol_valid - (Hsubsumption : protocol_constraint_subsumption constraint1 constraint2) + Lemma constraint_subsumption_input_valid + (Hsubsumption : input_valid_constraint_subsumption constraint1 constraint2) (l : label) (s : state) (om : option message) - (Hv : protocol_valid X1 l (s, om)) + (Hv : input_valid X1 l (s, om)) : vvalid X2 l (s, om). Proof. split; [apply Hv|apply Hsubsumption]. assumption. Qed. - Lemma constraint_subsumption_protocol_prop - (Hsubsumption : protocol_constraint_subsumption constraint1 constraint2) + Lemma constraint_subsumption_valid_state_message_preservation + (Hsubsumption : input_valid_constraint_subsumption constraint1 constraint2) (s : state) (om : option message) - (Hps : protocol_prop X1 s om) - : protocol_prop X2 s om. + (Hps : valid_state_message_prop X1 s om) + : valid_state_message_prop X2 s om. Proof. induction Hps. - - apply (protocol_initial X2);assumption. - - apply (protocol_generated X2) with s _om _s om l; try assumption. - apply constraint_subsumption_protocol_valid; [assumption|]. - apply protocol_generated_valid with _om _s; assumption. + - apply (valid_initial_state_message X2);assumption. + - apply (valid_generated_state_message X2) with s _om _s om l; try assumption. + apply constraint_subsumption_input_valid; [assumption|]. + split;[|split];[exists _om|exists _s|];assumption. Qed. Lemma constraint_subsumption_incl - (Hsubsumption : protocol_constraint_subsumption constraint1 constraint2) + (Hsubsumption : input_valid_constraint_subsumption constraint1 constraint2) : VLSM_incl X1 X2. Proof. apply basic_VLSM_incl; intro; intros. - assumption. - - apply initial_message_is_protocol. assumption. - - apply constraint_subsumption_protocol_valid; assumption. + - apply initial_message_is_valid. assumption. + - apply constraint_subsumption_input_valid; assumption. - apply H. Qed. - Lemma preloaded_constraint_subsumption_protocol_valid + Lemma preloaded_constraint_subsumption_input_valid (Hpre_subsumption : preloaded_constraint_subsumption constraint1 constraint2) (l : label) (s : state) (om : option message) - (Hv : protocol_valid (pre_loaded_with_all_messages_vlsm X1) l (s, om)) + (Hv : input_valid (pre_loaded_with_all_messages_vlsm X1) l (s, om)) : vvalid X2 l (s, om). Proof. split; [apply Hv|apply Hpre_subsumption]. assumption. @@ -583,18 +583,18 @@ Lemma [basic_VLSM_inclusion] : VLSM_incl (pre_loaded_with_all_messages_vlsm X1) (pre_loaded_with_all_messages_vlsm X2). Proof. apply basic_VLSM_incl; intro; intros; [assumption| | |apply H]. - - apply initial_message_is_protocol. assumption. - - apply preloaded_constraint_subsumption_protocol_valid; assumption. + - apply initial_message_is_valid. assumption. + - apply preloaded_constraint_subsumption_input_valid; assumption. Qed. Lemma preloaded_constraint_subsumption_stronger (Hpre_subsumption : preloaded_constraint_subsumption constraint1 constraint2) - : protocol_constraint_subsumption constraint1 constraint2. + : input_valid_constraint_subsumption constraint1 constraint2. Proof. intros l som Hv. apply (Hpre_subsumption l som). destruct som. revert Hv. - apply (VLSM_incl_protocol_valid (vlsm_incl_pre_loaded_with_all_messages_vlsm (composite_vlsm constraint1))). + apply (VLSM_incl_input_valid (vlsm_incl_pre_loaded_with_all_messages_vlsm (composite_vlsm constraint1))). Qed. Lemma strong_constraint_subsumption_strongest @@ -615,7 +615,7 @@ Lemma [basic_VLSM_inclusion] Qed. (* end hide *) - End protocol_constraint_subsumption. + End sec_constraint_subsumption. Lemma preloaded_constraint_free_incl (constraint : composite_label -> composite_state * option message -> Prop) @@ -669,13 +669,13 @@ Lemma [basic_VLSM_inclusion] - intro; intros. apply lift_to_composite_state_initial. assumption. Qed. - Lemma protocol_state_preloaded_composite_free_lift + Lemma valid_state_preloaded_composite_free_lift (j : index) (sj : vstate (IM j)) - (Hp : protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM j)) sj) - : protocol_state_prop (pre_loaded_with_all_messages_vlsm free_composite_vlsm) (lift_to_composite_state j sj). + (Hp : valid_state_prop (pre_loaded_with_all_messages_vlsm (IM j)) sj) + : valid_state_prop (pre_loaded_with_all_messages_vlsm free_composite_vlsm) (lift_to_composite_state j sj). Proof. - apply (VLSM_full_projection_protocol_state (lift_to_composite_preloaded_vlsm_full_projection j)) + apply (VLSM_full_projection_valid_state (lift_to_composite_preloaded_vlsm_full_projection j)) ; assumption. Qed. @@ -706,21 +706,21 @@ Lemma [basic_VLSM_inclusion] - rewrite state_update_neq; try assumption. apply Hs. Qed. - (** Updating a composite protocol state for the free composition with - a component initial state yields a composite protocol state *) + (** Updating a composite [valid_state] for the free composition with + a component initial state yields a composite [valid_state] *) Lemma composite_free_update_state_with_initial (s : vstate free_composite_vlsm) - (Hs : protocol_state_prop free_composite_vlsm s) + (Hs : valid_state_prop free_composite_vlsm s) (i : index) (si : vstate (IM i)) (Hsi : vinitial_state_prop (IM i) si) - : protocol_state_prop free_composite_vlsm (state_update s i si). + : valid_state_prop free_composite_vlsm (state_update s i si). Proof. - generalize dependent s. apply protocol_state_prop_ind; intros. + generalize dependent s. apply valid_state_prop_ind; intros. - remember (state_update s i si) as s'. assert (Hs' : vinitial_state_prop free_composite_vlsm s') by (subst; apply composite_free_update_initial_state_with_initial; assumption). - apply initial_is_protocol. + apply initial_state_is_valid. assumption. - destruct Ht as [[Hps [Hom Hv]] Ht]. unfold transition in Ht. simpl in Ht. @@ -734,7 +734,7 @@ Lemma [basic_VLSM_inclusion] destruct Hs as [_om Hs]. destruct Hom as [_s Hom]. specialize - (protocol_generated free_composite_vlsm _ _ Hs _ _ Hom (existT j lj)) + (valid_generated_state_message free_composite_vlsm _ _ Hs _ _ Hom (existT j lj)) as Hgen. assert (n' : j <> i) by (intro contra; subst; elim n; reflexivity). spec Hgen. @@ -744,21 +744,21 @@ Lemma [basic_VLSM_inclusion] rewrite Htj in Hgen. eexists _. apply Hgen. reflexivity. Qed. - End composite_vlsm. + End sec_composite_vlsm. End VLSM_composition. (** These basic projection lemmas relate - the [protocol_state_prop] and [protocol_transition] of + the [valid_state_prop] and [input_valid_transition] of a composite VLSM back to those conditions holding over projections to individual components of the state. Because the composition may have validly produced - messages that are not protocol in an individual + messages that are not valid for an individual component (by interaction between components), - We cannot just use properties [protocol_prop (IM i)] - or [protocol_transition (IM i)]. + We cannot just use properties [valid_state_message_prop (IM i)] + or [input_valid_transition (IM i)]. For simplicity these lemmas use [pre_loaded_with_all_messages_vlsm (IM i)]. @@ -775,17 +775,17 @@ End VLSM_composition. precise. *) -Lemma protocol_state_project_preloaded_to_preloaded +Lemma valid_state_project_preloaded_to_preloaded message `{EqDecision index} (IM : index -> VLSM message) constraint (X:=composite_vlsm IM constraint) (s: vstate (pre_loaded_with_all_messages_vlsm X)) i: - protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> - protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i). + valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i). Proof. intros [om Hproto]. - apply preloaded_protocol_state_prop_iff. + apply preloaded_valid_state_prop_iff. induction Hproto. - - apply preloaded_protocol_initial_state. + - apply preloaded_valid_initial_state. apply (Hs i). - destruct l as [j lj]. simpl in Ht. unfold vtransition in Ht. simpl in Ht. @@ -800,18 +800,18 @@ Proof. exact IHHproto1. Qed. -Lemma protocol_state_project_preloaded +Lemma valid_state_project_preloaded message `{EqDecision index} (IM : index -> VLSM message) constraint (X:=composite_vlsm IM constraint) (s: vstate X) i: - protocol_state_prop X s -> - protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i). + valid_state_prop X s -> + valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) (s i). Proof. change (vstate X) with (vstate (pre_loaded_with_all_messages_vlsm X)) in s. intros [om Hproto]. - apply protocol_state_project_preloaded_to_preloaded. + apply valid_state_project_preloaded_to_preloaded. exists om. - apply preloaded_weaken_protocol_prop. + apply preloaded_weaken_valid_state_message_prop. assumption. Qed. @@ -831,13 +831,13 @@ Proof. reflexivity. Qed. -Lemma protocol_transition_preloaded_project_active +Lemma input_valid_transition_preloaded_project_active {message} `{EqDecision V} {IM: V -> VLSM message} {constraint} (X := composite_vlsm IM constraint) l s im s' om: - protocol_transition (pre_loaded_with_all_messages_vlsm X) l (s,im) (s',om) -> - protocol_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) - (s (projT1 l), im) (s' (projT1 l), om). + input_valid_transition (pre_loaded_with_all_messages_vlsm X) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) + (s (projT1 l), im) (s' (projT1 l), om). Proof. intro Hptrans. destruct Hptrans as [Hpvalid Htrans]. @@ -845,36 +845,36 @@ Proof. - destruct Hpvalid as [Hproto_s [_ Hcvalid]]. split;[|split]. + revert Hproto_s. - apply protocol_state_project_preloaded_to_preloaded. - + apply any_message_is_protocol_in_preloaded. + apply valid_state_project_preloaded_to_preloaded. + + apply any_message_is_valid_in_preloaded. + destruct l. apply Hcvalid. - apply composite_transition_project_active in Htrans;assumption. Qed. -Lemma protocol_transition_project_active +Lemma input_valid_transition_project_active {message} `{EqDecision V} {IM: V -> VLSM message} {constraint} (X := composite_vlsm IM constraint) l s im s' om: - protocol_transition X l (s,im) (s',om) -> - protocol_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) - (s (projT1 l), im) (s' (projT1 l), om). + input_valid_transition X l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) + (s (projT1 l), im) (s' (projT1 l), om). Proof. intro Hptrans. - apply preloaded_weaken_protocol_transition in Hptrans. + apply preloaded_weaken_input_valid_transition in Hptrans. revert Hptrans. - apply protocol_transition_preloaded_project_active. + apply input_valid_transition_preloaded_project_active. Qed. -Lemma protocol_transition_preloaded_project_any {V} (i:V) +Lemma input_valid_transition_preloaded_project_any {V} (i:V) {message} `{EqDecision V} {IM: V -> VLSM message} {constraint} (X := composite_vlsm IM constraint) (l:vlabel X) s im s' om: - protocol_transition (pre_loaded_with_all_messages_vlsm X) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm X) l (s,im) (s',om) -> (s i = s' i \/ exists li, (l = existT i li) /\ - protocol_transition (pre_loaded_with_all_messages_vlsm (IM i)) - li - (s i,im) (s' i,om)). + input_valid_transition (pre_loaded_with_all_messages_vlsm (IM i)) + li + (s i,im) (s' i,om)). Proof. intro Hptrans. destruct l as [j lj]. @@ -884,7 +884,7 @@ Proof. exists lj. split;[reflexivity|]. revert Hptrans. - apply protocol_transition_preloaded_project_active. + apply input_valid_transition_preloaded_project_active. - left. destruct Hptrans as [Hpvalid Htrans]. cbn in Htrans. @@ -894,21 +894,21 @@ Proof. reflexivity. Qed. -Lemma protocol_transition_project_any {V} (i:V) +Lemma input_valid_transition_project_any {V} (i:V) {message} `{EqDecision V} {IM: V -> VLSM message} {constraint} (X := composite_vlsm IM constraint) (l:vlabel X) s im s' om: - protocol_transition X l (s,im) (s',om) -> + input_valid_transition X l (s,im) (s',om) -> (s i = s' i \/ exists li, (l = existT i li) /\ - protocol_transition (pre_loaded_with_all_messages_vlsm (IM i)) - li - (s i,im) (s' i,om)). + input_valid_transition (pre_loaded_with_all_messages_vlsm (IM i)) + li + (s i,im) (s' i,om)). Proof. intro Hproto. - apply preloaded_weaken_protocol_transition in Hproto. + apply preloaded_weaken_input_valid_transition in Hproto. revert Hproto. - apply protocol_transition_preloaded_project_any. + apply input_valid_transition_preloaded_project_any. Qed. (** If a message can be emitted by a composition, then it can be emited by one of the @@ -927,7 +927,7 @@ Proof. apply can_emit_iff. exists (s2 (projT1 l)). exists (s1 (projT1 l), oim), (projT2 l). - revert Ht. apply protocol_transition_preloaded_project_active. + revert Ht. apply input_valid_transition_preloaded_project_active. Qed. Section binary_free_composition. @@ -1020,8 +1020,8 @@ Section composite_plan_properties. to a [vstate Free] <>, knowing its effects on a different [vstate Free] <> which shares some relevant features with <>. *) - (* A transition on component <> is protocol from <> if it is - protocol from <> and their <>'th components are equal *) + (* A transition on component <> is [input_valid] from <> if it is + [input_valid] from <> and their <>'th components are equal. *) Lemma relevant_component_transition (s s' : vstate Free) @@ -1029,11 +1029,11 @@ Section composite_plan_properties. (input : option message) (i := projT1 l) (Heq : (s i) = (s' i)) - (Hprs : protocol_state_prop Free s') - (Hpr : protocol_valid Free l (s, input)) : - protocol_valid Free l (s', input). + (Hprs : valid_state_prop Free s') + (Hiv : input_valid Free l (s, input)) : + input_valid Free l (s', input). Proof. - unfold protocol_valid in *. + unfold input_valid in *. split; [intuition|intuition|..]. unfold valid in *; simpl in *. unfold constrained_composite_valid in *. @@ -1055,7 +1055,7 @@ Section composite_plan_properties. (input : option message) (i := projT1 l) (Heq : (s i) = (s' i)) - (Hprs : protocol_state_prop Free s') : + (Hprs : valid_state_prop Free s') : let (dest, output) := vtransition Free l (s, input) in let (dest', output') := vtransition Free l (s', input) in output = output' /\ (dest i) = (dest' i). @@ -1075,18 +1075,18 @@ Section composite_plan_properties. Lemma relevant_components_one (s s' : vstate Free) - (Hprs' : protocol_state_prop Free s') + (Hprs' : valid_state_prop Free s') (ai : vplan_item Free) (i := projT1 (label_a ai)) (Heq : (s i) = (s' i)) - (Hpr : finite_protocol_plan_from Free s [ai]) : + (Hpr : finite_valid_plan_from Free s [ai]) : let res' := snd (apply_plan Free s' [ai]) in let res := snd (apply_plan Free s [ai]) in - finite_protocol_plan_from Free s' [ai] /\ + finite_valid_plan_from Free s' [ai] /\ (res' i) = res i. Proof. simpl. - unfold finite_protocol_plan_from in *. + unfold finite_valid_plan_from in *. unfold apply_plan, _apply_plan in *. destruct ai; simpl in *. match goal with @@ -1099,16 +1099,16 @@ Section composite_plan_properties. end. inversion Hpr; subst. split. - - assert (Ht' : protocol_transition Free label_a (s', input_a) (s0, o)). { - unfold protocol_transition in *. + - assert (Ht' : input_valid_transition Free label_a (s', input_a) (s0, o)). { + unfold input_valid_transition in *. destruct Ht as [Hpr_valid Htrans]. apply relevant_component_transition with (s' := s') in Hpr_valid. all : intuition. } - apply finite_ptrace_extend. - apply finite_ptrace_empty. - apply protocol_transition_destination in Ht'; assumption. + apply finite_valid_trace_from_extend. + apply finite_valid_trace_from_empty. + apply input_valid_transition_destination in Ht'; assumption. assumption. - simpl. specialize (relevant_component_transition2 s s' label_a input_a) as Hrel. @@ -1199,25 +1199,25 @@ Section composite_plan_properties. Lemma relevant_components (s s' : vstate Free) - (Hprs' : protocol_state_prop Free s') + (Hprs' : valid_state_prop Free s') (a : plan Free) (a_indices := List.map (@projT1 _ _) (List.map (@label_a _ _) a)) (li : list index) (Heq : forall (i : index), i ∈ li -> (s' i) = (s i)) (Hincl : a_indices ⊆ li) - (Hpr : finite_protocol_plan_from Free s a) : + (Hpr : finite_valid_plan_from Free s a) : let res' := snd (apply_plan Free s' a) in let res := snd (apply_plan Free s a) in - finite_protocol_plan_from Free s' a /\ + finite_valid_plan_from Free s' a /\ (forall (i : index), i ∈ li -> (res' i) = res i). Proof. induction a using rev_ind. - split. - apply finite_protocol_plan_empty. + apply finite_valid_plan_empty. assumption. simpl. assumption. - simpl in *. - apply finite_protocol_plan_from_app_iff in Hpr. + apply finite_valid_plan_from_app_iff in Hpr. destruct Hpr as [Hrem Hsingle]. spec IHa. { @@ -1239,7 +1239,7 @@ Section composite_plan_properties. specialize (relevant_components_one (snd (apply_plan Free s a)) (snd (apply_plan Free s' a))) as Hrel. spec Hrel. { - apply apply_plan_last_protocol. + apply apply_plan_last_valid. all : intuition. } @@ -1259,7 +1259,7 @@ Section composite_plan_properties. specialize (Hrel Hsingle). destruct Hrel as [Hrelpr Hrelind]. split. - + apply finite_protocol_plan_from_app_iff. + + apply finite_valid_plan_from_app_iff. split; intuition. + intros i Hi. specialize (IHaind i Hi). @@ -1345,11 +1345,11 @@ Proof. elim (empty_composition_no_label l). Qed. -Lemma empty_composition_no_protocol_message - : forall m, ~ protocol_message_prop X m. +Lemma empty_composition_no_valid_message + : forall m, ~ valid_message_prop X m. Proof. intros m Hm. - apply can_emit_protocol_iff in Hm as [Hinit | Hemit]. + apply emitted_messages_are_valid_iff in Hm as [Hinit | Hemit]. - elim (empty_composition_no_initial_message _ Hinit). - elim (empty_composition_no_emit _ Hemit). Qed. @@ -1378,7 +1378,7 @@ If two indexed sets of VLSMs are extensionally-equal, then we can establish a [VLSM_full_projection] between their compositions with subsumable constraints (and pre-loaded with the same set of messages). *) -Section same_IM_full_projection. +Section sec_same_IM_full_projection. Context {message : Type} @@ -1403,7 +1403,7 @@ Context (constraint1 : composite_label IM1 -> composite_state IM1 * option message -> Prop) (constraint2 : composite_label IM2 -> composite_state IM2 * option message -> Prop) (constraint_projection - : forall s1, protocol_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM1)) s1 -> + : forall s1, valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM1)) s1 -> forall l1 om, constraint1 l1 (s1,om) -> constraint2 (same_IM_label_rew l1) (same_IM_state_rew s1, om)) (seed : message -> Prop) @@ -1419,7 +1419,7 @@ Proof. apply basic_VLSM_full_projection; intros l **. - destruct Hv as [Hs [Hom [Hv Hc]]]. apply constraint_projection in Hc; cycle 1. - + apply (VLSM_incl_protocol_state + + apply (VLSM_incl_valid_state (composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages IM1 constraint1 seed)). assumption. + split; [|assumption]. @@ -1436,7 +1436,7 @@ Proof. + subst. rewrite !state_update_eq. reflexivity. + rewrite !state_update_neq by congruence. reflexivity. - intros i. apply same_VLSM_initial_state_preservation, H. - - apply initial_message_is_protocol. + - apply initial_message_is_valid. destruct HmX as [[i [[im Him] Hi]] | Hseed]; [| right; assumption]. simpl in Hi. subst im. cbn. unfold composite_initial_message_prop. @@ -1458,19 +1458,19 @@ Proof. constructor. intros s1 tr1 Htr1. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (free_composite_vlsm IM1)) as Heq1. - apply (VLSM_eq_finite_protocol_trace Heq1) in Htr1. + apply (VLSM_eq_finite_valid_trace Heq1) in Htr1. clear Heq1. specialize (same_IM_full_projection (free_constraint IM1) (free_constraint IM2)) as Hproj. spec Hproj. { intros. exact I. } specialize (Hproj (fun _ => True)). - apply (VLSM_full_projection_finite_protocol_trace Hproj) in Htr1. + apply (VLSM_full_projection_finite_valid_trace Hproj) in Htr1. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (free_composite_vlsm IM2)) as Heq2. - apply (VLSM_eq_finite_protocol_trace Heq2). + apply (VLSM_eq_finite_valid_trace Heq2). assumption. Qed. -End same_IM_full_projection. +End sec_same_IM_full_projection. Arguments same_IM_label_rew {_ _ _ _} _ _ : assert. Arguments same_IM_state_rew {_ _ _ _} _ _ _ : assert. diff --git a/theories/VLSM/Core/Decisions.v b/theories/VLSM/Core/Decisions.v index 4bd0e9ea..df8c75b3 100644 --- a/theories/VLSM/Core/Decisions.v +++ b/theories/VLSM/Core/Decisions.v @@ -1,7 +1,7 @@ From stdpp Require Import prelude. From Coq Require Import FinFun Streams Logic.Epsilon Arith.Compare_dec Lia. From VLSM Require Import Lib.Preamble Lib.ListExtras Lib.ListSetExtras Lib.RealsExtras. -From VLSM Require Import Core.VLSM Core.Composition Core.ProjectionTraces. +From VLSM Require Import Core.VLSM Core.Composition Core.VLSMProjections Core.ProjectionTraces. (** * VLSM Decisions on Consensus Values *) @@ -28,7 +28,7 @@ Section CommuteSingleton. (* Definition of finality per document. *) Definition final_original : vdecision V -> Prop := - fun (D : vdecision V) => forall (tr : protocol_trace V), + fun (D : vdecision V) => forall (tr : valid_trace V), forall (n1 n2 : nat) (s1 s2 : state) (c1 c2 : C), (trace_nth (proj1_sig tr) n1 = Some s1) -> (trace_nth (proj1_sig tr) n2 = Some s2) -> @@ -44,16 +44,16 @@ Section CommuteSingleton. (D s2 = (Some c2)) -> c1 = c2. - (* 3.3.1 Initial protocol state bivalence *) + (* 3.3.1 Initial state bivalence *) Definition bivalent : vdecision V -> Prop := fun (D : vdecision V) => (* All initial states decide on None *) (forall (s0 : state), vinitial_state_prop V s0 -> D s0 = None) /\ - (* Every protocol trace (already beginning from an initial state) contains a state deciding on each consensus value *) + (* Every valid trace (already beginning from an initial state) contains a state deciding on each consensus value *) (forall (c : C) , - exists (tr : protocol_trace V) (s : state) (n : nat), + exists (tr : valid_trace V) (s : state) (n : nat), (trace_nth (proj1_sig tr) n) = Some s /\ D s = (Some c)). (* 3.3.2 No stuck states *) @@ -61,7 +61,7 @@ Section CommuteSingleton. Definition stuck_free : vdecision V -> Prop := fun (D : vdecision V) => (forall (s : state), - exists (tr : protocol_trace V) + exists (tr : valid_trace V) (decided_state : state) (n_s n_decided : nat) (c : C), @@ -102,7 +102,7 @@ Section CommuteIndexed. *) Definition consistent_original := - forall (tr : protocol_trace X), + forall (tr : valid_trace X), forall (n1 n2 : nat), forall (j k : index), forall (s1 s2 : vstate X), @@ -183,11 +183,13 @@ Section CommuteIndexed. Lemma final_and_consistent_implies_final (Hcons : final_and_consistent) (i : index) - (Hfr : finite_projection_friendly IM constraint i) + (Hfr : projection_friendly_prop (component_projection IM constraint i)) : final (composite_vlsm_constrained_projection IM constraint i) (ID i). Proof. intros s1 s2 c1 c2 Hfuturesi HD1 HD2. - specialize (projection_friendly_in_futures IM constraint i Hfr s1 s2 Hfuturesi) + specialize + (projection_friendly_in_futures (component_projection IM constraint i) + Hfr s1 s2 Hfuturesi) ; intros [sX1 [sX2 [Hs1 [Hs2 HfuturesX]]]]. subst. apply (Hcons sX1 sX2 HfuturesX i i c1 c2 HD1 HD2). @@ -242,7 +244,7 @@ Section Estimators. Lemma estimator_only_has_decision (Hde : decision_estimator_property) - (s : protocol_state X) + (s : valid_state X) (c c_other : C) (Hc : D (proj1_sig s) = (Some c)) (Hc_other : estimates (proj1_sig s) c_other) @@ -257,7 +259,7 @@ Section Estimators. Lemma estimator_surely_has_decision (Hde : decision_estimator_property) - (s : protocol_state X) + (s : valid_state X) (c : C) (Hc : D (proj1_sig s) = (Some c)) : estimates (proj1_sig s) c. @@ -294,7 +296,7 @@ Section Estimators. intros. unfold final. intros. - specialize (in_futures_protocol_snd X s1 s2 H0); intro Hps2. + specialize (in_futures_valid_snd X s1 s2 H0); intro Hps2. apply estimator_only_has_decision with (s := (exist _ s2 Hps2)). assumption. assumption. diff --git a/theories/VLSM/Core/Equivocation.v b/theories/VLSM/Core/Equivocation.v index 1007bc3a..b1b20fd0 100644 --- a/theories/VLSM/Core/Equivocation.v +++ b/theories/VLSM/Core/Equivocation.v @@ -140,7 +140,7 @@ Qed. [state_message_oracle]. Such an oracle can, given a state and a message, decide whether the message has been sent (or received) in the history leading to the current state. Formally, we say that a [message] [has_been_sent] - if we're in [state] iff every protocol trace which produces contains + if we're in [state] iff every valid trace which produces contains as a sent message somewhere along the way. The existence of such oracles, which practically imply endowing states with history, @@ -267,7 +267,7 @@ Section Simple. forall (start : state) (tr : list transition_item) - (Htr : finite_protocol_trace_init_to X start s tr), + (Htr : finite_valid_trace_init_to X start s tr), trace_has_message message_selector m tr. Definition selected_message_exists_in_all_preloaded_traces @@ -283,7 +283,7 @@ Section Simple. exists (start : state) (tr : list transition_item) - (Htr : finite_protocol_trace_init_to X start s tr), + (Htr : finite_valid_trace_init_to X start s tr), trace_has_message message_selector m tr. Definition selected_message_exists_in_some_preloaded_traces: forall @@ -303,7 +303,7 @@ Section Simple. forall (start : state) (tr : list transition_item) - (Htr : finite_protocol_trace_init_to X start s tr), + (Htr : finite_valid_trace_init_to X start s tr), ~trace_has_message message_selector m tr. Definition selected_message_exists_in_no_preloaded_trace := @@ -345,19 +345,19 @@ Section Simple. (m : message) (start : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from_to X start s tr) + (Htr : finite_valid_trace_from_to X start s tr) (Hsome : trace_has_message message_selector m tr) : specialized_selected_message_exists_in_some_traces X message_selector s m. Proof. - assert (protocol_state_prop X start) as Hstart - by (apply ptrace_first_pstate in Htr; assumption). - apply protocol_state_has_trace in Hstart. + assert (valid_state_prop X start) as Hstart + by (apply valid_trace_first_pstate in Htr; assumption). + apply valid_state_has_trace in Hstart. destruct Hstart as [is [tr' Htr']]. - assert (finite_protocol_trace_init_to X is s (tr'++tr)). + assert (finite_valid_trace_init_to X is s (tr'++tr)). { destruct Htr'. split; - [apply finite_protocol_trace_from_to_app with start|]; + [apply finite_valid_trace_from_to_app with start|]; assumption. } exists _, _, H. @@ -382,9 +382,9 @@ Section Simple. : ~ selected_message_exists_in_all_preloaded_traces message_selector s m. Proof. intro Hselected. - assert (Hps : protocol_state_prop pre_vlsm s) - by (apply initial_is_protocol;assumption). - assert (Htr : finite_protocol_trace_init_to pre_vlsm s s []). + assert (Hps : valid_state_prop pre_vlsm s) + by (apply initial_state_is_valid;assumption). + assert (Htr : finite_valid_trace_init_to pre_vlsm s s []). { split; try assumption. constructor. assumption. } specialize (Hselected s [] Htr). unfold trace_has_message in Hselected. @@ -392,11 +392,11 @@ Section Simple. assumption. Qed. -(** Checks if all [protocol_trace]s leading to a certain state contain a certain message. +(** Checks if all [valid_trace]s leading to a certain state contain a certain message. The [message_selector] argument specifices whether we're looking for received or sent messages. - Notably, the [protocol_trace]s over which we are iterating belong to the preloaded + Notably, the [valid_trace]s over which we are iterating belong to the preloaded version of the target VLSM. This is because we want VLSMs to have oracles which are valid irrespective of the composition they take part in. As we know, the behaviour preloaded VLSMs includes behaviours of its projections in any @@ -450,7 +450,7 @@ Section Simple. proper_sent: forall (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), (has_been_sent_prop has_been_sent s m); @@ -459,22 +459,22 @@ Section Simple. proper_not_sent: forall (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), has_not_been_sent_prop has_not_been_sent s m; }. (** Reverse implication for 'selected_messages_consistency_prop' always holds. *) - Lemma consistency_from_protocol_proj2 + Lemma consistency_from_valid_state_proj2 (s : state) - (Hs: protocol_state_prop pre_vlsm s) + (Hs: valid_state_prop pre_vlsm s) (m : message) (selector : message -> transition_item -> Prop) (Hall : selected_message_exists_in_all_preloaded_traces selector s m) : selected_message_exists_in_some_preloaded_traces selector s m. Proof. - apply protocol_state_has_trace in Hs. + apply valid_state_has_trace in Hs. destruct Hs as [is [tr Htr]]. exists _, _, Htr. apply (Hall _ _ Htr). @@ -483,7 +483,7 @@ Section Simple. Lemma has_been_sent_consistency {Hbs : HasBeenSentCapability} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : selected_messages_consistency_prop (field_selector output) s m. Proof. @@ -496,29 +496,29 @@ Section Simple. destruct Hsome as [is [tr [Htr Hmsg]]]. elim (Hsm _ _ Htr). assumption. - - apply consistency_from_protocol_proj2. + - apply consistency_from_valid_state_proj2. assumption. Qed. - Lemma protocol_generated_has_been_sent + Lemma can_produce_has_been_sent {Hbs : HasBeenSentCapability} (s : state) (m : message) - (Hsm : protocol_generated_prop pre_vlsm s m) + (Hsm : can_produce pre_vlsm s m) : has_been_sent s m. Proof. - assert (protocol_state_prop pre_vlsm s). - { apply protocol_generated_prop_protocol in Hsm. + assert (valid_state_prop pre_vlsm s). + { apply can_produce_valid in Hsm. eexists; exact Hsm. } apply proper_sent; [assumption|]. apply has_been_sent_consistency; [assumption|]. - apply non_empty_protocol_trace_from_protocol_generated_prop in Hsm. + apply non_empty_valid_trace_from_can_produce in Hsm. destruct Hsm as [is [tr [lst_tr [Htr [Hlst [Hs Hm]]]]]]. destruct_list_last tr tr' _lst_tr Heqtr; [inversion Hlst|]. rewrite last_error_is_last in Hlst. inversion Hlst; subst _lst_tr; clear Hlst. - apply ptrace_add_default_last in Htr. + apply valid_trace_add_default_last in Htr. rewrite finite_trace_last_is_last, Hs in Htr. eexists _, _, Htr. apply Exists_app. right. left. assumption. @@ -530,14 +530,14 @@ Section Simple. Lemma specialized_proper_sent {Hbs : HasBeenSentCapability} (s : state) - (Hs : protocol_state_prop vlsm s) + (Hs : valid_state_prop vlsm s) (m : message) (Hsome : specialized_selected_message_exists_in_some_traces vlsm (field_selector output) s m) : has_been_sent s m. Proof. destruct Hs as [_om Hs]. - assert (Hpres : protocol_state_prop pre_vlsm s). - { exists _om. apply (pre_loaded_with_all_messages_protocol_prop vlsm). assumption. } + assert (Hpres : valid_state_prop pre_vlsm s). + { exists _om. apply (pre_loaded_with_all_messages_valid_state_message_preservation vlsm). assumption. } apply proper_sent; [assumption|]. specialize (has_been_sent_consistency s Hpres m) as Hcons. apply Hcons. @@ -547,7 +547,7 @@ Section Simple. revert Htr. unfold pre_vlsm;clear. destruct vlsm as (T,(S,M)). - apply VLSM_incl_finite_protocol_trace_init_to. + apply VLSM_incl_finite_valid_trace_init_to. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. Qed. @@ -557,14 +557,14 @@ Section Simple. Lemma specialized_proper_sent_rev {Hbs : HasBeenSentCapability} (s : state) - (Hs : protocol_state_prop vlsm s) + (Hs : valid_state_prop vlsm s) (m : message) (Hsm : has_been_sent s m) : specialized_selected_message_exists_in_all_traces vlsm (field_selector output) s m. Proof. destruct Hs as [_om Hs]. - assert (Hpres : protocol_state_prop pre_vlsm s). - { exists _om. apply (pre_loaded_with_all_messages_protocol_prop vlsm). assumption. } + assert (Hpres : valid_state_prop pre_vlsm s). + { exists _om. apply (pre_loaded_with_all_messages_valid_state_message_preservation vlsm). assumption. } apply proper_sent in Hsm; [|assumption]. intros is tr Htr. specialize (Hsm is tr). @@ -572,7 +572,7 @@ Section Simple. revert Htr. unfold pre_vlsm;clear. destruct vlsm as (T,(S,M)). - apply VLSM_incl_finite_protocol_trace_init_to. + apply VLSM_incl_finite_valid_trace_init_to. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. Qed. @@ -603,7 +603,7 @@ Section Simple. proper_received: forall (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), (has_been_received_prop has_been_received s m); @@ -612,7 +612,7 @@ Section Simple. proper_not_received: forall (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), has_not_been_received_prop has_not_been_received s m; }. @@ -620,7 +620,7 @@ Section Simple. Lemma has_been_received_consistency {Hbs : HasBeenReceivedCapability} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : selected_messages_consistency_prop (field_selector input) s m. Proof. @@ -632,7 +632,7 @@ Section Simple. destruct Hsome as [is [tr [Htr Hsome]]]. elim (Hsm _ _ Htr). assumption. - - apply consistency_from_protocol_proj2. + - apply consistency_from_valid_state_proj2. assumption. Qed. @@ -672,7 +672,7 @@ Section Simple. Lemma sent_messages_proper (Hhbs : HasBeenSentCapability) (s : vstate vlsm) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_been_sent s m <-> exists (m' : sent_messages s), proj1_sig m' = m. Proof. @@ -693,7 +693,7 @@ Section Simple. Lemma received_messages_proper (Hhbs : HasBeenReceivedCapability) (s : vstate vlsm) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_been_received s m <-> exists (m' : received_messages s), proj1_sig m' = m. Proof. @@ -709,13 +709,13 @@ Section Simple. sent_messages_fn : vstate vlsm -> list message; sent_messages_full : - forall (s : vstate vlsm) (Hs : protocol_state_prop pre_vlsm s) (m : message), + forall (s : vstate vlsm) (Hs : valid_state_prop pre_vlsm s) (m : message), m ∈ (sent_messages_fn s) <-> exists (sm : sent_messages s), proj1_sig sm = m; sent_messages_consistency : forall (s : vstate vlsm) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), selected_messages_consistency_prop (field_selector output) s m }. @@ -725,8 +725,8 @@ Section Simple. (s : vinitial_state vlsm) : sent_messages_fn (proj1_sig s) = []. Proof. - assert (Hps : protocol_state_prop pre_vlsm (proj1_sig s)) - by (apply initial_is_protocol; apply proj2_sig). + assert (Hps : valid_state_prop pre_vlsm (proj1_sig s)) + by (apply initial_state_is_valid; apply proj2_sig). destruct s as [s Hs]. simpl in *. destruct (sent_messages_fn s) as [|m l] eqn:Hsm; try reflexivity. specialize (sent_messages_full s Hps m) as Hl. apply proj1 in Hl. @@ -755,7 +755,7 @@ Section Simple. Lemma ComputableSentMessages_has_been_sent_proper {Hsm : ComputableSentMessages} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_been_sent_prop ComputableSentMessages_has_been_sent s m. Proof. @@ -784,7 +784,7 @@ Section Simple. Lemma ComputableSentMessages_has_not_been_sent_proper {Hsm : ComputableSentMessages} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_not_been_sent_prop ComputableSentMessages_has_not_been_sent s m. Proof. @@ -823,13 +823,13 @@ Section Simple. received_messages_fn : vstate vlsm -> list message; received_messages_full : - forall (s : vstate vlsm) (Hs : protocol_state_prop pre_vlsm s) (m : message), + forall (s : vstate vlsm) (Hs : valid_state_prop pre_vlsm s) (m : message), m ∈ (received_messages_fn s) <-> exists (sm : received_messages s), proj1_sig sm = m; received_messages_consistency : forall (s : vstate vlsm) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message), selected_messages_consistency_prop (field_selector input) s m }. @@ -839,8 +839,8 @@ Section Simple. (s : vinitial_state vlsm) : received_messages_fn (proj1_sig s) = []. Proof. - assert (Hps : protocol_state_prop pre_vlsm (proj1_sig s)) - by (apply initial_is_protocol;apply proj2_sig). + assert (Hps : valid_state_prop pre_vlsm (proj1_sig s)) + by (apply initial_state_is_valid;apply proj2_sig). destruct s as [s Hs]. simpl in *. destruct (received_messages_fn s) as [|m l] eqn:Hrcv; try reflexivity. specialize (received_messages_full s Hps m) as Hl. apply proj1 in Hl. @@ -864,12 +864,12 @@ Section Simple. {Hsm : ComputableReceivedMessages} {eq_message : EqDecision message} : RelDecision ComputableReceivedMessages_has_been_received - := fun s m => decide_rel _ _ (received_messages_fn s). + := fun s m => decide_rel _ _ (received_messages_fn s). Lemma ComputableReceivedMessages_has_been_received_proper {Hsm : ComputableReceivedMessages} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_been_received_prop ComputableReceivedMessages_has_been_received s m. Proof. @@ -896,7 +896,7 @@ Section Simple. Lemma ComputableReceivedMessages_has_not_been_received_proper {Hsm : ComputableReceivedMessages} (s : state) - (Hs : protocol_state_prop pre_vlsm s) + (Hs : valid_state_prop pre_vlsm s) (m : message) : has_not_been_received_prop ComputableReceivedMessages_has_not_been_received s m. Proof. @@ -934,7 +934,7 @@ End Simple. one is that the oracle cannot only for any initial state, the other is that the oracle judgement is appropriately related for the starting and [destination] states of - any [protocol_transition]. + any [input_valid_transition]. These conditions are defined in the record [oracle_stepwise_props] *) @@ -948,7 +948,7 @@ Record oracle_stepwise_props forall m, ~oracle s m; oracle_step_update: forall l s im s' om, - protocol_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> forall msg, oracle s' msg <-> (message_selector msg {|l:=l; input:=im; destination:=s'; output:=om|} \/ oracle s msg) @@ -962,7 +962,7 @@ Lemma oracle_partial_trace_update [oracle: state_message_oracle vlsm] (Horacle: oracle_stepwise_props selector oracle) s0 s tr - (Htr: finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): + (Htr: finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): forall m, oracle s m <-> (trace_has_message selector m tr \/ oracle s0 m). @@ -985,7 +985,7 @@ Lemma oracle_initial_trace_update [oracle : state_message_oracle vlsm] (Horacle: oracle_stepwise_props selector oracle) s0 s tr - (Htr: finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): + (Htr: finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): forall m, oracle s m <-> trace_has_message selector m tr. Proof. @@ -998,13 +998,13 @@ Qed. (** Proving the trace properties from the stepwise properties begins with a lemma using induction along a trace to - prove that given a [finite_protocol_trace] to a state, + prove that given a [finite_valid_trace] to a state, the oracle holds at that state for some message iff a satsifying transition item exists in the trace. The theorems for [all_traces_have_message_prop] and [no_traces_have_message_prop] are mostly rearraning - quantifiers to use this lemma, also using [protocol_state_prop] + quantifiers to use this lemma, also using [valid_state_prop] to choose a trace to the state for the directions where one is not given. *) @@ -1017,9 +1017,9 @@ Section TraceFromStepwise. (oracle_props : oracle_stepwise_props selector oracle) . - Local Lemma H_protocol_trace_prop + Local Lemma H_valid_trace_prop [s0 s tr] - (Htr: finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): + (Htr: finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) s0 s tr): forall m, oracle s m <-> trace_has_message selector m tr. Proof. @@ -1033,7 +1033,7 @@ Section TraceFromStepwise. Lemma prove_all_have_message_from_stepwise: forall (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) (m : message), (all_traces_have_message_prop vlsm selector oracle s m). Proof. @@ -1041,42 +1041,42 @@ Section TraceFromStepwise. unfold all_traces_have_message_prop. split. - intros Hsent s0 tr Htr. - apply (H_protocol_trace_prop Htr). + apply (H_valid_trace_prop Htr). assumption. - intro H_all_traces. - apply protocol_state_has_trace in Hproto. + apply valid_state_has_trace in Hproto. destruct Hproto as [s0 [tr Htr]]. - apply (H_protocol_trace_prop Htr). + apply (H_valid_trace_prop Htr). specialize (H_all_traces s0 tr Htr). assumption. Qed. Lemma prove_none_have_message_from_stepwise: forall (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) (m : message), no_traces_have_message_prop vlsm selector (fun s m => ~oracle s m) s m. Proof. intros s Hproto m. - pose proof (H_protocol_trace_prop). + pose proof (H_valid_trace_prop). split. - intros H_not_sent start tr Htr. contradict H_not_sent. - apply (H_protocol_trace_prop Htr). + apply (H_valid_trace_prop Htr). assumption. - intros H_no_traces. - apply protocol_state_has_trace in Hproto. + apply valid_state_has_trace in Hproto. destruct Hproto as [s0 [tr Htr]]. specialize (H_no_traces s0 tr Htr). contradict H_no_traces. - apply (H_protocol_trace_prop Htr). + apply (H_valid_trace_prop Htr). assumption. Qed. Lemma selected_messages_consistency_prop_from_stepwise (oracle_dec: RelDecision oracle) (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) (m : message) : selected_messages_consistency_prop vlsm selector s m. Proof. @@ -1089,7 +1089,7 @@ Section TraceFromStepwise. destruct Hsome as [is [tr [Htr Hmsg]]]. elim (Hsm _ _ Htr). assumption. - - apply consistency_from_protocol_proj2. + - apply consistency_from_valid_state_proj2. assumption. Qed. @@ -1109,7 +1109,7 @@ End TraceFromStepwise. The stepwise properties are proven from the trace properties by considering the empty trace to prove the [oracle_no_inits] property, and by considering a trace that ends with the given - [protocol_transition] to prove the [oracle_step_update] property. + [input_valid_transition] to prove the [oracle_step_update] property. *) Section StepwiseFromTrace. Context @@ -1119,10 +1119,10 @@ Section StepwiseFromTrace. (oracle: state_message_oracle vlsm) (oracle_dec: RelDecision oracle) (Horacle_all_have: - forall s (Hs: protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) m, + forall s (Hs: valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) m, all_traces_have_message_prop vlsm selector oracle s m) (Hnot_oracle_none_have: - forall s (Hs: protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) m, + forall s (Hs: valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) m, no_traces_have_message_prop vlsm selector (fun m s => ~oracle m s) s m). Lemma oracle_no_inits_from_trace: @@ -1130,8 +1130,8 @@ Section StepwiseFromTrace. forall m, ~oracle s m. Proof. intros s Hinit m Horacle. - assert (Hproto : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) - by (apply initial_is_protocol;assumption). + assert (Hproto : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + by (apply initial_state_is_valid;assumption). apply Horacle_all_have in Horacle;[|assumption]. specialize (Horacle s nil). eapply Exists_nil;apply Horacle;clear Horacle. @@ -1140,14 +1140,14 @@ Section StepwiseFromTrace. Lemma examine_one_trace: forall is s tr, - finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> + finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> forall m, oracle s m <-> trace_has_message selector m tr. Proof. intros is s tr Htr m. - assert (protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) - by (apply ptrace_last_pstate in Htr;assumption). + assert (valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + by (apply valid_trace_last_pstate in Htr;assumption). split. - intros Horacle. apply Horacle_all_have in Horacle;[|assumption]. @@ -1165,7 +1165,7 @@ Section StepwiseFromTrace. Lemma oracle_step_property_from_trace: forall l s im s' om, - protocol_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> forall msg, oracle s' msg <-> (selector msg {| l:=l; input:=im; destination:=s'; output:=om |} \/ oracle s msg). @@ -1175,7 +1175,7 @@ Section StepwiseFromTrace. pose proof Htrans' as [[Hproto_s [Hproto_m Hvalid]] Htrans]. set (preloaded:= pre_loaded_with_all_messages_vlsm vlsm) in * |- *. - pose proof (protocol_state_has_trace _ _ Hproto_s) + pose proof (valid_state_has_trace _ _ Hproto_s) as [is [tr [Htr Hinit]]]. pose proof (Htr' := extend_right_finite_trace_from_to _ Htr Htrans'). @@ -1242,7 +1242,7 @@ Defined. Lemma has_been_sent_step_update `{Hhbs: HasBeenSentCapability message vlsm}: forall [l s im s' om], - protocol_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> forall m, has_been_sent vlsm s' m <-> (om = Some m \/ has_been_sent vlsm s m). Proof. @@ -1252,7 +1252,7 @@ Qed. Lemma has_been_sent_examine_one_trace `(Hhbs: HasBeenSentCapability message vlsm): forall is s tr, - finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> + finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> forall m, has_been_sent vlsm s m <-> trace_has_message (field_selector output) m tr. @@ -1295,7 +1295,7 @@ Defined. Lemma has_been_received_step_update `{Hhbs: HasBeenReceivedCapability message vlsm}: forall [l s im s' om], - protocol_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s,im) (s',om) -> forall m, has_been_received vlsm s' m <-> (im = Some m \/ has_been_received vlsm s m). Proof. @@ -1305,7 +1305,7 @@ Qed. Lemma has_been_received_examine_one_trace `(Hhbr: HasBeenReceivedCapability message vlsm): forall is s tr, - finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> + finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> forall m, has_been_received vlsm s m <-> trace_has_message (field_selector input) m tr. @@ -1317,14 +1317,14 @@ Qed. Lemma trace_to_initial_state_has_no_inputs `{Hbr: HasBeenReceivedCapability message vlsm} is s tr - (Htr : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr) + (Htr : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr) (Hs : vinitial_state_prop vlsm s) : forall item, In item tr -> input item = None. Proof. intros item Hitem. destruct (input item) as [m|] eqn:Heqm; [|reflexivity]. elim (selected_message_exists_in_all_traces_initial_state _ _ Hs (field_selector input) m). - apply has_been_received_consistency; [assumption|apply initial_is_protocol; assumption|]. + apply has_been_received_consistency; [assumption|apply initial_state_is_valid; assumption|]. eexists _,_, Htr. apply Exists_exists. exists item. split; [|assumption]. apply elem_of_list_In. assumption. @@ -1354,7 +1354,7 @@ Definition has_been_observed_no_inits `{Hhbo: HasBeenObservedCapability message Definition has_been_observed_step_update `{Hhbo: HasBeenObservedCapability message vlsm} : forall l s im s' om, - protocol_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s, im) (s', om) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm vlsm) l (s, im) (s', om) -> forall msg, has_been_observed vlsm s' msg <-> ((im = Some msg \/ om = Some msg) \/ has_been_observed vlsm s msg) @@ -1362,7 +1362,7 @@ Definition has_been_observed_step_update `{Hhbo: HasBeenObservedCapability messa Lemma proper_observed `(Hhbo: HasBeenObservedCapability message vlsm): forall (s:state), - protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s -> + valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s -> forall m, all_traces_have_message_prop vlsm item_sends_or_receives (has_been_observed vlsm) s m. Proof. @@ -1374,7 +1374,7 @@ Qed. Lemma proper_not_observed `(Hhbo: HasBeenObservedCapability message vlsm): forall (s:state), - protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s -> + valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s -> forall m, no_traces_have_message_prop vlsm item_sends_or_receives (fun s m => ~has_been_observed vlsm s m) s m. @@ -1388,7 +1388,7 @@ Qed. Lemma has_been_observed_examine_one_trace `(Hhbo: HasBeenObservedCapability message vlsm): forall is s tr, - finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> + finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm vlsm) is s tr -> forall m, has_been_observed vlsm s m <-> trace_has_message item_sends_or_receives m tr. @@ -1447,7 +1447,7 @@ Context Lemma has_been_observed_sent_received_iff {Hbo : HasBeenObservedCapability vlsm} (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) (m : message) : has_been_observed vlsm s m <-> has_been_received vlsm s m \/ has_been_sent vlsm s m. Proof. @@ -1456,7 +1456,7 @@ Proof. (has_been_observed vlsm) has_been_observed_stepwise_props _ Hs m) as Hall. split; [intro H | intros [H | H]]. - apply proj1 in Hall. specialize (Hall H). - apply consistency_from_protocol_proj2 in Hall; [|assumption]. + apply consistency_from_valid_state_proj2 in Hall; [|assumption]. destruct Hall as [is [tr [Htr Hexists]]]. apply Exists_or_inv in Hexists. destruct Hexists as [Hsent | Hreceived]. @@ -1502,7 +1502,7 @@ Proof. apply Exists_or. right. assumption. + apply proper_received in H; [|apply Hs]. specialize (H _ _ Htr). apply Exists_or. left. assumption. - - apply consistency_from_protocol_proj2 in H; [|assumption]. + - apply consistency_from_valid_state_proj2 in H; [|assumption]. destruct H as [is [tr [Htr Hexists]]]. apply Exists_or_inv in Hexists. destruct Hexists as [Hsent | Hreceived]. @@ -1544,7 +1544,7 @@ Local Program Instance HasBeenObservedCapability_from_sent_received Lemma has_been_observed_consistency {Hbo : HasBeenObservedCapability vlsm} (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm vlsm) s) (m : message) : selected_messages_consistency_prop vlsm item_sends_or_receives s m. Proof. @@ -1557,26 +1557,26 @@ Local Program Instance HasBeenObservedCapability_from_sent_received destruct Hsome as [is [tr [Htr Hmsg]]]. elim (Hsm _ _ Htr). assumption. - - apply consistency_from_protocol_proj2. + - apply consistency_from_valid_state_proj2. assumption. Qed. End sent_received_observed_capabilities. -Lemma sent_protocol +Lemma sent_valid [message] (X : VLSM message) {Hhbs: HasBeenSentCapability X} (s : state) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (m : message) (Hsent : has_been_sent X s m) : - protocol_message_prop X m. + valid_message_prop X m. Proof. - induction Hs using protocol_state_prop_ind. + induction Hs using valid_state_prop_ind. - contradict Hsent. apply (oracle_no_inits (has_been_sent_stepwise_from_trace Hhbs));assumption. - - apply protocol_transition_out in Ht as Hom'. - apply preloaded_weaken_protocol_transition in Ht. + - apply input_valid_transition_out in Ht as Hom'. + apply preloaded_weaken_input_valid_transition in Ht. apply (oracle_step_update (has_been_sent_stepwise_from_trace Hhbs) _ _ _ _ _ Ht) in Hsent. destruct Hsent. + simpl in H. subst om'. assumption. @@ -1647,12 +1647,12 @@ Section Composite. assert (forall j, s j = s' j \/ j = i). { intro j. - apply (protocol_transition_preloaded_project_any j) in Hproto. + apply (input_valid_transition_preloaded_project_any j) in Hproto. destruct Hproto;[left;assumption|right]. destruct H as [lj [Hlj _]]. congruence. } - apply protocol_transition_preloaded_project_active in Hproto;simpl in Hproto. + apply input_valid_transition_preloaded_project_active in Hproto;simpl in Hproto. apply (oracle_step_update (stepwise_props i)) with (msg:=msg) in Hproto. split. + intros [j Hj]. @@ -1750,7 +1750,7 @@ Section Composite. Lemma composite_proper_sent (s : state) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s) (m : message) : has_been_sent_prop (free_composite_vlsm IM) composite_has_been_sent s m. Proof. @@ -1872,7 +1872,7 @@ Section Composite. to produce that message Weak/strong nontriviality say that each validator should - be designated sender for at least one/all its protocol + be designated sender for at least one/all its valid messages. **) Definition sender_safety_prop : Prop := @@ -1914,6 +1914,8 @@ Section Composite. (** The [channel_authentication_prop]erty requires that any sent message must be originating with its <>. + Note that we don't require that <> is total, but rather that it is + defined for all messages which can be emitted. *) Definition channel_authentication_prop : Prop := forall i m, @@ -1940,15 +1942,15 @@ Section Composite. Definition no_initial_messages_in_IM_prop : Prop := forall i m, ~vinitial_message_prop (IM i) m. - Lemma composite_no_initial_protocol_messages_have_sender + Lemma composite_no_initial_valid_messages_have_sender (can_emit_signed : channel_authentication_prop) (no_initial_messages_in_IM : no_initial_messages_in_IM_prop) (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) - : forall (m : message) (Hm : protocol_message_prop X m), sender m <> None. + : forall (m : message) (Hm : valid_message_prop X m), sender m <> None. Proof. intros m Hm Hsender. - apply can_emit_protocol_iff in Hm. + apply emitted_messages_are_valid_iff in Hm. destruct Hm as [Hinitial | Hemit]. - destruct Hinitial as [i [[mi Hmi] Hom]]. simpl in Hom. subst. @@ -1960,36 +1962,36 @@ Section Composite. rewrite Hsender in can_emit_signed. simpl in can_emit_signed. spec can_emit_signed; [|congruence]. red. eexists _, _, _. - apply (VLSM_incl_protocol_transition (constraint_preloaded_free_incl IM _)) in Ht. - apply pre_loaded_with_all_messages_projection_protocol_transition_eq with (j := i) in Ht; [|reflexivity]. + apply (VLSM_incl_input_valid_transition (constraint_preloaded_free_incl IM _)) in Ht. + apply pre_loaded_with_all_messages_projection_input_valid_transition_eq with (j := i) in Ht; [|reflexivity]. exact Ht. Qed. Lemma has_been_sent_iff_by_sender (Hsender_safety : sender_safety_alt_prop) - [is s tr] (Htr : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) is s tr) + [is s tr] (Htr : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) is s tr) [m v] (Hsender : sender m = Some v): composite_has_been_sent s m <-> has_been_sent (IM (A v)) (s (A v)) m. Proof. split;[|exists (A v);assumption]. intros [i Hi]. rewrite (Hsender_safety _ _ Hsender i);[assumption|]. - assert (finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm (IM i)) (is i) (s i) + assert (finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm (IM i)) (is i) (s i) (VLSM_projection_trace_project (preloaded_component_projection IM i) tr)). { - rewrite <- (ptrace_get_last Htr). - apply ptrace_forget_last in Htr. + rewrite <- (valid_trace_get_last Htr). + apply valid_trace_forget_last in Htr. destruct Htr as [Htr Hinit]. - apply ptrace_add_last. - split. revert Htr;apply preloaded_finite_ptrace_projection. + apply valid_trace_add_last. + split. revert Htr;apply preloaded_finite_valid_trace_projection. exact (Hinit i). symmetry. apply (VLSM_projection_finite_trace_last (preloaded_component_projection IM i)). assumption. } - apply (can_emit_from_protocol_trace (pre_loaded_with_all_messages_vlsm (IM i)) (is i) m + apply (can_emit_from_valid_trace (pre_loaded_with_all_messages_vlsm (IM i)) (is i) m (VLSM_projection_trace_project (preloaded_component_projection IM i) tr)). - exact (ptrace_forget_last H). + exact (valid_trace_forget_last H). apply (oracle_initial_trace_update (has_been_sent_stepwise_from_trace (has_been_sent_capabilities i)) (is i) (s i)). assumption. assumption. @@ -2080,122 +2082,122 @@ Section Composite. let (s', om') := (composite_transition IM l som) in not_heavy s'. - Lemma sent_component_protocol_composed + Lemma messages_sent_from_component_of_valid_state_are_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (i : index) (m : message) (Hsent : has_been_sent (IM i) (s i) m) : - protocol_message_prop X m. + valid_message_prop X m. Proof. pose (Xhbs := composite_HasBeenSentCapability constraint). - apply (sent_protocol X s). assumption. exists i;assumption. + apply (sent_valid X s). assumption. exists i;assumption. Qed. - Lemma preloaded_sent_component_protocol_composed + Lemma preloaded_messages_sent_from_component_of_valid_state_are_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (seed : message -> Prop) (X := pre_loaded_vlsm (composite_vlsm IM constraint) seed) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (i : index) (m : message) (Hsent : has_been_sent (IM i) (s i) m) : - protocol_message_prop X m. + valid_message_prop X m. Proof. pose (Xhbs := preloaded_composite_HasBeenSentCapability constraint seed). - apply (sent_protocol X s). assumption. exists i;assumption. + apply (sent_valid X s). assumption. exists i;assumption. Qed. - Lemma received_component_protocol_composed + Lemma messages_received_from_component_of_valid_state_are_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (i : index) (m : message) (Hreceived : has_been_received (IM i) (s i) m) - : protocol_message_prop X m. + : valid_message_prop X m. Proof. assert (Hcomp : has_been_received Free s m) by (exists i; assumption). - apply protocol_state_has_trace in Hs as H'. + apply valid_state_has_trace in Hs as H'. destruct H' as [is [tr Hpr]]. - assert (Htr : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm Free) is s tr). - { revert Hpr. apply VLSM_incl_finite_protocol_trace_init_to. + assert (Htr : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm Free) is s tr). + { revert Hpr. apply VLSM_incl_finite_valid_trace_init_to. apply VLSM_incl_trans with (MY := machine Free). - apply constraint_free_incl with (constraint0 := constraint). - apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - apply protocol_trace_input_is_protocol with (is0:=is) (tr0:=tr). - - apply ptrace_forget_last in Hpr; apply Hpr. + apply valid_trace_input_is_valid with (is0:=is) (tr0:=tr). + - apply valid_trace_forget_last in Hpr; apply Hpr. - apply proper_received in Hcomp; [apply (Hcomp is tr Htr)|]. - apply finite_protocol_trace_init_to_last in Htr as Heqs. - subst s. apply finite_ptrace_last_pstate. - apply proj1, finite_protocol_trace_from_to_forget_last in Htr. + apply finite_valid_trace_init_to_last in Htr as Heqs. + subst s. apply finite_valid_trace_last_pstate. + apply proj1, finite_valid_trace_from_to_forget_last in Htr. assumption. Qed. - Lemma composite_sent_protocol + Lemma composite_sent_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (m : message) (Hsent : composite_has_been_sent s m) - : protocol_message_prop X m. + : valid_message_prop X m. Proof. destruct Hsent as [i Hsent]. - apply sent_component_protocol_composed with s i + apply messages_sent_from_component_of_valid_state_are_valid with s i ; assumption. Qed. - Lemma preloaded_composite_sent_protocol + Lemma preloaded_composite_sent_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (seed : message -> Prop) (X := pre_loaded_vlsm (composite_vlsm IM constraint) seed) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (m : message) (Hsent : composite_has_been_sent s m) - : protocol_message_prop X m. + : valid_message_prop X m. Proof. destruct Hsent as [i Hsent]. - apply preloaded_sent_component_protocol_composed with s i + apply preloaded_messages_sent_from_component_of_valid_state_are_valid with s i ; assumption. Qed. - Lemma composite_received_protocol + Lemma composite_received_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (m : message) (Hreceived : composite_has_been_received has_been_received_capabilities s m) - : protocol_message_prop X m. + : valid_message_prop X m. Proof. destruct Hreceived as [i Hreceived]. - apply received_component_protocol_composed with s i + apply messages_received_from_component_of_valid_state_are_valid with s i ; assumption. Qed. - Lemma composite_observed_protocol + Lemma composite_observed_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) (s : composite_state IM) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (m : message) (Hobserved : composite_has_been_observed s m) - : protocol_message_prop X m. + : valid_message_prop X m. Proof. destruct Hobserved as [i Hobserved]. apply (has_been_observed_sent_received_iff (IM i)) in Hobserved. - destruct Hobserved as [Hreceived | Hsent]. - + apply received_component_protocol_composed with s i; assumption. - + apply sent_component_protocol_composed with s i; assumption. - - revert Hs. apply protocol_state_project_preloaded. + + apply messages_received_from_component_of_valid_state_are_valid with s i; assumption. + + apply messages_sent_from_component_of_valid_state_are_valid with s i; assumption. + - revert Hs. apply valid_state_project_preloaded. Qed. End Composite. @@ -2232,7 +2234,7 @@ Arguments has_been_sent_iff_by_sender {message index}%type_scope {IndEqDec} _%fu (finite_index : Listing index_listing) (i : index) (s : vstate (IM i)) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s) (m : message) : composite_has_been_observed IM Hbo (lift_to_composite_state IM i s) m <-> has_been_observed (HasBeenObservedCapability := Hbo i) (IM i) s m. Proof. @@ -2241,14 +2243,14 @@ Arguments has_been_sent_iff_by_sender {message index}%type_scope {IndEqDec} _%fu pose (free_composite_HasBeenReceivedCapability IM finite_index Hbr) as Free_Hbr. pose (free_composite_HasBeenSentCapability IM finite_index Hbs) as Free_Hbs. assert - (Hlift_s : protocol_state_prop (pre_loaded_with_all_messages_vlsm Free) (lift_to_composite_state IM i s)). - { revert Hs. apply protocol_state_preloaded_composite_free_lift. } + (Hlift_s : valid_state_prop (pre_loaded_with_all_messages_vlsm Free) (lift_to_composite_state IM i s)). + { revert Hs. apply valid_state_preloaded_composite_free_lift. } split; intros Hobs. - apply (proper_observed (Hbo i)); [assumption|]. intros is tr Htr. apply (proper_observed Free_Hbo) in Hobs ; [|assumption]. - apply (VLSM_full_projection_finite_protocol_trace_init_to (lift_to_composite_preloaded_vlsm_full_projection IM i)) in Htr as Hpre_tr. + apply (VLSM_full_projection_finite_valid_trace_init_to (lift_to_composite_preloaded_vlsm_full_projection IM i)) in Htr as Hpre_tr. specialize (Hobs _ _ Hpre_tr). apply Exists_exists. apply Exists_exists in Hobs. @@ -2267,7 +2269,7 @@ Arguments has_been_sent_iff_by_sender {message index}%type_scope {IndEqDec} _%fu apply (proper_observed (Hbo i)) in Hobs ; [|assumption]. apply has_been_observed_consistency in Hobs; [|apply Hbo|assumption]. destruct Hobs as [is [tr [Htr Hobs]]]. - apply (VLSM_full_projection_finite_protocol_trace_init_to (lift_to_composite_preloaded_vlsm_full_projection IM i)) in Htr as Hpre_tr. + apply (VLSM_full_projection_finite_valid_trace_init_to (lift_to_composite_preloaded_vlsm_full_projection IM i)) in Htr as Hpre_tr. eexists. eexists. exists Hpre_tr. apply Exists_exists. apply Exists_exists in Hobs. @@ -2299,11 +2301,11 @@ Lemma state_received_not_sent_trace_iff (s : state) (is : state) (tr : list transition_item) - (Htr : finite_protocol_trace_init_to PreX is s tr) + (Htr : finite_valid_trace_init_to PreX is s tr) : state_received_not_sent s m <-> trace_received_not_sent_before_or_after tr m. Proof. - assert (Hs : protocol_state_prop PreX s). - { apply proj1 in Htr. apply ptrace_last_pstate in Htr. + assert (Hs : valid_state_prop PreX s). + { apply proj1 in Htr. apply valid_trace_last_pstate in Htr. assumption. } split; intros [Hbrm Hnbsm]. @@ -2334,7 +2336,7 @@ Lemma state_received_not_sent_invariant_trace_iff (s : state) (is : state) (tr : list transition_item) - (Htr : finite_protocol_trace_init_to PreX is s tr) + (Htr : finite_valid_trace_init_to PreX is s tr) : state_received_not_sent_invariant s P <-> trace_received_not_sent_before_or_after_invariant tr P. Proof. @@ -2349,7 +2351,7 @@ A sent message cannot have been previously sent or received. *) Definition cannot_resend_message_stepwise_prop : Prop := forall l s oim s' m, - protocol_transition (pre_loaded_with_all_messages_vlsm X) l (s,oim) (s',Some m) -> + input_valid_transition (pre_loaded_with_all_messages_vlsm X) l (s,oim) (s',Some m) -> ~has_been_sent X s m /\ ~has_been_received X s' m. Lemma cannot_resend_received_message_in_future @@ -2378,8 +2380,8 @@ Qed. Context (Hno_resend : cannot_resend_message_stepwise_prop). - Lemma protocol_transition_received_not_resent l s m s' om' - (Ht : protocol_transition (pre_loaded_with_all_messages_vlsm X) l (s,Some m) (s', om')) + Lemma input_valid_transition_received_not_resent l s m s' om' + (Ht : input_valid_transition (pre_loaded_with_all_messages_vlsm X) l (s,Some m) (s', om')) : om' <> Some m. Proof. destruct om' as [m'|]; [|congruence]. @@ -2388,7 +2390,7 @@ Qed. elim Hnbr_m. clear Hnbr_m. apply exists_right_finite_trace_from in Ht. destruct Ht as [is [tr [Htr Hs]]]. - apply proj1 in Htr as Hlst. apply finite_protocol_trace_from_to_last_pstate in Hlst. + apply proj1 in Htr as Hlst. apply finite_valid_trace_from_to_last_pstate in Hlst. apply proper_received; [assumption|]. apply has_been_received_consistency; [assumption|assumption|]. exists _,_,Htr. @@ -2400,14 +2402,14 @@ Qed. (tr: list transition_item) (Htrm: trace_received_not_sent_before_or_after_invariant tr P) (is: state) - (Htr: finite_protocol_trace PreX is tr) - : finite_protocol_trace (pre_loaded_vlsm X P) is tr. + (Htr: finite_valid_trace PreX is tr) + : finite_valid_trace (pre_loaded_vlsm X P) is tr. Proof. unfold trace_received_not_sent_before_or_after_invariant in Htrm. split; [|apply Htr]. - induction Htr using finite_protocol_trace_rev_ind; intros. - - rapply @finite_ptrace_empty. - apply initial_is_protocol. assumption. + induction Htr using finite_valid_trace_rev_ind; intros. + - rapply @finite_valid_trace_from_empty. + apply initial_state_is_valid. assumption. - assert (trace_received_not_sent_before_or_after_invariant tr P) as Htrm'. { intros m [Hrecv Hsend]. apply (Htrm m);clear Htrm. split;[apply Exists_app;left;assumption|]. @@ -2422,21 +2424,21 @@ Qed. erewrite oracle_partial_trace_update. - left;exact Hrecv. - apply has_been_received_stepwise_from_trace. - - apply ptrace_add_default_last. apply Htr. + - apply valid_trace_add_default_last. apply Htr. } specialize (IHHtr Htrm'). apply (extend_right_finite_trace_from _ IHHtr). repeat split;try apply Hx; - [apply finite_ptrace_last_pstate;assumption|]. - destruct iom as [m|];[|apply option_protocol_message_None]. - (* If m was sent during tr, it is protocol because it was + [apply finite_valid_trace_last_pstate;assumption|]. + destruct iom as [m|];[|apply option_valid_message_None]. + (* If m was sent during tr, it is valid because it was produced in a valid (by IHHtr) trace. If m was not sent during tr, *) assert (Decision (trace_has_message (field_selector output) m tr)) as [Hsent|Hnot_sent]. apply (@Exists_dec _). intros. apply decide_eq. - + exact (protocol_trace_output_is_protocol _ _ _ IHHtr _ Hsent). - + apply initial_message_is_protocol. + + exact (valid_trace_output_is_valid _ _ _ IHHtr _ Hsent). + + apply initial_message_is_valid. right. apply Htrm. split. * apply Exists_app. right;apply Exists_cons. left;reflexivity. @@ -2454,20 +2456,20 @@ Qed. (P : message -> Prop) (s: state) (Hequiv_s: state_received_not_sent_invariant s P) - (Hs: protocol_state_prop PreX s) - : protocol_state_prop (pre_loaded_vlsm X P) s. + (Hs: valid_state_prop PreX s) + : valid_state_prop (pre_loaded_vlsm X P) s. Proof. - apply protocol_state_has_trace in Hs as Htr. + apply valid_state_has_trace in Hs as Htr. destruct Htr as [is [tr Htr]]. specialize (lift_preloaded_trace_to_seeded P tr) as Hlift. spec Hlift. { revert Hequiv_s. apply state_received_not_sent_invariant_trace_iff with is; assumption. } - specialize (Hlift _ (ptrace_forget_last Htr)). + specialize (Hlift _ (valid_trace_forget_last Htr)). apply proj1 in Hlift. - apply finite_ptrace_last_pstate in Hlift. - rewrite <- (ptrace_get_last Htr). assumption. + apply finite_valid_trace_last_pstate in Hlift. + rewrite <- (valid_trace_get_last Htr). assumption. Qed. Lemma lift_generated_to_seeded @@ -2475,18 +2477,18 @@ Qed. (s : state) (Hequiv_s: state_received_not_sent_invariant s P) (m : message) - (Hgen : protocol_generated_prop PreX s m) - : protocol_generated_prop (pre_loaded_vlsm X P) s m. + (Hgen : can_produce PreX s m) + : can_produce (pre_loaded_vlsm X P) s m. Proof. - apply non_empty_protocol_trace_from_protocol_generated_prop. - apply non_empty_protocol_trace_from_protocol_generated_prop in Hgen. + apply non_empty_valid_trace_from_can_produce. + apply non_empty_valid_trace_from_can_produce in Hgen. destruct Hgen as [is [tr [item [Htr Hgen]]]]. exists is, tr, item. split; [|assumption]. specialize (lift_preloaded_trace_to_seeded P tr) as Hlift. spec Hlift. { revert Hequiv_s. apply state_received_not_sent_invariant_trace_iff with is. - apply ptrace_add_last. assumption. + apply valid_trace_add_last. assumption. apply last_error_destination_last. destruct Hgen as [Hlst [Hs _]]. rewrite Hlst. subst. reflexivity. } @@ -2514,7 +2516,7 @@ Section has_been_sent_irrelevance. Lemma has_been_sent_irrelevance (s : state) (m : message) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm X) s) : has_been_sent1 s m -> has_been_sent2 s m. Proof. intro H. @@ -2525,7 +2527,7 @@ Section has_been_sent_irrelevance. End has_been_sent_irrelevance. -Section all_traces_to_protocol_state_are_protocol. +Section all_traces_to_valid_state_are_valid. Context {message : Type} @@ -2547,64 +2549,64 @@ Existing Instance X_HasBeenReceivedCapability. (** Under [HasBeenReceivedCapability] assumptions, and given the fact that -any protocol state @s@ has a protocol trace leading to it, -in which all (received) messages are protocol, it follows that -any message which [has_been_received] for state @s@ is protocol. +any valid state <> has a valid trace leading to it, +in which all (received) messages are valid, it follows that +any message which [has_been_received] for state <> is valid. -Hence, given any pre_loaded trace leading to @s@, all messages received -within it must be protocol, thus the trace itself is protocol. +Hence, given any pre_loaded trace leading to <>, all messages received +within it must be valid, thus the trace itself is valid. *) -Lemma pre_traces_with_protocol_inputs_are_protocol is s tr - (Htr : finite_protocol_trace_init_to PreX is s tr) +Lemma pre_traces_with_valid_inputs_are_valid is s tr + (Htr : finite_valid_trace_init_to PreX is s tr) (Hobs : forall m, trace_has_message (field_selector input) m tr -> - protocol_message_prop X m + valid_message_prop X m ) - : finite_protocol_trace_init_to X is s tr. + : finite_valid_trace_init_to X is s tr. Proof. revert s Htr Hobs. induction tr using rev_ind; intros; split ; [|apply Htr| | apply Htr] ; destruct Htr as [Htr Hinit]. - inversion Htr. subst. - apply (finite_ptrace_from_to_empty X). - apply initial_is_protocol. assumption. - - apply finite_protocol_trace_from_to_last in Htr as Hlst. - apply finite_protocol_trace_from_to_app_split in Htr. + apply (finite_valid_trace_from_to_empty X). + apply initial_state_is_valid. assumption. + - apply finite_valid_trace_from_to_last in Htr as Hlst. + apply finite_valid_trace_from_to_app_split in Htr. destruct Htr as [Htr Hx]. specialize (IHtr _ (conj Htr Hinit)). spec IHtr. { intros. apply Hobs. apply trace_has_message_prefix. assumption. } - apply proj1, finite_protocol_trace_from_to_forget_last in IHtr. - apply finite_protocol_trace_from_add_last; [|assumption]. + apply proj1, finite_valid_trace_from_to_forget_last in IHtr. + apply finite_valid_trace_from_add_last; [|assumption]. inversion Hx. subst f tl s'. apply (extend_right_finite_trace_from X) ; [assumption|]. destruct Ht as [[_ [_ [Hv Hc]]] Ht]. - apply finite_ptrace_last_pstate in IHtr as Hplst. + apply finite_valid_trace_last_pstate in IHtr as Hplst. repeat split; try assumption. - destruct iom as [m|]; [|apply option_protocol_message_None]. - apply option_protocol_message_Some. + destruct iom as [m|]; [|apply option_valid_message_None]. + apply option_valid_message_Some. apply Hobs. apply Exists_app. right. apply Exists_cons. left. subst. reflexivity. Qed. -Lemma all_pre_traces_to_protocol_state_are_protocol +Lemma all_pre_traces_to_valid_state_are_valid s - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) is tr - (Htr : finite_protocol_trace_init_to PreX is s tr) - : finite_protocol_trace_init_to X is s tr. + (Htr : finite_valid_trace_init_to PreX is s tr) + : finite_valid_trace_init_to X is s tr. Proof. - apply pre_traces_with_protocol_inputs_are_protocol in Htr; [assumption|]. - apply ptrace_last_pstate in Htr as Hspre. + apply pre_traces_with_valid_inputs_are_valid in Htr; [assumption|]. + apply valid_trace_last_pstate in Htr as Hspre. intros. apply - (composite_received_protocol IM finite_index + (composite_received_valid IM finite_index has_been_received_capabilities _ _ Hs ). specialize (proper_received _ s Hspre m) as Hproper. @@ -2613,7 +2615,7 @@ Proof. exists is, tr, Htr. assumption. Qed. -End all_traces_to_protocol_state_are_protocol. +End all_traces_to_valid_state_are_valid. Section has_been_received_in_state. @@ -2624,27 +2626,27 @@ Section has_been_received_in_state. . Lemma has_been_received_in_state s1 m: - protocol_state_prop X s1 -> + valid_state_prop X s1 -> has_been_received X s1 m -> exists (s0 : state) (item : transition_item) (tr : list transition_item), input item = Some m /\ - finite_protocol_trace_from_to X s0 s1 (item :: tr). + finite_valid_trace_from_to X s0 s1 (item :: tr). Proof. intros Hpsp Hhbr. - pose proof (Hetr := protocol_state_has_trace _ _ Hpsp). + pose proof (Hetr := valid_state_has_trace _ _ Hpsp). destruct Hetr as [ist [tr Hetr]]. apply proper_received in Hhbr. - 2: { apply pre_loaded_with_all_messages_protocol_state_prop. + 2: { apply pre_loaded_with_all_messages_valid_state_prop. apply Hpsp. } unfold selected_message_exists_in_all_preloaded_traces in Hhbr. unfold specialized_selected_message_exists_in_all_traces in Hhbr. specialize (Hhbr ist tr). - unfold finite_protocol_trace_init_to in Hhbr. - unfold finite_protocol_trace_init_to in Hetr. + unfold finite_valid_trace_init_to in Hhbr. + unfold finite_valid_trace_init_to in Hetr. destruct Hetr as [Hfptf Hisp]. - pose proof (Hfptf' := preloaded_weaken_finite_protocol_trace_from_to _ _ _ _ Hfptf). + pose proof (Hfptf' := preloaded_weaken_finite_valid_trace_from_to _ _ _ _ Hfptf). specialize (Hhbr (conj Hfptf' Hisp)). clear Hfptf'. @@ -2654,7 +2656,7 @@ Section has_been_received_in_state. apply elem_of_list_split in Htritemin. destruct Htritemin as [l1 [l2 Heqtr]]. rewrite Heqtr in Hfptf. - apply (finite_protocol_trace_from_to_app_split X) in Hfptf. + apply (finite_valid_trace_from_to_app_split X) in Hfptf. destruct Hfptf as [Htr1 Htr2]. destruct tritem eqn:Heqtritem. simpl in Hintritem. subst input. @@ -2665,14 +2667,14 @@ Section has_been_received_in_state. Qed. Lemma has_been_received_in_state_preloaded s1 m: - protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s1 -> + valid_state_prop (pre_loaded_with_all_messages_vlsm X) s1 -> has_been_received X s1 m -> exists (s0 : state) (item : transition_item) (tr : list transition_item), input item = Some m /\ - finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm X) s0 s1 (item :: tr). + finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm X) s0 s1 (item :: tr). Proof. intros Hpsp Hhbr. - pose proof (Hetr := protocol_state_has_trace _ _ Hpsp). + pose proof (Hetr := valid_state_has_trace _ _ Hpsp). destruct Hetr as [ist [tr Hetr]]. apply proper_received in Hhbr. 2: { apply Hpsp. } @@ -2680,8 +2682,8 @@ Section has_been_received_in_state. unfold selected_message_exists_in_all_preloaded_traces in Hhbr. unfold specialized_selected_message_exists_in_all_traces in Hhbr. specialize (Hhbr ist tr). - unfold finite_protocol_trace_init_to in Hhbr. - unfold finite_protocol_trace_init_to in Hetr. + unfold finite_valid_trace_init_to in Hhbr. + unfold finite_valid_trace_init_to in Hetr. destruct Hetr as [Hfptf Hisp]. specialize (Hhbr (conj Hfptf Hisp)). @@ -2691,7 +2693,7 @@ Section has_been_received_in_state. apply elem_of_list_split in Htritemin. destruct Htritemin as [l1 [l2 Heqtr]]. rewrite Heqtr in Hfptf. - apply (finite_protocol_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm X)) in Hfptf. + apply (finite_valid_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm X)) in Hfptf. destruct Hfptf as [Htr1 Htr2]. destruct tritem eqn:Heqtritem. simpl in Hintritem. subst input. diff --git a/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v b/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v index fa586f4b..6e1932dc 100644 --- a/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v +++ b/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v @@ -9,7 +9,7 @@ In this section we define fixed equivocation for the regular composition. Assuming that a only a fixed subset of the nodes (here called equivocating) are allowed to equivocate, inspired by the results leading to -Lemma [equivocators_protocol_trace_from_project] which links state equivocation +Lemma [equivocators_valid_trace_from_project] which links state equivocation to free composition, we can say that a message can be produced by a network of nodes allowed to equivocate, if it can be produced by their free composition [free_equivocating_vlsm_composition] preloaded with the current information @@ -49,7 +49,7 @@ Definition free_equivocating_vlsm_composition (** [pre_loaded_free_equivocating_vlsm_composition] preloads the free composition -of equivocating nodes all with the messages given by the @messageSet@. +of equivocating nodes all with the messages given by the <>. *) Definition pre_loaded_free_equivocating_vlsm_composition (messageSet : message -> Prop) @@ -265,18 +265,29 @@ Proof. apply fixed_equivocation_index_incl_subsumption. Qed. +Lemma fixed_equivocation_vlsm_composition_index_incl + : VLSM_incl + (fixed_equivocation_vlsm_composition IM Hbs Hbr indices1) + (fixed_equivocation_vlsm_composition IM Hbs Hbr indices2). +Proof. + apply constraint_subsumption_incl. + apply preloaded_constraint_subsumption_stronger. + apply strong_constraint_subsumption_strongest. + apply fixed_equivocation_constraint_index_incl_subsumption. +Qed. + End fixed_equivocation_index_incl. -(** ** Restricting Fixed protocol traces to only the equivocators +(** ** Restricting Fixed valid traces to only the equivocators -In this section we study the properties of Fixed protocol traces +In this section we study the properties of Fixed valid traces when projected to the composition of equivocator nodes. -The main result, Lemma [fixed_finite_protocol_trace_sub_projection] -is a strengthening of Lemma [finite_protocol_trace_sub_projection] for the +The main result, Lemma [fixed_finite_valid_trace_sub_projection] +is a strengthening of Lemma [finite_valid_trace_sub_projection] for the [fixed_equivocation_constraint], showing that, when restricting -a Fixed protocol trace to only the transitions belonging to the equivocators, -the obtained trace is protocol for the composition of equivocators pre-loaded +a Fixed valid trace to only the transitions belonging to the equivocators, +the obtained trace is valid for the composition of equivocators pre-loaded only with those messages sent by the non-equivocators in the original trace. *) Section fixed_equivocator_sub_projection. @@ -345,35 +356,35 @@ Proof. apply Hpreserve. Qed. -Lemma strong_fixed_equivocation_eqv_protocol base_s m +Lemma strong_fixed_equivocation_eqv_valid_message base_s m (Hstrong : strong_fixed_equivocation IM Hbs equivocators base_s m) - : protocol_message_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) m. + : valid_message_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) m. Proof. destruct Hstrong as [Hobs | Hemit]. - - apply initial_message_is_protocol. right. assumption. - - apply can_emit_protocol in Hemit. assumption. + - apply initial_message_is_valid. right. assumption. + - apply emitted_messages_are_valid in Hemit. assumption. Qed. -Lemma strong_fixed_equivocation_eqv_protocol_in_futures base_s s +Lemma strong_fixed_equivocation_eqv_valid_message_in_futures base_s s (Hfuture_s : in_futures PreFree s base_s) m (Hstrong : strong_fixed_equivocation IM Hbs equivocators s m) - : protocol_message_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) m. + : valid_message_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) m. Proof. destruct (in_futures_preserves_strong_fixed_equivocation _ _ Hfuture_s _ Hstrong) as [Hobs | Hemit]. - - apply initial_message_is_protocol. right. assumption. - - apply can_emit_protocol in Hemit. assumption. + - apply initial_message_is_valid. right. assumption. + - apply emitted_messages_are_valid in Hemit. assumption. Qed. -Section fixed_finite_protocol_trace_sub_projection_helper_lemmas. +Section fixed_finite_valid_trace_sub_projection_helper_lemmas. (** The results in this section are not meant to be used directly. They are just -sub-lemmas of [fixed_finite_protocol_trace_sub_projection_helper]. +sub-lemmas of [fixed_finite_valid_trace_sub_projection_helper]. We make an assumption (<>) which is later shown to hold -for all Fixed protocol states. +for all Fixed valid states. We then restate (some of) these lemmas without the extra assumption. *) @@ -388,7 +399,7 @@ Context (** See Lemma [fixed_input_has_strong_fixed_equivocation] below. *) Local Lemma fixed_input_has_strong_fixed_equivocation_helper l m - (Hv : protocol_valid Fixed l (s, Some m)) + (Hv : input_valid Fixed l (s, Some m)) : strong_fixed_equivocation IM Hbs equivocators base_s m. Proof. destruct Hv as [_ [_ [_ [Hobs | Hemit]]]]. @@ -398,20 +409,20 @@ Proof. apply VLSM_incl_can_emit; simpl. apply basic_VLSM_incl; intros s0 **; [assumption | | apply Hv | apply H]. destruct HmX as [Him | Hobs]. - + apply initial_message_is_protocol. left. assumption. + + apply initial_message_is_valid. left. assumption. + apply Hobs_s_protocol in Hobs as [Hsent | Hemit]. - * apply initial_message_is_protocol. right. assumption. - * apply can_emit_protocol. assumption. + * apply initial_message_is_valid. right. assumption. + * apply emitted_messages_are_valid. assumption. Qed. -Local Lemma fixed_protocol_transition_sub_projection_helper - (Hs_pr: protocol_state_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) +Local Lemma fixed_input_valid_transition_sub_projection_helper + (Hs_pr: valid_state_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) (composite_state_sub_projection IM equivocators s)) l (e : sub_index_prop equivocators (projT1 l)) iom oom sf - (Ht : protocol_transition Fixed l (s, iom) (sf, oom)) - : protocol_transition (equivocators_composition_for_sent IM Hbs equivocators base_s) + (Ht : input_valid_transition Fixed l (s, iom) (sf, oom)) + : input_valid_transition (equivocators_composition_for_sent IM Hbs equivocators base_s) (composite_label_sub_projection IM equivocators l e) (composite_state_sub_projection IM equivocators s, iom) (composite_state_sub_projection IM equivocators sf, oom). @@ -420,8 +431,8 @@ Proof. repeat split. - assumption. - apply proj1 in Ht. - destruct iom as [im|]; [|apply option_protocol_message_None]. - apply strong_fixed_equivocation_eqv_protocol. + destruct iom as [im|]; [|apply option_valid_message_None]. + apply strong_fixed_equivocation_eqv_valid_message. apply (fixed_input_has_strong_fixed_equivocation_helper _ _ Ht). - apply Ht. - apply proj2 in Ht. simpl in Ht. @@ -441,17 +452,17 @@ Qed. (** See Lemma [fixed_output_has_strong_fixed_equivocation] below. *) Local Lemma fixed_output_has_strong_fixed_equivocation_helper - (Hs_pr: protocol_state_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) + (Hs_pr: valid_state_prop (equivocators_composition_for_sent IM Hbs equivocators base_s) (composite_state_sub_projection IM equivocators s)) sf (Hfuture : in_futures PreFree sf base_s) l iom om - (Ht : protocol_transition Fixed l (s, iom) (sf, Some om)) + (Ht : input_valid_transition Fixed l (s, iom) (sf, Some om)) : strong_fixed_equivocation IM Hbs equivocators base_s om. Proof. destruct (decide (projT1 l ∈ equivocators)). - apply - (fixed_protocol_transition_sub_projection_helper Hs_pr _ e) in Ht. + (fixed_input_valid_transition_sub_projection_helper Hs_pr _ e) in Ht. right. eexists _,_,_. exact Ht. - left. exists (projT1 l). split; [assumption|]. @@ -459,8 +470,8 @@ Proof. (preloaded_component_projection IM (projT1 l))) in Hfuture. apply in_futures_preserving_oracle_from_stepwise with (field_selector output) (sf (projT1 l)) ; [apply has_been_sent_stepwise_from_trace| assumption |]. - apply (VLSM_incl_protocol_transition Fixed_incl_Preloaded) in Ht. - specialize (VLSM_projection_protocol_transition + apply (VLSM_incl_input_valid_transition Fixed_incl_Preloaded) in Ht. + specialize (VLSM_projection_input_valid_transition (preloaded_component_projection IM (projT1 l)) l (projT2 l)) as Hproject. spec Hproject. { unfold composite_project_label. @@ -472,36 +483,36 @@ Proof. apply (has_been_sent_step_update Hproject). left. reflexivity. Qed. -End fixed_finite_protocol_trace_sub_projection_helper_lemmas. +End fixed_finite_valid_trace_sub_projection_helper_lemmas. (** Using the lemmas above we can now prove (by induction) a generic result, which has as part of the conclusion the "observability implies strong_fixed_equivocation" hypothesis assumed in the above section. *) -Lemma fixed_finite_protocol_trace_sub_projection_helper +Lemma fixed_finite_valid_trace_sub_projection_helper si s tr - (Htr: finite_protocol_trace_init_to Fixed si s tr) + (Htr: finite_valid_trace_init_to Fixed si s tr) base_s (Hfuture: in_futures PreFree s base_s) - : finite_protocol_trace_from_to (equivocators_composition_for_sent IM Hbs equivocators base_s) + : finite_valid_trace_from_to (equivocators_composition_for_sent IM Hbs equivocators base_s) (composite_state_sub_projection IM equivocators si) (composite_state_sub_projection IM equivocators s) (finite_trace_sub_projection IM equivocators tr) /\ forall m, composite_has_been_observed IM Hbo s m -> strong_fixed_equivocation IM Hbs equivocators base_s m. Proof. - induction Htr using finite_protocol_trace_init_to_rev_ind. + induction Htr using finite_valid_trace_init_to_rev_ind. - split. - + apply finite_ptrace_from_to_empty. + + apply finite_valid_trace_from_to_empty. apply (composite_initial_state_sub_projection IM equivocators si) in Hsi. - apply initial_is_protocol. assumption. + apply initial_state_is_valid. assumption. + intros m Hobs. apply @has_been_observed_no_inits with (m := m) (Hhbo := Free_hbo) in Hsi. contradiction. - - apply (VLSM_incl_protocol_transition Fixed_incl_Preloaded) in Ht as Hpre_t. + - apply (VLSM_incl_input_valid_transition Fixed_incl_Preloaded) in Ht as Hpre_t. assert (Hfuture_s : in_futures PreFree s base_s). { destruct Hfuture as [tr' Htr']. - specialize (finite_ptrace_from_to_extend _ _ _ _ Htr' _ _ _ _ Hpre_t) as Htr''. + specialize (finite_valid_trace_from_to_extend _ _ _ _ Htr' _ _ _ _ Hpre_t) as Htr''. eexists. exact Htr''. } specialize (IHHtr Hfuture_s) as [Htr_pr Htr_obs]. @@ -510,8 +521,8 @@ Proof. intros m Hobs. apply @has_been_observed_step_update with (msg := m) (Hhbo := Free_hbo) in Hpre_t. apply Hpre_t in Hobs. destruct Hobs as [Hitem | Hobs] ; [| apply Htr_obs in Hobs; assumption]. - apply ptrace_last_pstate in Htr. - apply ptrace_last_pstate in Htr_pr. + apply valid_trace_last_pstate in Htr. + apply valid_trace_last_pstate in Htr_pr. destruct Hitem as [Hm | Hm]; subst. + apply (fixed_input_has_strong_fixed_equivocation_helper _ _ Htr_obs _ _ (proj1 Ht)). + apply (fixed_output_has_strong_fixed_equivocation_helper _ _ Htr_obs Htr_pr _ Hfuture _ _ _ Ht). @@ -520,10 +531,10 @@ Proof. unfold pre_VLSM_projection_transition_item_project. simpl. unfold composite_label_sub_projection_option. case_decide. - + eapply finite_protocol_trace_from_to_app; [apply Htr_pr|]. - apply finite_ptrace_from_to_singleton. simpl. - apply ptrace_last_pstate in Htr_pr. - apply (fixed_protocol_transition_sub_projection_helper _ _ Htr_obs Htr_pr _ H _ _ _ Ht). + + eapply finite_valid_trace_from_to_app; [apply Htr_pr|]. + apply finite_valid_trace_from_to_singleton. simpl. + apply valid_trace_last_pstate in Htr_pr. + apply (fixed_input_valid_transition_sub_projection_helper _ _ Htr_obs Htr_pr _ H _ _ _ Ht). + rewrite app_nil_r. replace (composite_state_sub_projection _ _ sf) with (composite_state_sub_projection IM equivocators s) ; [assumption|]. @@ -539,74 +550,74 @@ Proof. Qed. (** -Main result of this section: when restricting a Fixed protocol trace to only -the transitions belonging to the equivocators, the obtained trace is protocol +Main result of this section: when restricting a Fixed valid trace to only +the transitions belonging to the equivocators, the obtained trace is valid for the composition of equivocators pre-loaded with messages sent by the non-equivocators in the original trace. -This is a simple corollary of Lemma [fixed_finite_protocol_trace_sub_projection_helper] +This is a simple corollary of Lemma [fixed_finite_valid_trace_sub_projection_helper] proved above. *) -Lemma fixed_finite_protocol_trace_sub_projection is f tr - (Htr : finite_protocol_trace_init_to Fixed is f tr) - : finite_protocol_trace_init_to +Lemma fixed_finite_valid_trace_sub_projection is f tr + (Htr : finite_valid_trace_init_to Fixed is f tr) + : finite_valid_trace_init_to (equivocators_composition_for_sent IM Hbs equivocators f) (composite_state_sub_projection IM equivocators is) (composite_state_sub_projection IM equivocators f) (finite_trace_sub_projection IM equivocators tr). Proof. - apply fixed_finite_protocol_trace_sub_projection_helper with (base_s := f) in Htr as Htr_pr. + apply fixed_finite_valid_trace_sub_projection_helper with (base_s := f) in Htr as Htr_pr. - split; [apply Htr_pr|]. apply proj2 in Htr. specialize (composite_initial_state_sub_projection IM equivocators is Htr). exact id. - - apply in_futures_refl. apply ptrace_last_pstate in Htr. - apply (VLSM_incl_protocol_state Fixed_incl_Preloaded). assumption. + - apply in_futures_refl. apply valid_trace_last_pstate in Htr. + apply (VLSM_incl_valid_state Fixed_incl_Preloaded). assumption. Qed. (** -Any message observed in a Fixed protocol state has either been sent by the +Any message observed in a Fixed valid state has either been sent by the non-equivocating nodes, or it can be generated by the equivocating nodes using only the messages sent by the non-equivocating nodes. *) Lemma fixed_observed_has_strong_fixed_equivocation f - (Hf : protocol_state_prop Fixed f) + (Hf : valid_state_prop Fixed f) m (Hobs: composite_has_been_observed IM Hbo f m) : strong_fixed_equivocation IM Hbs equivocators f m. Proof. - apply (VLSM_incl_protocol_state Fixed_incl_Preloaded) in Hf as Hfuture. + apply (VLSM_incl_valid_state Fixed_incl_Preloaded) in Hf as Hfuture. apply in_futures_refl in Hfuture. - apply protocol_state_has_trace in Hf as [is [tr Htr]]. - apply fixed_finite_protocol_trace_sub_projection_helper with (base_s := f) + apply valid_state_has_trace in Hf as [is [tr Htr]]. + apply fixed_finite_valid_trace_sub_projection_helper with (base_s := f) in Htr as Htr_pr ; [|assumption..]. revert m Hobs. apply Htr_pr. Qed. -Lemma fixed_protocol_state_sub_projection s f +Lemma fixed_valid_state_sub_projection s f (Hsf : in_futures Fixed s f) - : protocol_state_prop + : valid_state_prop (equivocators_composition_for_sent IM Hbs equivocators f) (composite_state_sub_projection IM equivocators s). Proof. destruct Hsf as [tr Htr]. - apply finite_protocol_trace_from_to_complete_left in Htr as [is [trs [Htr Hs]]]. - apply fixed_finite_protocol_trace_sub_projection in Htr as Hpr_tr. - apply proj1, finite_protocol_trace_from_to_app_split,proj1, ptrace_forget_last in Htr. + apply finite_valid_trace_from_to_complete_left in Htr as [is [trs [Htr Hs]]]. + apply fixed_finite_valid_trace_sub_projection in Htr as Hpr_tr. + apply proj1, finite_valid_trace_from_to_app_split,proj1, valid_trace_forget_last in Htr. rewrite (finite_trace_sub_projection_app IM equivocators) in Hpr_tr. - apply proj1, finite_protocol_trace_from_to_app_split,proj1, ptrace_last_pstate in Hpr_tr. + apply proj1, finite_valid_trace_from_to_app_split,proj1, valid_trace_last_pstate in Hpr_tr. subst s. simpl. rewrite <- (finite_trace_sub_projection_last_state IM _ _ _ _ Htr). assumption. Qed. -(** The input of a Fixed protocol transition has the [strong_fixed_equivocation] +(** The input of a Fixed input valid transition has the [strong_fixed_equivocation] property. *) Lemma fixed_input_has_strong_fixed_equivocation l s m - (Ht : protocol_valid Fixed l (s, Some m)) + (Ht : input_valid Fixed l (s, Some m)) : strong_fixed_equivocation IM Hbs equivocators s m. Proof. apply fixed_input_has_strong_fixed_equivocation_helper with (base_s := s) in Ht @@ -615,25 +626,25 @@ Proof. apply Ht. Qed. -(** The output of a Fixed protocol transition has the [strong_fixed_equivocation] +(** The output of a Fixed input valid transition has the [strong_fixed_equivocation] property for its destination. *) Lemma fixed_output_has_strong_fixed_equivocation l s iom sf om - (Ht : protocol_transition Fixed l (s, iom) (sf, Some om)) + (Ht : input_valid_transition Fixed l (s, iom) (sf, Some om)) : strong_fixed_equivocation IM Hbs equivocators sf om. Proof. - apply protocol_transition_origin in Ht as Hs. + apply input_valid_transition_origin in Ht as Hs. apply fixed_output_has_strong_fixed_equivocation_helper with s sf l iom. - intros m Hobs. apply in_futures_preserves_strong_fixed_equivocation with s. - + apply (VLSM_incl_protocol_transition Fixed_incl_Preloaded) in Ht. - revert Ht. apply (protocol_transition_in_futures PreFree). + + apply (VLSM_incl_input_valid_transition Fixed_incl_Preloaded) in Ht. + revert Ht. apply (input_valid_transition_in_futures PreFree). + revert Hobs. apply fixed_observed_has_strong_fixed_equivocation. assumption. - - apply protocol_transition_in_futures in Ht. - revert Ht. apply fixed_protocol_state_sub_projection. - - apply in_futures_refl. apply protocol_transition_destination in Ht. - revert Ht. apply (VLSM_incl_protocol_state Fixed_incl_Preloaded). + - apply input_valid_transition_in_futures in Ht. + revert Ht. apply fixed_valid_state_sub_projection. + - apply in_futures_refl. apply input_valid_transition_destination in Ht. + revert Ht. apply (VLSM_incl_valid_state Fixed_incl_Preloaded). - assumption. Qed. @@ -671,12 +682,12 @@ Context . (** -Given a Fixed protocol state, the composition of the equivocators +Given a Fixed valid state, the composition of the equivocators preloaded with all the observed messages in the state is not stronger than that preloaded with only the messages sent by non-equivocators. *) Lemma Equivocators_Fixed_Strong_incl base_s - (Hbase_s : protocol_state_prop Fixed base_s) + (Hbase_s : valid_state_prop Fixed base_s) : VLSM_incl (equivocators_composition_for_observed IM Hbs Hbr equivocators base_s) (equivocators_composition_for_sent IM Hbs equivocators base_s). @@ -684,8 +695,8 @@ Proof. apply basic_VLSM_incl; intro; intros. - assumption. - destruct HmX as [Hinit | Hobs] - ; [apply initial_message_is_protocol; left; assumption|]. - apply strong_fixed_equivocation_eqv_protocol. + ; [apply initial_message_is_valid; left; assumption|]. + apply strong_fixed_equivocation_eqv_valid_message. revert Hobs. apply (fixed_observed_has_strong_fixed_equivocation IM Hbs Hbr finite_index). assumption. @@ -694,7 +705,7 @@ Proof. Qed. Lemma Equivocators_Fixed_Strong_eq base_s - (Hbase_s : protocol_state_prop Fixed base_s) + (Hbase_s : valid_state_prop Fixed base_s) : VLSM_eq (equivocators_composition_for_observed IM Hbs Hbr equivocators base_s) (equivocators_composition_for_sent IM Hbs equivocators base_s). @@ -705,10 +716,10 @@ Proof. Qed. (** The subsumption between [fixed_equivocation_constraint] and -[strong_fixed_equivocation_constraint] holds under [protocol_valid] assumptions. +[strong_fixed_equivocation_constraint] holds under [input_valid] assumptions. *) Lemma fixed_strong_equivocation_subsumption - : protocol_constraint_subsumption IM + : input_valid_constraint_subsumption IM (fixed_equivocation_constraint IM Hbs Hbr equivocators) (strong_fixed_equivocation_constraint IM Hbs equivocators). Proof. @@ -741,9 +752,9 @@ only depends on the messages sent by non-equivocating nodes, it makes it much easier to establish results related to changing the behavior of equivocating nodes. -In this section we essentially prove that given a Fixed protocol state <>, -any protocol trace over the composition of equivocators pre-loaded with the -messages observed in <> can be "lifted" to a Fixed protocol trace in which +In this section we essentially prove that given a Fixed valid state <>, +any valid trace over the composition of equivocators pre-loaded with the +messages observed in <> can be "lifted" to a Fixed valid trace in which the non-equivocators remain in their corresponding component-state given by <> (Lemma [EquivPreloadedBase_Fixed_weak_full_projection]). *) @@ -877,23 +888,23 @@ Qed. Context (base_s : composite_state IM) - (Hbase_s : protocol_state_prop Fixed base_s) + (Hbase_s : valid_state_prop Fixed base_s) (EquivPreloadedBase := equivocators_composition_for_sent IM Hbs equivocators base_s) . (** As a corollary of the above projection result, the state obtained by replacing -the equivocator component of a Fixed protocol state with initial states is -still a Fixed protocol state. +the equivocator component of a Fixed valid state with initial states is +still a Fixed valid state. *) Lemma fixed_equivocator_lifting_initial_state : weak_full_projection_initial_state_preservation EquivPreloadedBase Fixed (lift_sub_state_to IM equivocators base_s). Proof. intros eqv_is Heqv_is. - apply (VLSM_incl_protocol_state (StrongFixed_incl_Fixed IM Hbs Hbr equivocators)). - apply (VLSM_projection_protocol_state (remove_equivocating_transitions_fixed_projection _ Heqv_is)). + apply (VLSM_incl_valid_state (StrongFixed_incl_Fixed IM Hbs Hbr equivocators)). + apply (VLSM_projection_valid_state (remove_equivocating_transitions_fixed_projection _ Heqv_is)). revert Hbase_s. - apply (VLSM_incl_protocol_state (Fixed_incl_StrongFixed IM Hbs Hbr finite_index equivocators)). + apply (VLSM_incl_valid_state (Fixed_incl_StrongFixed IM Hbs Hbr finite_index equivocators)). Qed. Lemma lift_sub_state_to_sent_are_observed s @@ -916,7 +927,7 @@ Qed. (** *** Main result of the section -Given a protocol state <> for the <> composition (composition of +Given a valid state <> for the <> composition (composition of nodes where only the <> are allowed to message-equivocate), we can "lift" any trace of the free composition of just the equivocators pre-loaded with the messages observed in <> to a trace over the <> @@ -944,7 +955,7 @@ Proof. + destruct om as [m|]; [|exact I]. simpl. destruct Hv as [_ [Hm _]]. - apply can_emit_protocol_iff in Hm. + apply emitted_messages_are_valid_iff in Hm. destruct Hm as [[Hinit | Hobs]| Hemit]. * destruct Hinit as [i [[im Him] Heqm]]. destruct_dec_sig i j Hj Heqi. subst. @@ -973,9 +984,9 @@ Proof. - apply fixed_equivocator_lifting_initial_state. assumption. - destruct HmX as [Hm | Hm]. + destruct Hm as [(i, Hi) [[im Him] Heqm]]. - apply initial_message_is_protocol. + apply initial_message_is_valid. exists i. exists (exist _ im Him). subst. reflexivity. - + apply (composite_observed_protocol IM finite_index Hbs Hbo Hbr _ _ Hbase_s). + + apply (composite_observed_valid IM finite_index Hbs Hbo Hbr _ _ Hbase_s). revert Hm. apply sent_by_non_equivocating_are_observed. Qed. @@ -1064,3 +1075,114 @@ Proof. Qed. End fixed_equivocation_no_equivocators. + +Section fixed_non_equivocator_lifting. + +Context + {message : Type} + {index : Type} + {IndEqDec : EqDecision index} + {finite_index : finite.Finite index} + (IM : index -> VLSM message) + (Hbs : forall i, HasBeenSentCapability (IM i)) + (Hbr : forall i, HasBeenReceivedCapability (IM i)) + (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) + (equivocators : list index) + (non_equivocators := set_diff (finite.enum index) equivocators) + (Free := free_composite_vlsm IM) + (Fixed := fixed_equivocation_vlsm_composition IM Hbs Hbr equivocators) + (FixedNonEquivocating:= induced_sub_projection IM non_equivocators + (fixed_equivocation_constraint IM Hbs Hbr + equivocators)) + (StrongFixed := strong_fixed_equivocation_vlsm_composition IM Hbs equivocators) + (StrongFixedNonEquivocating:= induced_sub_projection IM non_equivocators + (strong_fixed_equivocation_constraint IM Hbs + equivocators)) + (PreFree := pre_loaded_with_all_messages_vlsm Free) + (finite_listing := FinFunExtras.listing_from_finite index) + (Free_hbo := free_composite_HasBeenObservedCapability IM finite_listing Hbo) + (Free_hbr := free_composite_HasBeenReceivedCapability IM finite_listing Hbr) + (Free_hbs := free_composite_HasBeenSentCapability IM finite_listing Hbs) + . + +(** All valid traces in the induced projection of the composition under the +strong fixed-equivocation constraint to the non-equivocating nodes can be lifted +to valid traces of the constrained composition. +*) +Lemma lift_strong_fixed_non_equivocating + : VLSM_full_projection StrongFixedNonEquivocating StrongFixed + (lift_sub_label IM non_equivocators) + (lift_sub_state IM non_equivocators). +Proof. + apply induced_sub_projection_lift. + intros s1 s2 Heq l om. + destruct om as [m|]; [|intuition]. + cut + (forall m, sent_by_non_equivocating IM Hbs equivocators s1 m -> + sent_by_non_equivocating IM Hbs equivocators s2 m). + { + intros Hsent_impl [[j [Hj Hsent]] | Hemit]. + - left. apply Hsent_impl. exists j; split; assumption. + - right. revert Hemit. + apply VLSM_incl_can_emit. + apply pre_loaded_vlsm_incl. + assumption. + } + clear -Heq. + intros m [i [Hi Hsent]]. + exists i. split; [assumption|]. + assert (Hi' : i ∈ non_equivocators) + by (apply set_diff_intro; [apply finite_index|assumption]). + apply f_equal_dep with (x := dexist i Hi') in Heq. + cbv in Heq. + rewrite <- Heq. + assumption. +Qed. + +(** All valid traces in the induced projection of the composition under the +fixed-equivocation constraint to the non-equivocating nodes can be lifted +to valid traces of the constrained composition. +*) +Lemma lift_fixed_non_equivocating + : VLSM_full_projection FixedNonEquivocating Fixed + (lift_sub_label IM non_equivocators) + (lift_sub_state IM non_equivocators). +Proof. + constructor. + intros sX trX Htr. + apply + (VLSM_incl_finite_valid_trace + (StrongFixed_incl_Fixed IM Hbs Hbr equivocators)). + apply (VLSM_full_projection_finite_valid_trace lift_strong_fixed_non_equivocating). + revert Htr. + apply VLSM_incl_finite_valid_trace. + apply induced_sub_projection_constraint_subsumption_incl. + apply fixed_strong_equivocation_subsumption with (finite.enum index). + apply FinFunExtras.listing_from_finite. +Qed. + +Lemma fixed_non_equivocating_projection_friendliness + : projection_friendly_prop + (induced_sub_projection_is_projection IM non_equivocators + (fixed_equivocation_constraint IM Hbs Hbr equivocators)). +Proof. + apply induced_sub_projection_friendliness. + apply lift_fixed_non_equivocating. +Qed. + +(** The valid traces of the induced projection of the composition under the +fixed-equivocation constraint to the non-equivocating nodes are precisely +projections of traces of the constrained composition. +*) +Lemma fixed_non_equivocating_traces_char is tr + : finite_valid_trace FixedNonEquivocating is tr <-> + exists eis etr, + finite_valid_trace Fixed eis etr /\ + composite_state_sub_projection IM non_equivocators eis = is /\ + finite_trace_sub_projection IM non_equivocators etr = tr. +Proof. + apply (projection_friendly_trace_char (induced_sub_projection_is_projection _ _ _)). + apply fixed_non_equivocating_projection_friendliness. +Qed. + +End fixed_non_equivocator_lifting. diff --git a/theories/VLSM/Core/Equivocation/FullNode.v b/theories/VLSM/Core/Equivocation/FullNode.v index f33a37e0..4f0cbdd7 100644 --- a/theories/VLSM/Core/Equivocation/FullNode.v +++ b/theories/VLSM/Core/Equivocation/FullNode.v @@ -28,12 +28,12 @@ Section full_node_constraint. Existing Instance Free_HasBeenSentCapability. (** - Given a composite state @s@, a message @m@, and a node index @i@ - if there is a machine we say that message @m@ can be - [node_generated_without_further_equivocation] by node @i@ if the message - can be produced by node @i@ pre_loaded with all messages in a trace in which + Given a composite state <>, a message <>, and a node index <> + if there is a machine we say that message <> can be + [node_generated_without_further_equivocation] by node <> if the message + can be produced by node <> pre_loaded with all messages in a trace in which all message equivocation is done through messages causing - [no_additional_equivocations] to state @s@ (message [has_been_observed] in @s@). + [no_additional_equivocations] to state <> (message [has_been_observed] in <>). *) Definition node_generated_without_further_equivocation (s : composite_state IM) @@ -41,13 +41,13 @@ Section full_node_constraint. (i : index) : Prop := exists (si : vstate (IM i)), - protocol_generated_prop (pre_loaded_with_all_messages_vlsm (IM i)) si m /\ + can_produce (pre_loaded_with_all_messages_vlsm (IM i)) si m /\ state_received_not_sent_invariant (IM i) si (composite_has_been_observed IM has_been_observed_capabilities s). (** Similar to the condition above, but now the message is required to be generated by the machine pre-loaded only with messages causing - [no_additional_equivocations] to state @s@. + [no_additional_equivocations] to state <>. *) Definition node_generated_without_further_equivocation_alt (s : composite_state IM) @@ -58,8 +58,8 @@ Section full_node_constraint. (** The equivocation-based abstract definition of the full node condition - stipulates that a message can be received in a state @s@ if either it causes - [no_additional_equivocations] to state @s@, or it can be + stipulates that a message can be received in a state <> if either it causes + [no_additional_equivocations] to state <>, or it can be [node_generated_without_further_equivocation] by an admissible node. *) Definition full_node_condition_for_admissible_equivocators @@ -102,7 +102,7 @@ Section full_node_constraint. (i : index) (Hno_resend : cannot_resend_message_stepwise_prop (IM i)) (s: composite_state IM) - (Hs: protocol_state_prop + (Hs: valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s) (m : message) (Hsmi : node_generated_without_further_equivocation s m i) @@ -133,7 +133,7 @@ Section full_node_constraint. specialize (Hno_resend i). apply node_generated_without_further_equivocation_weaken; [assumption| |assumption]. revert Hs. - apply VLSM_incl_protocol_state. + apply VLSM_incl_valid_state. apply preloaded_constraint_free_incl. Qed. diff --git a/theories/VLSM/Core/Equivocation/LimitedEquivocation.v b/theories/VLSM/Core/Equivocation/LimitedEquivocation.v index 0c7139a4..95b563ad 100644 --- a/theories/VLSM/Core/Equivocation/LimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocation/LimitedEquivocation.v @@ -60,13 +60,13 @@ Definition limited_equivocation_vlsm_composition := composite_vlsm IM limited_equivocation_constraint. -Lemma full_node_limited_equivocation_protocol_state_weight s - : protocol_state_prop limited_equivocation_vlsm_composition s -> +Lemma full_node_limited_equivocation_valid_state_weight s + : valid_state_prop limited_equivocation_vlsm_composition s -> tracewise_not_heavy s. Proof. intro Hs. unfold tracewise_not_heavy, not_heavy. - induction Hs using protocol_state_prop_ind. + induction Hs using valid_state_prop_ind. - replace (equivocation_fault s) with 0%R by (symmetry;apply initial_state_equivocators_weight;assumption). destruct threshold. simpl. apply Rge_le. assumption. @@ -85,7 +85,7 @@ Section fixed_limited_message_equivocation. In this section we show that if the set of allowed equivocators for a fixed equivocation constraint is of weight smaller than the threshold accepted for -limited message equivocation, then any protocol trace for the fixed equivocation +limited message equivocation, then any valid trace for the fixed equivocation constraint is also a trace under the limited equivocation constraint. *) @@ -119,8 +119,8 @@ Context (tracewise_equivocating_validators := @equivocating_validators _ _ _ _ Htracewise_BasicEquivocation) . -Lemma StrongFixed_protocol_state_not_heavy s - (Hs : protocol_state_prop StrongFixed s) +Lemma StrongFixed_valid_state_not_heavy s + (Hs : valid_state_prop StrongFixed s) : tracewise_not_heavy s. Proof. cut (tracewise_equivocating_validators s ⊆ equivocators). @@ -137,8 +137,8 @@ Proof. - apply (constraint_free_incl IM (strong_fixed_equivocation_constraint IM Hbs equivocators)). - apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - apply protocol_state_has_trace in Hs as [is [tr Htr]]. - apply (VLSM_incl_finite_protocol_trace_init_to StrongFixedinclPreFree) in Htr as Hpre_tr. + apply valid_state_has_trace in Hs as [is [tr Htr]]. + apply (VLSM_incl_finite_valid_trace_init_to StrongFixedinclPreFree) in Htr as Hpre_tr. intros v Hv. apply equivocating_validators_is_equivocating_tracewise_iff in Hv as Hvs'. specialize (Hvs' _ _ Hpre_tr). @@ -146,11 +146,11 @@ Proof. rewrite Heqtr in Htr. destruct Htr as [Htr Hinit]. change (pre ++ item::suf) with (pre ++ [item] ++ suf) in Htr. - apply (finite_protocol_trace_from_to_app_split StrongFixed) in Htr. + apply (finite_valid_trace_from_to_app_split StrongFixed) in Htr. destruct Htr as [Hpre Hitem]. - apply (VLSM_incl_finite_protocol_trace_from_to StrongFixedinclPreFree) in Hpre as Hpre_pre. - apply ptrace_last_pstate in Hpre_pre as Hs_pre. - apply (finite_protocol_trace_from_to_app_split StrongFixed), proj1 in Hitem. + apply (VLSM_incl_finite_valid_trace_from_to StrongFixedinclPreFree) in Hpre as Hpre_pre. + apply valid_trace_last_pstate in Hpre_pre as Hs_pre. + apply (finite_valid_trace_from_to_app_split StrongFixed), proj1 in Hitem. inversion Hitem; subst; clear Htl Hitem. simpl in Hm0. subst. destruct Ht as [[_ [_ [_ Hc]]] _]. destruct Hc as [[i [Hi Hsenti]] | Hemit]. @@ -170,8 +170,8 @@ Proof. intros (i, li) (s, om) Hpv. unfold limited_equivocation_constraint. destruct (composite_transition _ _ _) as (s', om') eqn:Ht. - specialize (protocol_transition_destination StrongFixed (conj Hpv Ht)) as Hs'. - apply StrongFixed_protocol_state_not_heavy in Hs'. + specialize (input_valid_transition_destination StrongFixed (conj Hpv Ht)) as Hs'. + apply StrongFixed_valid_state_not_heavy in Hs'. assumption. Qed. @@ -192,7 +192,7 @@ Section has_limited_equivocation. (** ** Limited Equivocation derived from Fixed Equivocation We say that a trace has the [fixed_limited_equivocation_prop]erty if it is -protocol for the composition using a [generalized_fixed_equivocation_constraint] +valid for the composition using a [generalized_fixed_equivocation_constraint] induced by a subset of indices whose weight is less than the allowed [ReachableThreshold]. *) @@ -213,7 +213,7 @@ Definition fixed_limited_equivocation_prop : Prop := exists (equivocators : list index) (Fixed := fixed_equivocation_vlsm_composition IM Hbs Hbr equivocators), (sum_weights (remove_dups equivocators) <= `threshold)%R /\ - finite_protocol_trace Fixed s tr. + finite_valid_trace Fixed s tr. Context {index_listing : list index} @@ -224,23 +224,23 @@ Context (Limited : VLSM message := limited_equivocation_vlsm_composition IM finite_index sender) . -(** Traces with the [fixed_limited_equivocation_prop]erty are protocol for the +(** Traces with the [fixed_limited_equivocation_prop]erty are valid for the composition using a [limited_equivocation_constraint]. *) -Lemma trace_exhibits_limited_equivocation_protocol +Lemma traces_exhibiting_limited_equivocation_are_valid (Hsender_safety : sender_safety_alt_prop IM (fun i => i) sender) - : forall s tr, fixed_limited_equivocation_prop s tr -> finite_protocol_trace Limited s tr. + : forall s tr, fixed_limited_equivocation_prop s tr -> finite_valid_trace Limited s tr. Proof. intros s tr [equivocators [Hlimited Htr]]. - eapply VLSM_incl_finite_protocol_trace; [| eassumption]. + eapply VLSM_incl_finite_valid_trace; [| eassumption]. apply Fixed_incl_Limited; assumption. Qed. (** Traces having the [strong_trace_witnessing_equivocation_prop]erty, which -are protocol for the free composition and whose final state is [not_heavy] have +are valid for the free composition and whose final state is [not_heavy] have the [fixed_limited_equivocation_prop]erty. *) -Lemma trace_exhibits_limited_equivocation_protocol_rev +Lemma traces_exhibiting_limited_equivocation_are_valid_rev (Hke : WitnessedEquivocationCapability IM id sender finite_index) (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) @@ -251,15 +251,14 @@ Lemma trace_exhibits_limited_equivocation_protocol_rev := equivocation_dec_tracewise IM (fun i => i) sender finite_index) (tracewise_not_heavy := @not_heavy _ _ _ _ Htracewise_basic_equivocation) : forall is s tr, strong_trace_witnessing_equivocation_prop IM id sender finite_index is tr -> - finite_protocol_trace_init_to (free_composite_vlsm IM) is s tr -> + finite_valid_trace_init_to (free_composite_vlsm IM) is s tr -> tracewise_not_heavy s -> fixed_limited_equivocation_prop is tr. Proof. intros is s tr Hstrong Htr Hnot_heavy. exists (equivocating_validators s). - split; [|split]; cycle 1. - - by eapply ptrace_forget_last, strong_witness_has_fixed_equivocation. - - apply Htr. + split; cycle 1. + - eapply valid_trace_forget_last, strong_witness_has_fixed_equivocation; eassumption. - replace (sum_weights _) with (equivocation_fault s); [assumption|]. apply set_eq_nodup_sum_weight_eq. + apply equivocating_validators_nodup. @@ -269,10 +268,10 @@ Proof. Qed. (** Traces with the [strong_trace_witnessing_equivocation_prop]erty, which are -protocol for the composition using a [limited_equivocation_constraint] +valid for the composition using a [limited_equivocation_constraint] have the [fixed_limited_equivocation_prop]erty. *) -Lemma limited_trace_exhibits_limited_equivocation_protocol_rev +Lemma limited_traces_exhibiting_limited_equivocation_are_valid_rev (Hke : WitnessedEquivocationCapability IM id sender finite_index) (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) @@ -280,47 +279,47 @@ Lemma limited_trace_exhibits_limited_equivocation_protocol_rev (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) (can_emit_signed : channel_authentication_prop IM id sender) : forall s tr, strong_trace_witnessing_equivocation_prop IM id sender finite_index s tr -> - finite_protocol_trace Limited s tr -> fixed_limited_equivocation_prop s tr. + finite_valid_trace Limited s tr -> fixed_limited_equivocation_prop s tr. Proof. intros s tr Hstrong Htr. apply proj1 in Htr as Hnot_heavy. - apply finite_ptrace_last_pstate, full_node_limited_equivocation_protocol_state_weight in Hnot_heavy. - assert (Hfree_tr : finite_protocol_trace (free_composite_vlsm IM) s tr). { - revert Htr. apply VLSM_incl_finite_protocol_trace. + apply finite_valid_trace_last_pstate, full_node_limited_equivocation_valid_state_weight in Hnot_heavy. + assert (Hfree_tr : finite_valid_trace (free_composite_vlsm IM) s tr). { + revert Htr. apply VLSM_incl_finite_valid_trace. apply constraint_free_incl. } clear Htr. - apply ptrace_add_default_last in Hfree_tr. - apply trace_exhibits_limited_equivocation_protocol_rev with (finite_trace_last s tr); assumption. + apply valid_trace_add_default_last in Hfree_tr. + apply traces_exhibiting_limited_equivocation_are_valid_rev with (finite_trace_last s tr); assumption. Qed. -(** Any state which is protocol for limited equivocation can be produced by +(** Any state which is valid for limited equivocation can be produced by a trace having the [fixed_limited_equivocation_prop]erty. *) -Lemma limited_protocol_state_has_trace_exhibiting_limited_equivocation +Lemma limited_valid_state_has_trace_exhibiting_limited_equivocation (Hke : WitnessedEquivocationCapability IM id sender finite_index) (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) (can_emit_signed : channel_authentication_prop IM id sender) - : forall s, protocol_state_prop Limited s -> + : forall s, valid_state_prop Limited s -> exists is tr, finite_trace_last is tr = s /\ fixed_limited_equivocation_prop is tr. Proof. intros s Hs. - assert (Hfree_s : protocol_state_prop (free_composite_vlsm IM) s). { + assert (Hfree_s : valid_state_prop (free_composite_vlsm IM) s). { revert Hs. - apply VLSM_incl_protocol_state. + apply VLSM_incl_valid_state. apply constraint_free_incl. } destruct (free_has_strong_trace_witnessing_equivocation_prop IM finite_index Hbs Hbr id sender finite_index _ s Hfree_s) as [is [tr [Htr Heqv]]]. exists is, tr. - apply ptrace_get_last in Htr as Hlst. + apply valid_trace_get_last in Htr as Hlst. split; [assumption|]. - apply full_node_limited_equivocation_protocol_state_weight in Hs. - apply trace_exhibits_limited_equivocation_protocol_rev with s. + apply full_node_limited_equivocation_valid_state_weight in Hs. + apply traces_exhibiting_limited_equivocation_are_valid_rev with s. all: assumption. Qed. diff --git a/theories/VLSM/Core/Equivocation/NoEquivocation.v b/theories/VLSM/Core/Equivocation/NoEquivocation.v index efa7ac13..5a7b908e 100644 --- a/theories/VLSM/Core/Equivocation/NoEquivocation.v +++ b/theories/VLSM/Core/Equivocation/NoEquivocation.v @@ -48,7 +48,7 @@ End no_equivocations. In this section we show that under [no_equivocations] assumptions: -- for any protocol state all messages [observed_were_sent]. +- for any valid state all messages [observed_were_sent]. - the [pre_loaded_with_all_messages_vlsm] is equal to the [no_equivocations] VLSM. *) @@ -58,7 +58,7 @@ Section NoEquivocationInvariants. (X: VLSM message) (Hhbs: HasBeenSentCapability X) (Hhbo: HasBeenObservedCapability X) - (Henforced: forall l s om, protocol_valid (pre_loaded_with_all_messages_vlsm X) l (s,om) -> no_equivocations X l (s,om)) + (Henforced: forall l s om, input_valid (pre_loaded_with_all_messages_vlsm X) l (s,om) -> no_equivocations X l (s,om)) . (** @@ -82,13 +82,13 @@ any message that tests as [has_been_observed] in a state also tests as Qed. Lemma observed_were_sent_preserved l s im s' om: - protocol_transition X l (s,im) (s',om) -> + input_valid_transition X l (s,im) (s',om) -> observed_were_sent s -> observed_were_sent s'. Proof. intros Hptrans Hprev msg Hobs. specialize (Hprev msg). - apply preloaded_weaken_protocol_transition in Hptrans. + apply preloaded_weaken_input_valid_transition in Hptrans. apply (oracle_step_update has_been_observed_stepwise_props _ _ _ _ _ Hptrans) in Hobs. simpl in Hobs. specialize (Henforced l s (Some msg)). @@ -106,11 +106,11 @@ any message that tests as [has_been_observed] in a state also tests as Qed. Lemma observed_were_sent_invariant s: - protocol_state_prop X s -> + valid_state_prop X s -> observed_were_sent s. Proof. intro Hproto. - induction Hproto using protocol_state_prop_ind. + induction Hproto using valid_state_prop_ind. - apply observed_were_sent_initial. assumption. - revert Ht IHHproto. apply observed_were_sent_preserved. Qed. @@ -124,25 +124,25 @@ any message that tests as [has_been_observed] in a state also tests as Lemma no_equivocations_preloaded_traces (is : state) (tr : list transition_item) - : finite_protocol_trace (pre_loaded_with_all_messages_vlsm X) is tr -> finite_protocol_trace X is tr. + : finite_valid_trace (pre_loaded_with_all_messages_vlsm X) is tr -> finite_valid_trace X is tr. Proof. intro Htr. - induction Htr using finite_protocol_trace_rev_ind. + induction Htr using finite_valid_trace_rev_ind. - split;[|assumption]. - rapply @finite_ptrace_empty. - apply initial_is_protocol. + rapply @finite_valid_trace_from_empty. + apply initial_state_is_valid. assumption. - destruct IHHtr as [IHtr His]. split; [|assumption]. rapply extend_right_finite_trace_from;[assumption|]. - apply finite_ptrace_last_pstate in IHtr as Hs. - cut (option_protocol_message_prop X iom);[firstorder|]. - destruct iom as [m|];[|apply option_protocol_message_None]. + apply finite_valid_trace_last_pstate in IHtr as Hs. + cut (option_valid_message_prop X iom);[firstorder|]. + destruct iom as [m|];[|apply option_valid_message_None]. destruct Hx as [Hv _]. apply Henforced in Hv. destruct Hv as [Hbsm | []]. revert Hbsm. - apply sent_protocol. + apply sent_valid. assumption. Qed. @@ -218,7 +218,7 @@ Section CompositeNoEquivocationInvariants. forall msg, composite_has_been_observed IM Hbo s msg -> composite_has_been_sent IM Hbs s msg. Lemma composite_observed_were_sent_invariant s: - protocol_state_prop X s -> + valid_state_prop X s -> composite_observed_were_sent s. Proof. intro Hs. diff --git a/theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v b/theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v index f5739e63..42b4af62 100644 --- a/theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v +++ b/theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v @@ -148,7 +148,7 @@ Definition is_equivocating_tracewise : Prop := forall is tr - (Hpr : finite_protocol_trace_init_to PreFree is s tr), + (Hpr : finite_valid_trace_init_to PreFree is s tr), exists (m : message), (sender m = Some v) /\ exists prefix elem suffix (lprefix := finite_trace_last is prefix), @@ -166,7 +166,7 @@ Definition is_equivocating_tracewise_no_has_been_sent : Prop := forall is tr - (Htr : finite_protocol_trace_init_to PreFree is s tr), + (Htr : finite_valid_trace_init_to PreFree is s tr), exists (m : message), (sender m = Some v) /\ equivocation_in_trace PreFree m tr. @@ -176,7 +176,7 @@ Lemma is_equivocating_tracewise_no_has_been_sent_equivocating_senders_in_trace (v : validator) : is_equivocating_tracewise_no_has_been_sent s v <-> forall is tr - (Htr : finite_protocol_trace_init_to PreFree is s tr), + (Htr : finite_valid_trace_init_to PreFree is s tr), v ∈ equivocating_senders_in_trace tr. Proof. split; intros Heqv is tr Htr; specialize (Heqv _ _ Htr) @@ -198,10 +198,10 @@ Proof. apply and_proper_l; intro Hsender. apply exist_proper; intro prefix. apply exist_proper; intro item. apply exist_proper; intro suffix. apply and_proper_l. intro Htreq. apply and_iff_compat_l. apply not_iff_compat. - assert (Hpre : finite_protocol_trace_init_to PreFree is (finite_trace_last is prefix) prefix). + assert (Hpre : finite_valid_trace_init_to PreFree is (finite_trace_last is prefix) prefix). { split; [|apply Htr]. subst tr. - apply proj1, finite_protocol_trace_from_to_app_split, proj1 in Htr. + apply proj1, finite_valid_trace_from_to_app_split, proj1 in Htr. assumption. } pose proof (CHbs := composite_has_been_sent_stepwise_props IM Hbs (free_constraint IM)). @@ -212,7 +212,7 @@ Qed. Lemma transition_is_equivocating_tracewise_char l s om s' om' - (Ht : protocol_transition PreFree l (s, om) (s', om')) + (Ht : input_valid_transition PreFree l (s, om) (s', om')) (v : validator) : is_equivocating_tracewise_no_has_been_sent s' v -> is_equivocating_tracewise_no_has_been_sent s v \/ @@ -238,7 +238,7 @@ Qed. Lemma transition_receiving_no_sender_reflects_is_equivocating_tracewise l s om s' om' - (Ht : protocol_transition PreFree l (s, om) (s', om')) + (Ht : input_valid_transition PreFree l (s, om) (s', om')) (Hno_sender : option_bind _ _ sender om = None) (v : validator) : is_equivocating_tracewise_no_has_been_sent s' v -> is_equivocating_tracewise_no_has_been_sent s v. @@ -253,9 +253,9 @@ Lemma is_equivocating_statewise_implies_is_equivocating_tracewise s v Proof. intros [j [m [Hm [Hnbs_m Hbr_m]]]] is tr Htr. exists m. split; [assumption|]. - apply preloaded_finite_ptrace_init_to_projection with (j0 := j) in Htr as Htrj. + apply preloaded_finite_valid_trace_init_to_projection with (j0 := j) in Htr as Htrj. apply proj1 in Htrj as Hlstj. - apply finite_protocol_trace_from_to_last_pstate in Hlstj. + apply finite_valid_trace_from_to_last_pstate in Hlstj. apply proper_received in Hbr_m; [|assumption]. specialize (Hbr_m _ _ Htrj). @@ -270,12 +270,12 @@ Proof. subst. clear -Hnbs_m Htr. - apply preloaded_finite_ptrace_init_to_projection with (j := A v) in Htr as Htrv. - apply proj1, finite_protocol_trace_from_to_app_split,proj1 - , preloaded_finite_ptrace_from_to_projection with (j := A v) - , finite_protocol_trace_from_to_last in Htr. + apply preloaded_finite_valid_trace_init_to_projection with (j := A v) in Htr as Htrv. + apply proj1, finite_valid_trace_from_to_app_split,proj1 + , preloaded_finite_valid_trace_from_to_projection with (j := A v) + , finite_valid_trace_from_to_last in Htr. rewrite (VLSMProjections.VLSM_projection_trace_project_app (preloaded_component_projection IM (A v))) in Htrv. - apply proj1, (finite_protocol_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm (IM (A v)))),proj2 in Htrv. + apply proj1, (finite_valid_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm (IM (A v)))),proj2 in Htrv. rewrite Htr in Htrv. intro Hbs_m. elim Hnbs_m. clear Hnbs_m. @@ -293,7 +293,7 @@ Lemma initial_state_not_is_equivocating_tracewise Proof. intros Heqv. specialize (Heqv s []). - spec Heqv. { split; [|assumption]. constructor. apply initial_is_protocol. assumption. } + spec Heqv. { split; [|assumption]. constructor. apply initial_state_is_valid. assumption. } destruct Heqv as [m [_ [prefix [suf [item [Heq _]]]]]]. destruct prefix; inversion Heq. Qed. @@ -341,13 +341,13 @@ Proof. assumption. Qed. -Lemma protocol_transition_receiving_no_sender_reflects_equivocating_validators +Lemma input_valid_transition_receiving_no_sender_reflects_equivocating_validators l s om s' om' - (Ht : protocol_transition PreFree l (s, om) (s', om')) + (Ht : input_valid_transition PreFree l (s, om) (s', om')) (Hno_sender : option_bind _ _ sender om = None) : equivocating_validators s' ⊆ equivocating_validators s. Proof. - intro v. + intro v. intro Hs'. apply equivocating_validators_is_equivocating_tracewise_iff in Hs'. apply equivocating_validators_is_equivocating_tracewise_iff. apply transition_receiving_no_sender_reflects_is_equivocating_tracewise with l om s' om' @@ -367,11 +367,11 @@ Qed. Lemma composite_transition_no_sender_equivocators_weight l s om s' om' - (Ht : protocol_transition PreFree l (s, om) (s', om')) + (Ht : input_valid_transition PreFree l (s, om) (s', om')) (Hno_sender : option_bind _ _ sender om = None) : (equivocation_fault s' <= equivocation_fault s)%R. Proof. - specialize (protocol_transition_receiving_no_sender_reflects_equivocating_validators _ _ _ _ _ Ht Hno_sender) as Heqv. + specialize (input_valid_transition_receiving_no_sender_reflects_equivocating_validators _ _ _ _ _ Ht Hno_sender) as Heqv. revert Heqv. apply incl_equivocating_validators_equivocation_fault. Qed. diff --git a/theories/VLSM/Core/Equivocation/WitnessedEquivocation.v b/theories/VLSM/Core/Equivocation/WitnessedEquivocation.v index e470528c..8ddc02dc 100644 --- a/theories/VLSM/Core/Equivocation/WitnessedEquivocation.v +++ b/theories/VLSM/Core/Equivocation/WitnessedEquivocation.v @@ -15,14 +15,14 @@ The witnessed equivocation assumption is a possible way to address this issue. Starting from the (reasonable) assumption that for any state <>, there is a trace ending in <> whose [equivocating_senders_in_trace] are precisely the equivocators of <> (the [WitnessedEquivocationCapability]), -we can show that for each Free protocol state there exists -a protocol trace with the [strong_trace_witnessing_equivocation_prop]erty, +we can show that for each Free valid state there exists +a valid trace with the [strong_trace_witnessing_equivocation_prop]erty, i.e., a trace whose every prefix is a witness for its corresponding end state (Lemma [free_has_strong_trace_witnessing_equivocation_prop]). In particular, the set of equivocators is monotonically increasing for such a trace (Lemma [strong_witness_equivocating_validators_prefix_monotonicity]). -We then use this result to show that any Free protocol state is also a protocol +We then use this result to show that any Free valid state is also a valid state for the composition of nodes under the [fixed_equivocation_constraint] induced by its set of equivocators. *) @@ -91,8 +91,8 @@ valid states there exist a trace witnessing its equivocation. Class WitnessedEquivocationCapability := { is_equivocating_tracewise_witness : - forall s, protocol_state_prop PreFree s -> - exists is tr, finite_protocol_trace_init_to PreFree is s tr /\ + forall s, valid_state_prop PreFree s -> + exists is tr, finite_valid_trace_init_to PreFree is s tr /\ trace_witnessing_equivocation_prop is tr }. @@ -125,7 +125,7 @@ its final transition is monotonic w.r.t. the [equivocating_validators]. Lemma equivocating_validators_witness_monotonicity (is s : composite_state IM) (tr : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_init_to PreFree is s tr) + (Htr : finite_valid_trace_init_to PreFree is s tr) (item : composite_transition_item IM) (Hwitness : trace_witnessing_equivocation_prop is (tr ++ [item])) (s' := destination item) @@ -148,17 +148,17 @@ are included in the [equivocating_validators] for the source of its last transition, the the trace without its last transition also has the [trace_witnessing_equivocation_prop]erty. *) -Lemma protocol_transition_reflects_trace_witnessing_equivocation_prop +Lemma input_valid_transition_reflects_trace_witnessing_equivocation_prop (is s: composite_state IM) (tr: list (composite_transition_item IM)) - (Htr: finite_protocol_trace_init_to PreFree is s tr) + (Htr: finite_valid_trace_init_to PreFree is s tr) (item : composite_transition_item IM) (Hwitness : trace_witnessing_equivocation_prop is (tr ++ [item])) (s' := destination item) (Hincl : equivocating_validators s' ⊆ equivocating_validators s) : trace_witnessing_equivocation_prop is tr. Proof. - apply finite_protocol_trace_init_to_last in Htr as Hlst. + apply finite_valid_trace_init_to_last in Htr as Hlst. simpl in Hlst. intros v. split; intros; simpl in *; rewrite Hlst in *. - apply equivocating_validators_is_equivocating_tracewise_iff in H. @@ -179,14 +179,14 @@ message is not sent by any trace witnessing the source of the transition. *) Lemma equivocating_validators_step_update l s om s' om' - (Ht : protocol_transition PreFree l (s, om) (s', om')) + (Ht : input_valid_transition PreFree l (s, om) (s', om')) v : v ∈ equivocating_validators s' -> v ∈ equivocating_validators s \/ (exists m, om = Some m /\ sender m = Some v /\ forall is tr - (Htr : finite_protocol_trace_init_to PreFree is s tr) + (Htr : finite_valid_trace_init_to PreFree is s tr) (Hwitness : trace_witnessing_equivocation_prop is tr), ~ trace_has_message (field_selector output) m tr). Proof. @@ -210,7 +210,7 @@ Proof. inversion Heq_m. assumption. - elim Hnv. spec Hwitness v. - apply finite_protocol_trace_from_to_last in Htr as Hs. + apply finite_valid_trace_from_to_last in Htr as Hs. simpl in Hs, Hwitness. rewrite Hs in Hwitness. apply Hwitness. exists m'. split; [assumption|]. exists prefix, item, suffix'. @@ -226,7 +226,7 @@ Given a non-empty trace with the [trace_witnessing_equivocation_prop]erty, there are two disjoint possibilities concerning its last transition. (1) either it preserves the set of [equivocating_validators] and, in that case, -the trace without the last transition has the +the trace without the last transition has the [trace_witnessing_equivocation_prop]erty as well; or (2) The set of [equivocating_validators] of its destination is obtained @@ -242,7 +242,7 @@ Lemma equivocating_validators_witness_last_char (s' : composite_state IM) (om' : option message) (item := {| l := l; input := om; destination := s'; output := om' |}) - (Htr_item : finite_protocol_trace_init_to PreFree is s' (tr ++ [item])) + (Htr_item : finite_valid_trace_init_to PreFree is s' (tr ++ [item])) (Hwitness : trace_witnessing_equivocation_prop is (tr ++ [item])) (s := finite_trace_last is tr) : set_eq (equivocating_validators s) (equivocating_validators s') @@ -253,19 +253,19 @@ Lemma equivocating_validators_witness_last_char v ∉ equivocating_validators s /\ set_eq (equivocating_validators s') (set_add v (equivocating_validators s)) /\ forall (is : _composite_state IM) (tr : list transition_item), - finite_protocol_trace_init_to PreFree is s tr -> + finite_valid_trace_init_to PreFree is s tr -> trace_witnessing_equivocation_prop is tr -> ~ trace_has_message (field_selector output) m tr ). Proof. destruct Htr_item as [Htr Hinit]. - apply finite_protocol_trace_from_to_app_split in Htr. + apply finite_valid_trace_from_to_app_split in Htr. destruct Htr as [Htr Hitem]. inversion_clear Hitem. clear Htl. subst s. apply equivocating_validators_witness_monotonicity with (s := (finite_trace_last is tr)) in Hwitness as Hincl ; [| split; assumption]. - specialize (protocol_transition_receiving_no_sender_reflects_equivocating_validators IM A sender finite_validator _ _ _ _ _ Ht) + specialize (input_valid_transition_receiving_no_sender_reflects_equivocating_validators IM A sender finite_validator _ _ _ _ _ Ht) as Hreflect. remember (finite_trace_last is tr) as s. destruct (option_bind _ _ sender om) as [v|] eqn:Heq_v. @@ -273,7 +273,7 @@ Proof. destruct (decide (set_eq (equivocating_validators s) (equivocating_validators s'))) ; [left;split; [assumption|]|]. + apply - (protocol_transition_reflects_trace_witnessing_equivocation_prop + (input_valid_transition_reflects_trace_witnessing_equivocation_prop _ _ _ (conj Htr Hinit) _ Hwitness). subst. apply s0. + right. exists m; split; [reflexivity|]. exists v. split; [assumption|]. @@ -307,7 +307,7 @@ Proof. left. split. + split; subst; assumption. + apply - (protocol_transition_reflects_trace_witnessing_equivocation_prop + (input_valid_transition_reflects_trace_witnessing_equivocation_prop _ _ _ (conj Htr Hinit) _ Hwitness ). apply Hreflect. @@ -328,7 +328,7 @@ is that is guantees monotonicity of [equivocating_validators] along the trace. Lemma strong_witness_equivocating_validators_prefix_monotonicity (is s : composite_state IM) (tr : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_init_to PreFree is s tr) + (Htr : finite_valid_trace_init_to PreFree is s tr) (Hwitness : strong_trace_witnessing_equivocation_prop is tr) prefix suffix (Heqtr : prefix ++ suffix = tr) @@ -336,7 +336,7 @@ Lemma strong_witness_equivocating_validators_prefix_monotonicity : equivocating_validators ps ⊆ equivocating_validators s. Proof. revert prefix suffix Heqtr ps. - induction Htr using finite_protocol_trace_init_to_rev_ind; intros. + induction Htr using finite_valid_trace_init_to_rev_ind; intros. - apply app_eq_nil in Heqtr. destruct Heqtr; subst. reflexivity. - remember {| input := iom |} as item. @@ -345,7 +345,7 @@ Proof. specialize (Hwitness pre (suf ++ [item])). apply Hwitness. subst. apply app_assoc. - } + } destruct_list_last suffix suffix' _item Heqsuffix. + rewrite app_nil_r in Heqtr. subst. subst ps. rewrite finite_trace_last_is_last. simpl. reflexivity. @@ -369,9 +369,9 @@ the proof of Lemma [preloaded_has_strong_trace_witnessing_equivocation_prop]. Lemma strong_trace_witnessing_equivocation_prop_extend_eq s is tr' - (Htr': finite_protocol_trace_init_to PreFree is s tr') + (Htr': finite_valid_trace_init_to PreFree is s tr') is' tr'' - (Htr'': finite_protocol_trace_init_to PreFree is' s tr'') + (Htr'': finite_valid_trace_init_to PreFree is' s tr'') (Hprefix : strong_trace_witnessing_equivocation_prop is' tr'') item (Hwitness : trace_witnessing_equivocation_prop is (tr' ++ [item])) @@ -385,7 +385,7 @@ Proof. intro v. rewrite finite_trace_last_is_last. specialize (Hprefix tr'' []). spec Hprefix; [apply app_nil_r|]. - apply ptrace_get_last in Htr'' as Hlst'. + apply valid_trace_get_last in Htr'' as Hlst'. split. + intros Hv. apply Heq in Hv. rewrite <- Hlst' in Hv. @@ -403,7 +403,7 @@ Proof. * destruct Heqv as [Heq_om Heqv]. assert (Heqv' : ~ trace_has_message (field_selector output) m tr'). { intro Heqv'. elim Heqv. - apply ptrace_last_pstate in Htr' as Htr'_lst. + apply valid_trace_last_pstate in Htr' as Htr'_lst. destruct (has_been_sent_consistency Free _ Htr'_lst m @@ -431,7 +431,7 @@ Qed. Lemma strong_trace_witnessing_equivocation_prop_extend_neq s is tr - (Htr: finite_protocol_trace_init_to PreFree is s tr) + (Htr: finite_valid_trace_init_to PreFree is s tr) (Hprefix : strong_trace_witnessing_equivocation_prop is tr) item msg @@ -449,7 +449,7 @@ Proof. intro v'. rewrite finite_trace_last_is_last. simpl. specialize (Hprefix tr []). spec Hprefix; [apply app_nil_r|]. - apply ptrace_get_last in Htr as Hlst'. + apply valid_trace_get_last in Htr as Hlst'. split. + intros Hv'. apply Hneq in Hv'. apply set_add_iff in Hv'. @@ -496,16 +496,16 @@ the induction hypothesis via property (2). The conclusion then follows by the two helper lemmas above. *) Lemma preloaded_has_strong_trace_witnessing_equivocation_prop s - (Hs : protocol_state_prop PreFree s) + (Hs : valid_state_prop PreFree s) : exists is' tr', - finite_protocol_trace_init_to PreFree is' s tr' /\ + finite_valid_trace_init_to PreFree is' s tr' /\ strong_trace_witnessing_equivocation_prop is' tr'. Proof. apply is_equivocating_tracewise_witness in Hs. destruct Hs as [is [tr [Htr Hwitness]]]. - apply finite_protocol_trace_init_to_last in Htr as Hlst. + apply finite_valid_trace_init_to_last in Htr as Hlst. subst s. - apply finite_protocol_trace_init_to_forget_last in Htr. + apply finite_valid_trace_init_to_forget_last in Htr. remember (length tr) as n. remember (length (equivocating_validators (finite_trace_last is tr))) as m. revert m n is tr Heqm Heqn Htr Hwitness. @@ -514,11 +514,11 @@ Proof. forall is tr, m = length (equivocating_validators (finite_trace_last is tr)) -> n = length tr -> - finite_protocol_trace PreFree is tr -> + finite_valid_trace PreFree is tr -> trace_witnessing_equivocation_prop is tr -> let s := finite_trace_last is tr in exists (is' : state) (tr' : list transition_item), - finite_protocol_trace_init_to PreFree is' s tr' /\ + finite_valid_trace_init_to PreFree is' s tr' /\ (forall prefix suffix : list transition_item, prefix ++ suffix = tr' -> trace_witnessing_equivocation_prop is' prefix)) @@ -530,14 +530,14 @@ Proof. intros _ _ Htr _. exists is, []. split. - + apply finite_protocol_trace_init_add_last + + apply finite_valid_trace_init_add_last ; [assumption | reflexivity]. + intros prefix suffix Heq_tr. apply app_eq_nil in Heq_tr. destruct Heq_tr. subst. apply initial_state_witnessing_equivocation_prop. apply Htr. - rewrite finite_trace_last_is_last. intros Hm Hn Htr'_item Hwitness. - apply finite_protocol_trace_init_add_last + apply finite_valid_trace_init_add_last with (sf := destination item) in Htr'_item ; [|apply finite_trace_last_is_last]. @@ -560,17 +560,17 @@ Proof. spec IHn ; [subst s m; apply le_antisym; assumption|]. specialize (IHn eq_refl). destruct Htr'_item as [Htr'_item Hinit]. - apply finite_protocol_trace_from_to_app_split in Htr'_item. + apply finite_valid_trace_from_to_app_split in Htr'_item. destruct Htr'_item as [Htr' Hitem]. spec IHn. { split; [|assumption]. - apply finite_protocol_trace_from_to_forget_last in Htr'. + apply finite_valid_trace_from_to_forget_last in Htr'. assumption. } spec IHn Hwitness'. destruct IHn as [is' [tr'' [[Htr'' Hinit'] Hprefix]]]. specialize - (finite_protocol_trace_from_to_app PreFree _ _ _ _ _ Htr'' Hitem) + (finite_valid_trace_from_to_app PreFree _ _ _ _ _ Htr'' Hitem) as Htr''_item. eexists is', _. split; [exact (conj Htr''_item Hinit')|]. @@ -583,8 +583,8 @@ Proof. end. specialize (is_equivocating_tracewise_witness s) as Hwitness'. spec Hwitness'. - { apply proj1, finite_protocol_trace_from_to_app_split, proj1 - , finite_protocol_trace_from_to_last_pstate + { apply proj1, finite_valid_trace_from_to_app_split, proj1 + , finite_valid_trace_from_to_last_pstate in Htr'_item. subst. assumption. } @@ -600,21 +600,21 @@ Proof. rewrite <- set_add_length in Hlen1; assumption. } spec IHm is' tr''. - apply finite_protocol_trace_init_to_last in Htr'' as Htr''_lst. + apply finite_valid_trace_init_to_last in Htr'' as Htr''_lst. simpl in *. rewrite Htr''_lst in IHm. specialize (IHm eq_refl eq_refl). spec IHm. - { apply finite_protocol_trace_init_to_forget_last in Htr''. + { apply finite_valid_trace_init_to_forget_last in Htr''. assumption. } spec IHm Hwitness'. destruct IHm as [is'' [tr''' [[Htr''' Hinit'] Hprefix]]]. - apply proj1, finite_protocol_trace_from_to_app_split, proj2 in Htr'_item as Hitem. + apply proj1, finite_valid_trace_from_to_app_split, proj2 in Htr'_item as Hitem. simpl in *. rewrite <- Heqs in Hitem. specialize - (finite_protocol_trace_from_to_app PreFree _ _ _ _ _ Htr''' Hitem) + (finite_valid_trace_from_to_app PreFree _ _ _ _ _ Htr''' Hitem) as Htr'''_item. eexists is'', _. split; [exact (conj Htr'''_item Hinit')|]. @@ -632,21 +632,21 @@ Proof. Qed. (** A version of Lemma [preloaded_has_strong_trace_witnessing_equivocation_prop] -guaranteeing that for any [protocol_state] w.r.t. the Free composition there is -a trace ending in that state which is protocol w.r.t. the Free composition and +guaranteeing that for any [valid_state] w.r.t. the Free composition there is +a trace ending in that state which is valid w.r.t. the Free composition and it has the [strong_trace_witnessing_equivocation_prop]erty. *) Lemma free_has_strong_trace_witnessing_equivocation_prop s - (Hs : protocol_state_prop Free s) + (Hs : valid_state_prop Free s) : exists is' tr', - finite_protocol_trace_init_to Free is' s tr' /\ + finite_valid_trace_init_to Free is' s tr' /\ strong_trace_witnessing_equivocation_prop is' tr'. Proof. - apply (VLSM_incl_protocol_state (vlsm_incl_pre_loaded_with_all_messages_vlsm Free)) + apply (VLSM_incl_valid_state (vlsm_incl_pre_loaded_with_all_messages_vlsm Free)) in Hs as Hpre_s. apply preloaded_has_strong_trace_witnessing_equivocation_prop in Hpre_s. destruct Hpre_s as [is [tr [Htr Hwitness]]]. - apply (all_pre_traces_to_protocol_state_are_protocol IM _ Hbr finite_index) in Htr + apply (all_pre_traces_to_valid_state_are_valid IM _ Hbr finite_index) in Htr ; [|assumption]. exists is, tr. split; assumption. Qed. @@ -658,8 +658,8 @@ End witnessed_equivocation. The main result of this module is that, under witnessed equivocation assumptions, any trace with the [strong_trace_witnessing_equivocation_prop]erty -which is protocol for the free composition (guaranteed to exist by -Lemma [free_has_strong_trace_witnessing_equivocation_prop]) is also protocol +which is valid for the free composition (guaranteed to exist by +Lemma [free_has_strong_trace_witnessing_equivocation_prop]) is also valid for the composition constrained by the [fixed_equivocation_constrained] induced by the [equivocating_validators] of its final state. *) @@ -695,7 +695,7 @@ Context (Hsender_safety : sender_safety_alt_prop IM id sender := channel_authentication_sender_safety IM id sender can_emit_signed) (Free_has_sender := - composite_no_initial_protocol_messages_have_sender IM id sender + composite_no_initial_valid_messages_have_sender IM id sender can_emit_signed no_initial_messages_in_IM (free_constraint IM)) . @@ -716,20 +716,8 @@ Definition equivocating_validators_fixed_equivocation_constraint := fixed_equivocation_constraint IM Hbs Hbr (equivocating_validators s). -(** The [equivocating_validators_fixed_equivocation_constraint] induced by a -larger set of [equivocating_validators] is weaker. -*) -Lemma equivocating_validators_fixed_equivocation_subsumption - s s' - (Hincl : equivocating_validators s ⊆ equivocating_validators s') - : strong_constraint_subsumption IM (equivocating_validators_fixed_equivocation_constraint s) (equivocating_validators_fixed_equivocation_constraint s'). -Proof. - apply fixed_equivocation_constraint_index_incl_subsumption. - assumption. -Qed. - Lemma equivocators_can_emit_free m - (Hm : protocol_message_prop Free m) + (Hm : valid_message_prop Free m) v (Hsender: sender m = Some v) sf @@ -740,7 +728,7 @@ Lemma equivocators_can_emit_free m (equivocators_composition_for_observed IM Hbs Hbr (equivocating_validators sf) s) m. Proof. - apply can_emit_protocol_iff in Hm. + apply emitted_messages_are_valid_iff in Hm. destruct Hm as [[_v [[_im Him] Heqim]] | Hiom] ; [elim (no_initial_messages_in_IM _v _im); assumption|]. apply (VLSM_incl_can_emit (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))) @@ -771,30 +759,31 @@ Qed. (** *** Main result of the section -Any Free protocol trace with the -[strong_trace_witnessing_equivocation_prop]erty is also protocol w.r.t. the +Any Free valid trace with the +[strong_trace_witnessing_equivocation_prop]erty is also valid w.r.t. the composition using the [equivocating_validators_fixed_equivocation_constraint] induced by its final state. -The proof proceeds by induction on the protocol trace property. -Lemmas [equivocating_validators_witness_monotonicity] and -[equivocating_validators_fixed_equivocation_subsumption] are used to restate the +The proof proceeds by induction on the valid trace property. +Lemmas [equivocating_validators_witness_monotonicity] and +[fixed_equivocation_vlsm_composition_index_incl] are used to restate the induction hypothesis in terms of the final state after the last transition. *) Lemma strong_witness_has_fixed_equivocation is s tr - (Htr : finite_protocol_trace_init_to (free_composite_vlsm IM) is s tr) + (Htr : finite_valid_trace_init_to (free_composite_vlsm IM) is s tr) (Heqv: strong_trace_witnessing_equivocation_prop IM id sender finite_index is tr) - : finite_protocol_trace_from_to (fixed_equivocation_vlsm_composition IM Hbs Hbr (equivocating_validators s)) is s tr. + : finite_valid_trace_init_to (fixed_equivocation_vlsm_composition IM Hbs Hbr (equivocating_validators s)) is s tr. Proof. - induction Htr using finite_protocol_trace_init_to_rev_ind. - - eapply (finite_ptrace_from_to_empty (fixed_equivocation_vlsm_composition IM Hbs Hbr (equivocating_validators si))). - apply initial_is_protocol. assumption. + split; [|apply Htr]. + induction Htr using finite_valid_trace_init_to_rev_ind. + - eapply (finite_valid_trace_from_to_empty (fixed_equivocation_vlsm_composition IM Hbs Hbr (equivocating_validators si))). + apply initial_state_is_valid. assumption. - spec IHHtr. { intros prefix. intros. apply (Heqv prefix (suffix ++ [{| l := l; input := iom; destination := sf; output := oom |}])). subst. apply app_assoc. } - apply (VLSM_incl_finite_protocol_trace_init_to (vlsm_incl_pre_loaded_with_all_messages_vlsm Free)) + apply (VLSM_incl_finite_valid_trace_init_to (vlsm_incl_pre_loaded_with_all_messages_vlsm Free)) in Htr as Hpre_tr. specialize (equivocating_validators_witness_monotonicity IM id sender finite_index @@ -809,20 +798,17 @@ Proof. assumption. } assert - (Htr_sf : finite_protocol_trace_from_to + (Htr_sf : finite_valid_trace_from_to (fixed_equivocation_vlsm_composition IM Hbs Hbr (equivocating_validators sf)) si s tr). { revert IHHtr. - apply VLSM_incl_finite_protocol_trace_from_to. - apply constraint_subsumption_incl. - apply preloaded_constraint_subsumption_stronger. - apply strong_constraint_subsumption_strongest. - apply equivocating_validators_fixed_equivocation_subsumption. + apply VLSM_incl_finite_valid_trace_from_to. + apply fixed_equivocation_vlsm_composition_index_incl. assumption. } clear IHHtr. apply (extend_right_finite_trace_from_to _ Htr_sf). destruct Ht as [[Hs [Hiom [Hv _]]] Ht]. - apply finite_protocol_trace_from_to_last_pstate in Htr_sf as Hs'. + apply finite_valid_trace_from_to_last_pstate in Htr_sf as Hs'. specialize (Heqv (tr ++ [{| l := l; input := iom; destination := sf; output := oom |}]) @@ -831,7 +817,7 @@ Proof. destruct iom as [im|]. 2:{ repeat split - ; [assumption| apply option_protocol_message_None | assumption..]. + ; [assumption| apply option_valid_message_None | assumption..]. } apply Free_has_sender in Hiom as _Hsender. destruct (sender im) as [v|] eqn:Hsender; [|congruence]. @@ -839,12 +825,12 @@ Proof. spec Heqv v. rewrite finite_trace_last_is_last in Heqv. simpl in Heqv. - assert (Hpre_s : protocol_state_prop (pre_loaded_with_all_messages_vlsm Free) s). - { apply proj1, finite_protocol_trace_from_to_last_pstate in Hpre_tr. assumption. } + assert (Hpre_s : valid_state_prop (pre_loaded_with_all_messages_vlsm Free) s). + { apply proj1, finite_valid_trace_from_to_last_pstate in Hpre_tr. assumption. } destruct (@decide _ (composite_has_been_observed_dec IM finite_index Hbo s im)). { repeat split - ; [assumption| apply option_protocol_message_Some | assumption| | assumption]. - - apply (composite_observed_protocol IM finite_index Hbs Hbo Hbr _ s); assumption. + ; [assumption| apply option_valid_message_Some | assumption| | assumption]. + - apply (composite_observed_valid IM finite_index Hbs Hbo Hbr _ s); assumption. - left. assumption. } assert (Hequivocating_v : v ∈ equivocating_validators sf). @@ -864,7 +850,7 @@ Proof. specialize (equivocators_can_emit_free _ Hiom _ Hsender _ Hequivocating_v _ _ Hv) as Hemit_im. repeat split ; [assumption| | assumption| right; assumption | assumption]. - apply can_emit_protocol. + apply emitted_messages_are_valid. specialize (EquivPreloadedBase_Fixed_weak_full_projection IM Hbs Hbr _ finite_index _ Hs') as Hproj. spec Hproj. @@ -875,22 +861,22 @@ Proof. Qed. (** -As a corollary of the above, every protocol state for the free composition is -also a protocol state for the composition with the +As a corollary of the above, every valid state for the free composition is +also a valid state for the composition with the [equivocating_validators_fixed_equivocation_constraint] induced by it. *) Lemma equivocating_validators_fixed_equivocation_characterization : forall s, - protocol_state_prop Free s -> - protocol_state_prop + valid_state_prop Free s -> + valid_state_prop (composite_vlsm IM (equivocating_validators_fixed_equivocation_constraint s)) s. Proof. intros s Hs. destruct (free_has_strong_trace_witnessing_equivocation_prop IM finite_index Hbs Hbr id sender finite_index _ s Hs) as [is [tr [Htr Heqv]]]. - cut (finite_protocol_trace_from_to (composite_vlsm IM (equivocating_validators_fixed_equivocation_constraint s)) is s tr). - { intro Htr'. apply finite_protocol_trace_from_to_last_pstate in Htr'. + cut (finite_valid_trace_from_to (composite_vlsm IM (equivocating_validators_fixed_equivocation_constraint s)) is s tr). + { intro Htr'. apply finite_valid_trace_from_to_last_pstate in Htr'. assumption. } clear Hs. diff --git a/theories/VLSM/Core/EquivocationProjections.v b/theories/VLSM/Core/EquivocationProjections.v index d20bf7a3..7d810ce4 100644 --- a/theories/VLSM/Core/EquivocationProjections.v +++ b/theories/VLSM/Core/EquivocationProjections.v @@ -41,14 +41,14 @@ Lemma VLSM_projection_oracle_reflect (oracleY : vstate Y -> message -> Prop) (HstepwiseX : oracle_stepwise_props (vlsm := X) selectorX oracleX) (HstepwiseY : oracle_stepwise_props (vlsm := Y) selectorY oracleY) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, oracleY (state_project s) m -> oracleX s m. Proof. intros s Hs m Hm. apply (prove_all_have_message_from_stepwise _ _ _ _ HstepwiseX _ Hs m). intros isX trX HtrX. - apply (VLSM_projection_finite_protocol_trace_init_to Hsimul) in HtrX. - apply (VLSM_projection_protocol_state Hsimul) in Hs as HsY. + apply (VLSM_projection_finite_valid_trace_init_to Hsimul) in HtrX. + apply (VLSM_projection_valid_state Hsimul) in Hs as HsY. apply (prove_all_have_message_from_stepwise _ _ _ _ HstepwiseY _ HsY m) in Hm. specialize (Hm _ _ HtrX). apply Exists_exists in Hm as [itemY [HitemY Hm]]. @@ -70,7 +70,7 @@ End selectors. Lemma VLSM_projection_has_been_sent_reflect {HbsX : HasBeenSentCapability X} {HbsY : HasBeenSentCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_sent Y (state_project s) m -> has_been_sent X s m. Proof. apply VLSM_projection_oracle_reflect with (field_selector output) (field_selector output). @@ -83,7 +83,7 @@ Qed. Lemma VLSM_projection_has_been_received_reflect {HbrX : HasBeenReceivedCapability X} {HbrY : HasBeenReceivedCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_received Y (state_project s) m -> has_been_received X s m. Proof. apply VLSM_projection_oracle_reflect with (field_selector input) (field_selector input). @@ -100,7 +100,7 @@ Lemma VLSM_projection_has_been_observed_reflect {HbsY : HasBeenSentCapability Y} {HbrY : HasBeenReceivedCapability Y} (HboY := HasBeenObservedCapability_from_sent_received Y) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_observed Y (state_project s) m -> has_been_observed X s m. Proof. apply VLSM_projection_oracle_reflect with item_sends_or_receives item_sends_or_receives. @@ -141,13 +141,13 @@ Lemma VLSM_weak_full_projection_selected_message_exists_in_some_preloaded_traces Proof. intros [is [tr [Htr Hm]]]. destruct Htr as [Htr His]. - apply (VLSM_weak_full_projection_finite_protocol_trace_from_to Hsimul) in Htr. - apply ptrace_first_pstate in Htr as Hpr_is. - apply protocol_state_has_trace in Hpr_is as [is' [tr_is Hpr_is]]. + apply (VLSM_weak_full_projection_finite_valid_trace_from_to Hsimul) in Htr. + apply valid_trace_first_pstate in Htr as Hpr_is. + apply valid_state_has_trace in Hpr_is as [is' [tr_is Hpr_is]]. exists is', (tr_is ++ (VLSM_weak_full_projection_finite_trace_project Hsimul tr)). destruct Hpr_is as [Hpr_is His']. repeat split; [|assumption|]. - - apply (finite_protocol_trace_from_to_app (pre_loaded_with_all_messages_vlsm Y)) + - apply (finite_valid_trace_from_to_app (pre_loaded_with_all_messages_vlsm Y)) with (state_project is) ; assumption. - apply Exists_app. right. @@ -168,13 +168,13 @@ Lemma VLSM_weak_full_projection_oracle (HstepwiseY : oracle_stepwise_props (vlsm := Y) selectorY oracleY) (HoracleX_dec : RelDecision oracleX) (HoracleY_dec : RelDecision oracleY) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, oracleX s m -> oracleY (state_project s) m. Proof. intros s Hs m Hm. apply (prove_all_have_message_from_stepwise _ _ _ _ HstepwiseX _ Hs m) in Hm. apply (selected_messages_consistency_prop_from_stepwise _ _ _ _ HstepwiseX HoracleX_dec _ Hs) in Hm. - apply (VLSM_weak_full_projection_protocol_state Hsimul) in Hs as HsY. + apply (VLSM_weak_full_projection_valid_state Hsimul) in Hs as HsY. apply (prove_all_have_message_from_stepwise _ _ _ _ HstepwiseY _ HsY m). apply (selected_messages_consistency_prop_from_stepwise _ _ _ _ HstepwiseY HoracleY_dec _ HsY). revert Hm. @@ -186,7 +186,7 @@ End selectors. Lemma VLSM_weak_full_projection_has_been_sent {HbsX : HasBeenSentCapability X} {HbsY : HasBeenSentCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_sent X s m -> has_been_sent Y (state_project s) m. Proof. apply VLSM_weak_full_projection_oracle with (field_selector output) (field_selector output). @@ -201,7 +201,7 @@ Qed. Lemma VLSM_weak_full_projection_has_been_received {HbrX : HasBeenReceivedCapability X} {HbrY : HasBeenReceivedCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_received X s m -> has_been_received Y (state_project s) m. Proof. apply VLSM_weak_full_projection_oracle with (field_selector input) (field_selector input). @@ -220,7 +220,7 @@ Lemma VLSM_weak_full_projection_has_been_observed {HbsY : HasBeenSentCapability Y} {HbrY : HasBeenReceivedCapability Y} (HboY := HasBeenObservedCapability_from_sent_received Y) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_observed X s m -> has_been_observed Y (state_project s) m. Proof. apply VLSM_weak_full_projection_oracle with item_sends_or_receives item_sends_or_receives. @@ -248,14 +248,14 @@ Context Definition VLSM_full_projection_has_been_sent {HbsX : HasBeenSentCapability X} {HbsY : HasBeenSentCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_sent X s m -> has_been_sent Y (state_project s) m := VLSM_weak_full_projection_has_been_sent (VLSM_full_projection_weaken Hsimul). Definition VLSM_full_projection_has_been_received {HbsX : HasBeenReceivedCapability X} {HbsY : HasBeenReceivedCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_received X s m -> has_been_received Y (state_project s) m := VLSM_weak_full_projection_has_been_received (VLSM_full_projection_weaken Hsimul). @@ -266,21 +266,21 @@ Definition VLSM_full_projection_has_been_observed {HbsY : HasBeenSentCapability Y} {HbrY : HasBeenReceivedCapability Y} (HboY := HasBeenObservedCapability_from_sent_received Y) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_observed X s m -> has_been_observed Y (state_project s) m := VLSM_weak_full_projection_has_been_observed (VLSM_full_projection_weaken Hsimul). Definition VLSM_full_projection_has_been_sent_reflect {HbsX : HasBeenSentCapability X} {HbsY : HasBeenSentCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_sent Y (state_project s) m -> has_been_sent X s m := VLSM_projection_has_been_sent_reflect (VLSM_full_projection_is_projection Hsimul). Definition VLSM_full_projection_has_been_received_reflect {HbrX : HasBeenReceivedCapability X} {HbrY : HasBeenReceivedCapability Y} - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_received Y (state_project s) m -> has_been_received X s m := VLSM_projection_has_been_received_reflect (VLSM_full_projection_is_projection Hsimul). @@ -291,7 +291,7 @@ Definition VLSM_full_projection_has_been_observed_reflect {HbsY : HasBeenSentCapability Y} {HbrY : HasBeenReceivedCapability Y} (HboY := HasBeenObservedCapability_from_sent_received Y) - : forall s, protocol_state_prop (pre_loaded_with_all_messages_vlsm X) s -> + : forall s, valid_state_prop (pre_loaded_with_all_messages_vlsm X) s -> forall m, has_been_observed Y (state_project s) m -> has_been_observed X s m := VLSM_projection_has_been_observed_reflect (VLSM_full_projection_is_projection Hsimul). @@ -318,7 +318,7 @@ Existing Instance HbsFree2. [has_been_sent] predicate is preserved through the [same_IM_state_rew] map. *) Lemma same_IM_composite_has_been_sent_preservation s1 m - (Hs1 : protocol_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM1)) s1) + (Hs1 : valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM1)) s1) : composite_has_been_sent IM1 Hbs1 s1 m -> composite_has_been_sent IM2 Hbs2 (same_IM_state_rew Heq s1) m. Proof. @@ -360,7 +360,7 @@ Proof. specialize (Hsender_safety _ _ Hsender). intros [(s0,om0) [(i, li) [s1 Hemitted]]]. specialize (preloaded_component_projection IM i) as Hproj. - specialize (VLSM_projection_protocol_transition Hproj (existT i li) li) + specialize (VLSM_projection_input_valid_transition Hproj (existT i li) li) as Htransition. spec Htransition. { clear. unfold composite_project_label. simpl. diff --git a/theories/VLSM/Core/Equivocators/Common.v b/theories/VLSM/Core/Equivocators/Common.v index 7cab1acd..109d546a 100644 --- a/theories/VLSM/Core/Equivocators/Common.v +++ b/theories/VLSM/Core/Equivocators/Common.v @@ -10,9 +10,11 @@ regular machine X, and then, at any moment: - can perform [valid] [transition]s on any of its internal machines - can fork any of its internal machines by duplicating its state and then using the given label and message to [transition] on the new fork. + + Note that we only allow forking if a transition is then taken on the neq fork. *) -Section equivocator_vlsm. +Section sec_equivocator_vlsm. Context {message : Type} (X : VLSM message) @@ -451,7 +453,7 @@ Proof. unfold equivocator_state_n, equivocator_state_append, equivocator_state_last in *. simpl in *. lia. -Qed. +Qed. Local Ltac destruct_equivocator_state_append_project' es es' i Hi k Hk Hpr := let Hi' := fresh "Hi" in @@ -535,7 +537,7 @@ Proof. lia. - destruct (decide (n = k)). + subst. rewrite equivocator_state_update_size in Hi. - rewrite equivocator_state_update_project_eq; + rewrite equivocator_state_update_project_eq; [|rewrite equivocator_state_append_size; lia | reflexivity ]. @@ -661,7 +663,7 @@ Lemma mk_singleton_initial_state split;[reflexivity|assumption]. Qed. -End equivocator_vlsm. +End sec_equivocator_vlsm. Arguments Spawn {_ _} _: assert. Arguments ContinueWith {_ _} _ _: assert. @@ -718,7 +720,7 @@ Ltac destruct_equivocator_state_extend_project es s i Hi := destruct_equivocator_state_extend_project' es s i Hi Hpr ; clear Hpr. -Section equivocator_vlsm_protocol_state_projections. +Section equivocator_vlsm_valid_state_projections. Context {message : Type} @@ -956,41 +958,41 @@ Proof. Qed. (** -Protocol messages in the [equivocator_vlsm] are also protocol in the -original machine. All components of a protocol state in the -[equivocator_vlsm] are also protocol in the original machine. +Valid messages in the [equivocator_vlsm] are also valid in the +original machine. All components of a valid state in the +[equivocator_vlsm] are also valid in the original machine. *) -Lemma preloaded_equivocator_state_project_protocol +Lemma preloaded_equivocator_state_projection_preserves_validity (seed : message -> Prop) (bs : vstate equivocator_vlsm) (om : option message) - (Hbs : protocol_prop (pre_loaded_vlsm equivocator_vlsm seed) bs om) - : option_protocol_message_prop (pre_loaded_vlsm X seed) om /\ + (Hbs : valid_state_message_prop (pre_loaded_vlsm equivocator_vlsm seed) bs om) + : option_valid_message_prop (pre_loaded_vlsm X seed) om /\ forall i si, equivocator_state_project bs i = Some si -> - protocol_state_prop (pre_loaded_vlsm X seed) si. + valid_state_prop (pre_loaded_vlsm X seed) si. Proof. induction Hbs. - - split; [apply option_initial_message_is_protocol;assumption|]. + - split; [apply option_initial_message_is_valid;assumption|]. intros. destruct Hs as [Hn0 Hinit]. apply equivocator_state_project_Some_rev in H as Hi. unfold is_singleton_state in Hn0. assert (i = 0) by lia. subst. rewrite equivocator_state_project_zero in H. - inversion H. apply initial_is_protocol. assumption. - - specialize (protocol_generated (pre_loaded_vlsm X seed)) as Hgen. + inversion H. apply initial_state_is_valid. assumption. + - specialize (valid_generated_state_message (pre_loaded_vlsm X seed)) as Hgen. apply proj2 in IHHbs1. apply proj1 in IHHbs2. (* destruction tactic for a valid equivocator transition vtransition equivocator_vlsm l (s, om) = (s', om') *) destruct l as [sn|i l|i l] - ; [ inversion_clear Ht; split; [apply option_protocol_message_None|] + ; [ inversion_clear Ht; split; [apply option_valid_message_None|] ; intros ; destruct_equivocator_state_extend_project s sn i Hi ; [ apply IHHbs1 in H; assumption - | inversion H; subst; apply initial_is_protocol; apply Hv + | inversion H; subst; apply initial_state_is_valid; apply Hv | discriminate] |..] ; cbn in Hv, Ht @@ -1014,77 +1016,77 @@ Proof. * discriminate. Qed. -Lemma preloaded_with_equivocator_state_project_protocol_state +Lemma preloaded_with_equivocator_state_project_valid_state (seed : message -> Prop) (bs : vstate equivocator_vlsm) - (Hbs : protocol_state_prop (pre_loaded_vlsm equivocator_vlsm seed) bs) + (Hbs : valid_state_prop (pre_loaded_vlsm equivocator_vlsm seed) bs) : forall i si, equivocator_state_project bs i = Some si -> - protocol_state_prop (pre_loaded_vlsm X seed) si. + valid_state_prop (pre_loaded_vlsm X seed) si. Proof. destruct Hbs as [om Hbs]. - apply preloaded_equivocator_state_project_protocol, proj2 in Hbs. + apply preloaded_equivocator_state_projection_preserves_validity, proj2 in Hbs. assumption. Qed. -Lemma preloaded_with_equivocator_state_project_protocol_message +Lemma preloaded_with_equivocator_state_project_valid_message (seed : message -> Prop) (om : option message) - (Hom : option_protocol_message_prop (pre_loaded_vlsm equivocator_vlsm seed) om) + (Hom : option_valid_message_prop (pre_loaded_vlsm equivocator_vlsm seed) om) : - option_protocol_message_prop (pre_loaded_vlsm X seed) om. + option_valid_message_prop (pre_loaded_vlsm X seed) om. Proof. destruct Hom as [s Hm]. - apply preloaded_equivocator_state_project_protocol, proj1 in Hm. + apply preloaded_equivocator_state_projection_preserves_validity, proj1 in Hm. assumption. Qed. -Lemma equivocator_state_project_protocol_state +Lemma equivocator_state_project_valid_state (bs : vstate equivocator_vlsm) - (Hbs : protocol_state_prop equivocator_vlsm bs) + (Hbs : valid_state_prop equivocator_vlsm bs) : forall i si, - equivocator_state_project bs i = Some si -> protocol_state_prop X si. + equivocator_state_project bs i = Some si -> valid_state_prop X si. Proof. intros i si Hpr. - apply (VLSM_eq_protocol_state (vlsm_is_pre_loaded_with_False equivocator_vlsm)) in Hbs. - specialize (preloaded_with_equivocator_state_project_protocol_state _ _ Hbs _ _ Hpr) as Hsi. - apply (VLSM_eq_protocol_state (vlsm_is_pre_loaded_with_False X)) in Hsi. + apply (VLSM_eq_valid_state (vlsm_is_pre_loaded_with_False equivocator_vlsm)) in Hbs. + specialize (preloaded_with_equivocator_state_project_valid_state _ _ Hbs _ _ Hpr) as Hsi. + apply (VLSM_eq_valid_state (vlsm_is_pre_loaded_with_False X)) in Hsi. destruct X as (T, (S, M)). assumption. Qed. -Lemma equivocator_state_project_protocol_message +Lemma equivocator_state_project_valid_message (om : option message) - (Hom : option_protocol_message_prop equivocator_vlsm om) + (Hom : option_valid_message_prop equivocator_vlsm om) : - option_protocol_message_prop X om. + option_valid_message_prop X om. Proof. - destruct om as [m|]; [|apply option_protocol_message_None]. + destruct om as [m|]; [|apply option_valid_message_None]. specialize (vlsm_is_pre_loaded_with_False_initial_message equivocator_vlsm) as Hinit. - apply (VLSM_incl_protocol_message (VLSM_eq_proj1 (vlsm_is_pre_loaded_with_False equivocator_vlsm))) in Hom + apply (VLSM_incl_valid_message (VLSM_eq_proj1 (vlsm_is_pre_loaded_with_False equivocator_vlsm))) in Hom ; [| assumption]. - apply preloaded_with_equivocator_state_project_protocol_message in Hom. + apply preloaded_with_equivocator_state_project_valid_message in Hom. specialize (vlsm_is_pre_loaded_with_False_initial_message_rev X) as Hinit_rev. - apply (VLSM_incl_protocol_message (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Hom + apply (VLSM_incl_valid_message (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Hom ; destruct X as (T, (S, M)) ; assumption. Qed. (** -All components of a protocol states of the [pre_loaded_with_all_messages_vlsm] corresponding -to an [equivocator_vlsm] are also protocol for the [pre_loaded_with_all_messages_vlsm] +All components of valid states of the [pre_loaded_with_all_messages_vlsm] corresponding +to an [equivocator_vlsm] are also valid for the [pre_loaded_with_all_messages_vlsm] corresponding to the original machine. *) -Lemma preloaded_equivocator_state_project_protocol_state +Lemma preloaded_equivocator_state_project_valid_state (bs : vstate equivocator_vlsm) - (Hbs : protocol_state_prop (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs) + (Hbs : valid_state_prop (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs) : forall i si, - equivocator_state_project bs i = Some si -> protocol_state_prop (pre_loaded_with_all_messages_vlsm X) si. + equivocator_state_project bs i = Some si -> valid_state_prop (pre_loaded_with_all_messages_vlsm X) si. Proof. intros i si Hpr. - apply (VLSM_eq_protocol_state (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True equivocator_vlsm)) in Hbs. - specialize (preloaded_with_equivocator_state_project_protocol_state _ _ Hbs _ _ Hpr) as Hsi. - apply (VLSM_eq_protocol_state (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X)) in Hsi. + apply (VLSM_eq_valid_state (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True equivocator_vlsm)) in Hbs. + specialize (preloaded_with_equivocator_state_project_valid_state _ _ Hbs _ _ Hpr) as Hsi. + apply (VLSM_eq_valid_state (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X)) in Hsi. destruct X as (T, (S, M)). assumption. Qed. @@ -1237,7 +1239,7 @@ Proof. apply equivocator_state_project_Some_rev in Hsi. assumption. Qed. -End equivocator_vlsm_protocol_state_projections. +End equivocator_vlsm_valid_state_projections. Arguments NewMachine {_ _} _: assert. Arguments Existing {_ _} _ : assert. diff --git a/theories/VLSM/Core/Equivocators/Composition/Common.v b/theories/VLSM/Core/Equivocators/Composition/Common.v index 48f4a047..f24c8a8b 100644 --- a/theories/VLSM/Core/Equivocators/Composition/Common.v +++ b/theories/VLSM/Core/Equivocators/Composition/Common.v @@ -189,7 +189,7 @@ Proof. apply equivocator_transition_cannot_decrease_state_size in Htj. inversion Ht. subst. clear Ht. intro eqv. - destruct (decide (j = eqv)). + destruct (decide (j = eqv)). - subst. rewrite state_update_eq. assumption. - rewrite state_update_neq by congruence. lia. Qed. @@ -221,7 +221,7 @@ Qed. Lemma equivocators_pre_trace_cannot_decrease_state_size (Pre := pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) s s' tr - (Htr : finite_protocol_trace_from_to Pre s s' tr) + (Htr : finite_valid_trace_from_to Pre s s' tr) : forall eqv, equivocator_state_n (s eqv) <= equivocator_state_n (s' eqv). Proof. apply trace_to_plan_to_trace_from_to in Htr. @@ -234,7 +234,7 @@ Qed. Lemma equivocators_pre_trace_preserves_equivocating_state (Pre := pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) s s' tr - (Htr : finite_protocol_trace_from_to Pre s s' tr) + (Htr : finite_valid_trace_from_to Pre s s' tr) : forall eqv, is_equivocating_state (IM eqv) (s eqv) -> is_equivocating_state (IM eqv) (s' eqv). Proof. unfold is_equivocating_state, is_singleton_state. @@ -266,7 +266,7 @@ Lemma equivocators_no_equivocations_vlsm_incl_equivocators_free : VLSM_incl equivocators_no_equivocations_vlsm equivocators_free_vlsm. Proof. apply basic_VLSM_incl; intro; intros; [assumption|..]. - - apply initial_message_is_protocol; assumption. + - apply initial_message_is_valid; assumption. - split; [|exact I]. apply Hv. - apply H. Qed. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v index 45ba696b..185bb3c1 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v @@ -117,12 +117,12 @@ Definition equivocators_fixed_equivocations_vlsm Lemma not_equivocating_index_has_singleton_state (s : composite_state equivocator_IM) - (Hs : protocol_state_prop equivocators_fixed_equivocations_vlsm s) + (Hs : valid_state_prop equivocators_fixed_equivocations_vlsm s) (i : index) (Hi : i ∉ equivocating) : is_singleton_state (IM i) (s i). Proof. - apply protocol_state_has_trace in Hs. + apply valid_state_has_trace in Hs. destruct Hs as [is [tr [Htr His]]]. assert (is_singleton_state (IM i) (is i)) by (apply His). @@ -140,12 +140,12 @@ Proof. assumption. Qed. -Lemma protocol_has_fixed_equivocation +Lemma valid_state_has_fixed_equivocation (s : composite_state equivocator_IM) - (Hs : protocol_state_prop equivocators_fixed_equivocations_vlsm s) + (Hs : valid_state_prop equivocators_fixed_equivocations_vlsm s) : state_has_fixed_equivocation s. Proof. - apply protocol_state_has_trace in Hs. + apply valid_state_has_trace in Hs. destruct Hs as [is [tr [Htr Hinit]]]. assert (state_has_fixed_equivocation is) as His. { intros i Hin. @@ -299,9 +299,9 @@ Section from_equivocators_to_nodes. (** ** From composition of equivocators to composition of simple nodes -In this section we show that the projection of [protocol_trace]s of a +In this section we show that the projection of [valid_trace]s of a composition of equivocators with a fixed equivocation constraint are -[protocol_trace]s for the composition of nodes with a similar fixed +[valid_trace]s for the composition of nodes with a similar fixed equivocation constraint. *) @@ -354,7 +354,7 @@ Definition proper_fixed_equivocator_descriptors *) Lemma not_equivocating_equivocator_descriptors_proper_fixed (s : composite_state (equivocator_IM IM)) - (Hs : protocol_state_prop XE s) + (Hs : valid_state_prop XE s) (eqv_descriptors : equivocator_descriptors IM) (Heqv_descriptors : not_equivocating_equivocator_descriptors IM eqv_descriptors s) : proper_fixed_equivocator_descriptors eqv_descriptors s. @@ -379,7 +379,7 @@ Projections of (valid) traces of the composition of equivocators preserve Lemma equivocators_trace_project_preserves_proper_fixed_equivocator_descriptors (is : composite_state (equivocator_IM IM)) (tr : list (composite_transition_item (equivocator_IM IM))) - (Htr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm FreeE) is tr) + (Htr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm FreeE) is tr) (s := finite_trace_last is tr) (descriptors : equivocator_descriptors IM) (idescriptors : equivocator_descriptors IM) @@ -389,7 +389,7 @@ Lemma equivocators_trace_project_preserves_proper_fixed_equivocator_descriptors Proof. intros [Hproper Hfixed]. specialize - (preloaded_equivocators_protocol_trace_project_proper_initial IM + (preloaded_equivocators_valid_trace_project_proper_initial IM _ _ Htr _ _ _ HtrX Hproper ) as Hiproper. @@ -425,7 +425,7 @@ Existing Instance Free_HasBeenSentCapability. (** This is a property of the [fixed_equivocation_constraint] which also trivially holds for the free constraint. This property is sufficient for -proving the [_equivocators_protocol_trace_project] lemma, +proving the [_equivocators_valid_trace_project] lemma, which lets that lemma be used for both the composition of equivocators with fixed state-equivocation and the free composition. @@ -438,7 +438,7 @@ Definition constraint_has_been_sent_prop : Prop := forall (s : composite_state (equivocator_IM IM)) - (Hs : protocol_state_prop XE s) + (Hs : valid_state_prop XE s) (descriptors : equivocator_descriptors IM) (Hdescriptors : proper_fixed_equivocator_descriptors descriptors s) (sX := @equivocators_state_project _ _ IndEqDec IM descriptors s) @@ -449,7 +449,7 @@ Definition constraint_has_been_sent_prop (** Generic proof that the projection of a trace of the composition of equivocators -with no message equivocation and fixed state equivocation is protocol w.r.t. +with no message equivocation and fixed state equivocation is valid w.r.t. the composition of the regular nodes constrained by any constraint satisfying several properties, including the [constraint_has_been_sent_prop]erty. @@ -458,20 +458,20 @@ performing an analysis on the final transition item of the trace. It uses the fact that the trace hase no message equivocation to extract a subtrace producing the message being received at the last transition and -proves that it's a protocol message for the destination machine by using +proves that it's a valid message for the destination machine by using the induction hypothesis (that is why well-founded induction is used rather than a simpler induction principle). The constraint satisfaction for the projection of the final transition is -for now assumes as hypothesis @Hconstraint_hbs@. +for now assumes as hypothesis <>. *) -Lemma _equivocators_protocol_trace_project +Lemma _equivocators_valid_trace_project (final_descriptors : equivocator_descriptors IM) (is : composite_state (equivocator_IM IM)) (tr : list (composite_transition_item (equivocator_IM IM))) (final_state := finite_trace_last is tr) (Hproper : proper_fixed_equivocator_descriptors final_descriptors final_state) - (Htr : finite_protocol_trace XE is tr) + (Htr : finite_valid_trace XE is tr) (constraint : composite_label IM -> composite_state IM * option message -> Prop) (HconstraintNone : forall l s, constraint l (s, None)) (Hconstraint_hbs : constraint_has_been_sent_prop constraint) @@ -484,7 +484,7 @@ Lemma _equivocators_protocol_trace_project proper_fixed_equivocator_descriptors initial_descriptors is /\ equivocators_trace_project IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project IM final_descriptors final_state = final_stateX /\ - finite_protocol_trace X' isX trX. + finite_valid_trace X' isX trX. Proof. remember (length tr) as len_tr. generalize dependent final_descriptors. generalize dependent tr. @@ -496,7 +496,7 @@ Proof. remember (equivocators_state_project IM final_descriptors is) as isx. cut (vinitial_state_prop X' isx). { intro. split; [|assumption]. constructor. - apply protocol_state_prop_iff. left. + apply valid_state_prop_iff. left. exists (exist _ _ H). reflexivity. } subst. @@ -505,7 +505,7 @@ Proof. - specialize (H (length tr')) as H'. spec H'. { rewrite app_length. simpl. lia. } destruct Htr as [Htr Hinit]. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hlst]. specialize (H' tr' (conj Htr Hinit) eq_refl). specialize (equivocators_transition_item_project_proper_characterization IM final_descriptors lst) as Hproperx. @@ -538,16 +538,16 @@ Proof. specialize (H' _ Hproper'). destruct H' as [trX' [initial_descriptors [_ [Htr_project [Hstate_project HtrX']]]]]. assert - (Htr' : finite_protocol_trace FreeE is tr'). + (Htr' : finite_valid_trace FreeE is tr'). { split; [|assumption]. - revert Htr. apply VLSM_incl_finite_protocol_trace_from. + revert Htr. apply VLSM_incl_finite_valid_trace_from. apply equivocators_fixed_equivocations_vlsm_incl_free. } assert - (Htr'pre : finite_protocol_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr'). + (Htr'pre : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr'). { split; [|assumption]. specialize (vlsm_incl_pre_loaded_with_all_messages_vlsm FreeE) as Hincl. - apply (VLSM_incl_finite_protocol_trace_from Hincl). apply Htr'. + apply (VLSM_incl_finite_valid_trace_from Hincl). apply Htr'. } specialize (equivocators_trace_project_preserves_proper_fixed_equivocator_descriptors _ _ (proj1 Htr'pre) _ _ _ Htr_project Hproper') @@ -567,15 +567,15 @@ Proof. split; [assumption|]. destruct HtrX' as [HtrX' His]. split; [|assumption]. - apply finite_protocol_trace_from_app_iff. + apply finite_valid_trace_from_app_iff. split; [assumption|]. change [item] with ([] ++ [item]). match goal with - |- finite_protocol_trace_from _ ?l _ => remember l as lst + |- finite_valid_trace_from _ ?l _ => remember l as lst end. destruct item. - assert (Hplst : protocol_state_prop X' lst). - { apply finite_ptrace_last_pstate in HtrX'. subst. assumption. } + assert (Hplst : valid_state_prop X' lst). + { apply finite_valid_trace_last_pstate in HtrX'. subst. assumption. } apply (extend_right_finite_trace_from X'); [constructor; assumption|]. simpl in Hl. subst. simpl in Hc. @@ -583,15 +583,15 @@ Proof. simpl in Htx,Hvx,Hstate_project. rewrite Hstate_project in Hvx, Htx. destruct input as [input|] - ; [|repeat split; [assumption| apply option_protocol_message_None |assumption| apply HconstraintNone |assumption]]. + ; [|repeat split; [assumption| apply option_valid_message_None |assumption| apply HconstraintNone |assumption]]. simpl in Hno_equiv. apply or_comm in Hno_equiv. destruct Hno_equiv as [Hinit_input | Hno_equiv] ; [contradiction|]. assert - (Hs_free : protocol_state_prop FreeE (finite_trace_last is tr')). + (Hs_free : valid_state_prop FreeE (finite_trace_last is tr')). { destruct Hs as [_om Hs]. - apply (constraint_subsumption_protocol_prop (equivocator_IM IM)) + apply (constraint_subsumption_valid_state_message_preservation (equivocator_IM IM)) with (constraint2 := free_constraint (equivocator_IM IM)) in Hs as Hs_free ; [|intro x; intros; exact I]. @@ -599,12 +599,12 @@ Proof. } specialize (specialized_proper_sent_rev FreeE _ Hs_free _ Hno_equiv) as Hall. - specialize (Hall is tr' (ptrace_add_default_last Htr')). + specialize (Hall is tr' (valid_trace_add_default_last Htr')). destruct (equivocators_trace_project_output_reflecting_inv IM _ _ (proj1 Htr'pre) _ Hall) as [final_descriptors_m [initial_descriptors_m [trXm [_Hfinal_descriptors_m [Hproject_trXm Hex]]]]]. assert (Hfinal_descriptors_m : proper_fixed_equivocator_descriptors final_descriptors_m (finite_trace_last is tr')). { apply not_equivocating_equivocator_descriptors_proper_fixed; [|assumption]. - apply finite_ptrace_last_pstate. assumption. + apply finite_valid_trace_last_pstate. assumption. } specialize (H (length tr')). spec H. { rewrite app_length. simpl. lia. } @@ -613,16 +613,16 @@ Proof. simpl in *. rewrite Hproject_trXm in Hproject_trXm'. inversion Hproject_trXm'. subst trXm' initial_descriptors_m'. clear Hproject_trXm'. repeat split; [assumption| |assumption| |assumption] - ; [ apply option_protocol_message_Some - ; apply (protocol_trace_output_is_protocol X' _ _ (proj1 HtrXm) _ Hex) + ; [ apply option_valid_message_Some + ; apply (valid_trace_output_is_valid X' _ _ (proj1 HtrXm) _ Hex) |]. rewrite <- Hstate_project. apply Hconstraint_hbs; [assumption| apply Hproper'|]. - assert (Hlst'pre : protocol_state_prop (pre_loaded_with_all_messages_vlsm FreeE) (finite_trace_last is tr')). - { apply finite_ptrace_last_pstate. apply Htr'pre. } + assert (Hlst'pre : valid_state_prop (pre_loaded_with_all_messages_vlsm FreeE) (finite_trace_last is tr')). + { apply finite_valid_trace_last_pstate. apply Htr'pre. } apply proper_sent; [assumption|]. apply has_been_sent_consistency; [assumption| assumption| ]. - exists is, tr', (ptrace_add_default_last Htr'pre). assumption. + exists is, tr', (valid_trace_add_default_last Htr'pre). assumption. + exists trX'. exists initial_descriptors. subst foldx. split; [assumption|]. split; [apply Htr_project|]. split; [|assumption]. subst tr. clear -Hstate_project Hx. @@ -634,13 +634,13 @@ Qed. (** Instantiating the lemma above with the free constraint. *) -Lemma free_equivocators_protocol_trace_project +Lemma free_equivocators_valid_trace_project (final_descriptors : equivocator_descriptors IM) (is : composite_state (equivocator_IM IM)) (tr : list (composite_transition_item (equivocator_IM IM))) (final_state := @finite_trace_last _ (@type _ XE) is tr) (Hproper : proper_fixed_equivocator_descriptors final_descriptors final_state) - (Htr : finite_protocol_trace XE is tr) + (Htr : finite_valid_trace XE is tr) : exists (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors IM) @@ -649,9 +649,9 @@ Lemma free_equivocators_protocol_trace_project proper_fixed_equivocator_descriptors initial_descriptors is /\ equivocators_trace_project IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project IM final_descriptors final_state = final_stateX /\ - finite_protocol_trace (free_composite_vlsm IM) isX trX. + finite_valid_trace (free_composite_vlsm IM) isX trX. Proof. - apply _equivocators_protocol_trace_project; [assumption | assumption| ..] + apply _equivocators_valid_trace_project; [assumption | assumption| ..] ; unfold constraint_has_been_sent_prop; intros; exact I. Qed. @@ -661,7 +661,7 @@ projection of the final state. Lemma not_equivocating_sent_message_has_been_observed_in_projection (is: vstate XE) (tr: list (composite_transition_item (equivocator_IM IM))) - (Htr: finite_protocol_trace XE is tr) + (Htr: finite_valid_trace XE is tr) (lst := finite_trace_last is tr) (item: transition_item) (Hitem: item ∈ tr) @@ -672,23 +672,23 @@ Lemma not_equivocating_sent_message_has_been_observed_in_projection (Hdescriptors: proper_fixed_equivocator_descriptors descriptors lst) : has_been_observed Free (equivocators_state_project IM descriptors lst) m. Proof. - destruct (free_equivocators_protocol_trace_project descriptors is tr Hdescriptors Htr) + destruct (free_equivocators_valid_trace_project descriptors is tr Hdescriptors Htr) as [trX [initial_descriptors [_ [Htr_project [Hfinal_state HtrX_Free]]]]]. subst lst. rewrite Hfinal_state. - assert (Htr_Pre : finite_protocol_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr). - { revert Htr. apply VLSM_incl_finite_protocol_trace. + assert (Htr_Pre : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr). + { revert Htr. apply VLSM_incl_finite_valid_trace. apply VLSM_incl_trans with (machine FreeE); [|apply vlsm_incl_pre_loaded_with_all_messages_vlsm]. apply equivocators_fixed_equivocations_vlsm_incl_free. } - assert (HtrX_Pre : finite_protocol_trace (pre_loaded_with_all_messages_vlsm Free) (equivocators_state_project IM initial_descriptors is) trX ). - { revert HtrX_Free. apply VLSM_incl_finite_protocol_trace. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } + assert (HtrX_Pre : finite_valid_trace (pre_loaded_with_all_messages_vlsm Free) (equivocators_state_project IM initial_descriptors is) trX ). + { revert HtrX_Free. apply VLSM_incl_finite_valid_trace. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - assert (Hlst_preX : protocol_state_prop (pre_loaded_with_all_messages_vlsm Free) (finite_trace_last (equivocators_state_project IM initial_descriptors is) trX)). - { apply (finite_ptrace_last_pstate (pre_loaded_with_all_messages_vlsm Free)). + assert (Hlst_preX : valid_state_prop (pre_loaded_with_all_messages_vlsm Free) (finite_trace_last (equivocators_state_project IM initial_descriptors is) trX)). + { apply (finite_valid_trace_last_pstate (pre_loaded_with_all_messages_vlsm Free)). apply HtrX_Pre. } @@ -696,7 +696,7 @@ Proof. right. apply proper_sent; [assumption|]. apply has_been_sent_consistency; [assumption| assumption |]. exists (equivocators_state_project IM initial_descriptors is), trX, - (ptrace_add_default_last HtrX_Pre). + (valid_trace_add_default_last HtrX_Pre). apply elem_of_list_split in Hitem. destruct Hitem as [pre [suf Hitem]]. @@ -708,9 +708,9 @@ Proof. apply (not_equivocating_index_has_singleton_state IM Hbs _ finite_index equivocating); [|assumption]. apply proj1 in Htr. rewrite app_assoc in Htr. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr _]. - apply finite_ptrace_last_pstate in Htr. + apply finite_valid_trace_last_pstate in Htr. rewrite finite_trace_last_is_last in Htr. assumption. } @@ -719,14 +719,14 @@ Proof. subst trX. apply Exists_app. right. apply proj1 in Htr_Pre. - apply finite_protocol_trace_from_app_iff in Htr_Pre. + apply finite_valid_trace_from_app_iff in Htr_Pre. destruct Htr_Pre as [Hpre'_free Hsuf'_free]. apply equivocators_trace_project_app_iff in Hpr_suf. destruct Hpr_suf as [pre_itemX [sufX'' [eqv_descriptors'' [Hpr_suf' [Hpr_pre_item HsufX']]]]]. subst sufX'. apply Exists_app. left. - apply finite_protocol_trace_from_app_iff in Hsuf'_free. + apply finite_valid_trace_from_app_iff in Hsuf'_free. destruct Hsuf'_free as [Hpre_item_free Hsuf'_free]. assert (Heqv_descriptors'' : forall i, i ∉ equivocating -> eqv_descriptors'' i = Existing 0). { specialize (equivocators_trace_project_preserves_zero_descriptors IM _ _ Hsuf'_free descriptors) as Hpr. @@ -771,7 +771,7 @@ Proof. Qed. (** -Consider a [protocol_trace] for the composition of equivocators with +Consider a [valid_trace] for the composition of equivocators with no message equivocation and fixed state equivocation. Because any of its projections to the composition of original nodes contains @@ -783,13 +783,13 @@ Therefore if seeding the composition of equivocating nodes with these messages, the restriction of the initial trace to only the equivocating nodes will satisfy the [trace_sub_item_input_is_seeded_or_sub_previously_sent] property w.r.t. these messages, a sufficient condition for it being -protocol ([finite_protocol_trace_sub_projection]). +valid ([finite_valid_trace_sub_projection]). *) Lemma equivocators_trace_sub_item_input_is_seeded_or_sub_previously_sent (is : vstate XE) (tr : list (vtransition_item XE)) (s := finite_trace_last is tr) - (Htr : finite_protocol_trace XE is tr) + (Htr : finite_valid_trace XE is tr) (descriptors: equivocator_descriptors IM) (Hproper: proper_fixed_equivocator_descriptors descriptors s) (lst_trX := equivocators_state_project IM descriptors s) @@ -798,56 +798,57 @@ Lemma equivocators_trace_sub_item_input_is_seeded_or_sub_previously_sent (no_additional_equivocations (free_composite_vlsm IM) lst_trX) tr. Proof. intros pre item suf m Heq Hm Hitem. - destruct (free_equivocators_protocol_trace_project descriptors is tr Hproper Htr) + destruct (free_equivocators_valid_trace_project descriptors is tr Hproper Htr) as [trX [initial_descriptors [Hinitial_descriptors [Htr_project [Hfinal_state HtrXFree]]]]]. - assert (HtrXPre : finite_protocol_trace (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) (equivocators_state_project IM initial_descriptors is) trX). - { revert HtrXFree. apply VLSM_incl_finite_protocol_trace. + assert (HtrXPre : finite_valid_trace (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) (equivocators_state_project IM initial_descriptors is) trX). + { revert HtrXFree. apply VLSM_incl_finite_valid_trace. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - assert (Hlst_trX : protocol_state_prop (pre_loaded_with_all_messages_vlsm Free) lst_trX). + assert (Hlst_trX : valid_state_prop (pre_loaded_with_all_messages_vlsm Free) lst_trX). { subst lst_trX. subst s. simpl. simpl in Hfinal_state. rewrite Hfinal_state. - apply finite_ptrace_last_pstate. apply HtrXPre. + apply finite_valid_trace_last_pstate. apply HtrXPre. } - assert (Htr_free : finite_protocol_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr). - { revert Htr. apply VLSM_incl_finite_protocol_trace. + assert (Htr_free : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) is tr). + { revert Htr. apply VLSM_incl_finite_valid_trace. apply VLSM_incl_trans with (machine FreeE) ; [apply equivocators_fixed_equivocations_vlsm_incl_free|]. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } subst tr. destruct Htr as [Htr His]. - apply finite_protocol_trace_from_app_iff in Htr as Hsuf. destruct Hsuf as [_ Hsuf]. + apply finite_valid_trace_from_app_iff in Htr as Hsuf. destruct Hsuf as [_ Hsuf]. rewrite app_assoc in Htr. - apply finite_protocol_trace_from_app_iff in Htr as Hpre. destruct Hpre as [Hpre _]. - apply finite_protocol_trace_from_app_iff in Hpre. destruct Hpre as [Hpre Hpt]. + apply finite_valid_trace_from_app_iff in Htr as Hpre. destruct Hpre as [Hpre _]. + apply finite_valid_trace_from_app_iff in Hpre. destruct Hpre as [Hpre Hivt]. + apply first_transition_valid in Hivt. destruct (Free_no_additional_equivocation_decidable lst_trX m) ; [left; assumption|right]. unfold no_additional_equivocations in n. - assert (Hsuf_free : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm FreeE) (finite_trace_last is pre) ([item] ++ suf)). - { revert Hsuf. apply VLSM_incl_finite_protocol_trace_from. + assert (Hsuf_free : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm FreeE) (finite_trace_last is pre) ([item] ++ suf)). + { revert Hsuf. apply VLSM_incl_finite_valid_trace_from. apply VLSM_incl_trans with (machine FreeE) ; [apply equivocators_fixed_equivocations_vlsm_incl_free|]. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - assert (Hs_free : protocol_state_prop (pre_loaded_with_all_messages_vlsm FreeE) s). - { apply finite_ptrace_last_pstate in Hsuf_free. subst s. + assert (Hs_free : valid_state_prop (pre_loaded_with_all_messages_vlsm FreeE) s). + { apply finite_valid_trace_last_pstate in Hsuf_free. subst s. rewrite finite_trace_last_app. assumption. } - destruct item as (l, iom, s0, oom). apply first_transition_valid in Hpt. + destruct item as (l, iom, s0, oom). simpl in Hm. subst iom. - destruct Hpt as [[_ [_ [_ [[Hc _] Hfixed]]]] Ht]. + destruct Hivt as [[_ [_ [_ [[Hc _] Hfixed]]]] Ht]. simpl in Ht, Hfixed. rewrite Ht in Hfixed. simpl in Hfixed. clear Ht. destruct Hc as [Hc | Hinit]; [|contradiction]. - assert (Hpre_free : finite_protocol_trace FreeE is pre). - { split; [|assumption]. revert Hpre. apply VLSM_incl_finite_protocol_trace_from. + assert (Hpre_free : finite_valid_trace FreeE is pre). + { split; [|assumption]. revert Hpre. apply VLSM_incl_finite_valid_trace_from. apply equivocators_fixed_equivocations_vlsm_incl_free. } - assert (Hpre_lst_free : protocol_state_prop FreeE (finite_trace_last is pre)). - { apply (finite_ptrace_last_pstate FreeE). apply Hpre_free. } + assert (Hpre_lst_free : valid_state_prop FreeE (finite_trace_last is pre)). + { apply (finite_valid_trace_last_pstate FreeE). apply Hpre_free. } specialize (specialized_proper_sent_rev FreeE _ Hpre_lst_free m) as Hproper_sent. apply Hproper_sent in Hc. clear Hproper_sent. - specialize (Hc is pre (ptrace_add_default_last Hpre_free)). + specialize (Hc is pre (valid_trace_add_default_last Hpre_free)). apply Exists_exists in Hc. destruct Hc as [pre_item [Hpre_item Hpre_m]]. exists pre_item. split; [assumption|]. split; [assumption|]. @@ -858,7 +859,7 @@ Proof. assert (Hfinal' : proper_fixed_equivocator_descriptors final_descriptors' (finite_trace_last is pre)). { split. - apply proj1 in Hproper. subst s. rewrite finite_trace_last_app in Hproper. - destruct (preloaded_equivocators_protocol_trace_from_project IM _ _ _ Hproper Hsuf_free) + destruct (preloaded_equivocators_valid_trace_from_project IM _ _ _ Hproper Hsuf_free) as [_sufX [_final_descriptors' [_Hsuf_project [Hproper' _]]]]. rewrite Hsuf_project in _Hsuf_project. inversion _Hsuf_project. subst _sufX _final_descriptors'. clear _Hsuf_project. @@ -877,21 +878,21 @@ Proof. rewrite app_assoc in Hpre_item. rewrite app_assoc in Htr_free. destruct Htr_free as [Htr_free _]. - apply finite_protocol_trace_from_app_iff in Htr_free. + apply finite_valid_trace_from_app_iff in Htr_free. destruct Htr_free as [Htr_s0 _]. subst pre. clear -Htr_s0. rewrite <- app_assoc in Htr_s0. - apply (finite_protocol_trace_from_app_iff (pre_loaded_with_all_messages_vlsm FreeE)) in Htr_s0. + apply (finite_valid_trace_from_app_iff (pre_loaded_with_all_messages_vlsm FreeE)) in Htr_s0. destruct Htr_s0 as [_ Hfuture]. rewrite finite_trace_last_is_last in Hfuture. eexists. - apply ptrace_add_last;[apply Hfuture|]. + apply valid_trace_add_last;[apply Hfuture|]. rewrite finite_trace_last_is_last;reflexivity. } apply (@in_futures_reflects_fixed_equivocation _ _ _ IM index_listing equivocating) in Hfutures ; [|assumption]. - destruct (free_equivocators_protocol_trace_project final_descriptors' is pre Hfinal' (conj Hpre His)) + destruct (free_equivocators_valid_trace_project final_descriptors' is pre Hfinal' (conj Hpre His)) as [_preX [_initial_descriptors [_ [_Htr_project [Hpre_final_state _]]]]]. rewrite Htr_project in _Htr_project. inversion _Htr_project. subst _preX _initial_descriptors. clear _Htr_project. @@ -916,8 +917,8 @@ Proof. exists sufX. clear -HtrXPre. apply proj1 in HtrXPre. - apply finite_protocol_trace_from_app_iff in HtrXPre. - apply ptrace_add_default_last. + apply finite_valid_trace_from_app_iff in HtrXPre. + apply valid_trace_add_default_last. apply HtrXPre. Qed. @@ -930,7 +931,7 @@ in any projection. Lemma equivocator_vlsm_trace_project_reflect_non_equivocating (is: composite_state (equivocator_IM IM)) (tr: list (composite_transition_item (equivocator_IM IM))) - (Htr: finite_protocol_trace XE is tr) + (Htr: finite_valid_trace XE is tr) (m: message) (final_descriptors: equivocator_descriptors IM) (Hproper: proper_fixed_equivocator_descriptors final_descriptors (finite_trace_last is tr)) @@ -957,15 +958,15 @@ Proof. apply Exists_app. left. destruct Htr as [Htr _]. rewrite app_assoc in Htr. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Hpre Hsuf]. - apply finite_protocol_trace_from_app_iff in Hpre. + apply finite_valid_trace_from_app_iff in Hpre. destruct Hpre as [_ Hitem]. rewrite app_assoc,finite_trace_last_app in Hproper. rewrite finite_trace_last_is_last in Hsuf, Hproper. - assert (Hsufpre : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm FreeE) (destination item) suf). + assert (Hsufpre : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm FreeE) (destination item) suf). { - revert Hsuf. apply VLSM_incl_finite_protocol_trace_from. + revert Hsuf. apply VLSM_incl_finite_valid_trace_from. apply VLSM_incl_trans with (machine FreeE). - apply equivocators_fixed_equivocations_vlsm_incl_free. - apply vlsm_incl_pre_loaded_with_all_messages_vlsm. @@ -983,7 +984,7 @@ Proof. spec Hex. { apply (not_equivocating_index_has_singleton_state _ Hbs _ finite_index equivocating); [|assumption]. - apply finite_ptrace_last_pstate in Hitem. assumption. + apply finite_valid_trace_last_pstate in Hitem. assumption. } destruct item as (l, iom, s, oom). apply first_transition_valid in Hitem. simpl in Hitem. destruct Hitem as [[Hs [Hiom [Hv Hc]]] Ht]. @@ -1003,7 +1004,7 @@ the nodes allowed to equivocate. Lemma projection_has_not_been_observed_is_equivocating (is: _composite_state (equivocator_IM IM)) (tr: list (composite_transition_item (equivocator_IM IM))) - (Htr: finite_protocol_trace XE is tr) + (Htr: finite_valid_trace XE is tr) (s := @finite_trace_last _ (composite_type (equivocator_IM IM)) is tr) (descriptors: equivocator_descriptors IM) (Hproper: proper_fixed_equivocator_descriptors descriptors s) @@ -1013,26 +1014,26 @@ Lemma projection_has_not_been_observed_is_equivocating : forall item : composite_transition_item (equivocator_IM IM), item ∈ tr -> output item = Some m -> (projT1 (l item)) ∈ equivocating. Proof. - destruct (free_equivocators_protocol_trace_project descriptors is tr Hproper Htr) + destruct (free_equivocators_valid_trace_project descriptors is tr Hproper Htr) as [trX [initial_descriptors [_ [Htr_project [Hlast_state HtrX]]]]]. intros item Hitem Houtput. destruct (decide ((projT1 (l item)) ∈ equivocating)); [assumption|]. elim Hno. clear Hno. apply composite_has_been_observed_sent_received_iff. left. - assert (HtrX_free : finite_protocol_trace (pre_loaded_with_all_messages_vlsm Free) (equivocators_state_project IM initial_descriptors is) trX). + assert (HtrX_free : finite_valid_trace (pre_loaded_with_all_messages_vlsm Free) (equivocators_state_project IM initial_descriptors is) trX). { - revert HtrX. clear. apply VLSM_incl_finite_protocol_trace. + revert HtrX. clear. apply VLSM_incl_finite_valid_trace. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - specialize (finite_ptrace_last_pstate _ _ _ (proj1 HtrX_free)) as Hfree_lst. + specialize (finite_valid_trace_last_pstate _ _ _ (proj1 HtrX_free)) as Hfree_lst. subst sX s. simpl in *. rewrite Hlast_state. apply (composite_proper_sent IM finite_index Hbs) ; [assumption|]. apply has_been_sent_consistency; [assumption| assumption |]. - exists (equivocators_state_project IM initial_descriptors is), trX, (ptrace_add_default_last HtrX_free). + exists (equivocators_state_project IM initial_descriptors is), trX, (valid_trace_add_default_last HtrX_free). clear HtrX HtrX_free Hfree_lst. apply equivocator_vlsm_trace_project_reflect_non_equivocating with is tr descriptors initial_descriptors item ; assumption. @@ -1066,7 +1067,7 @@ Lemma pre_loaded_equivocators_composition_sub_projection_commute Proof. apply basic_VLSM_incl ; intro; intros; [assumption|..]. - - apply initial_message_is_protocol. + - apply initial_message_is_valid. simpl. destruct HmX. + left; assumption. @@ -1084,9 +1085,9 @@ Proof. exists subi. revert Hibs. apply has_been_sent_irrelevance. simpl. - apply (preloaded_protocol_state_projection (equivocator_IM (sub_IM IM equivocating)) subi). + apply (preloaded_valid_state_projection (equivocator_IM (sub_IM IM equivocating)) subi). revert Hs. - apply VLSM_incl_protocol_state. + apply VLSM_incl_valid_state. apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages. - apply H. Qed. @@ -1111,7 +1112,7 @@ Lemma pre_loaded_equivocators_composition_sub_projection_commute_inv Proof. apply basic_VLSM_incl ; intro; intros; [assumption|..]. - - apply initial_message_is_protocol. + - apply initial_message_is_valid. simpl. destruct HmX; [left; assumption|]. right. @@ -1128,9 +1129,9 @@ Proof. exists subi. revert Hibs. apply has_been_sent_irrelevance. simpl. - apply (preloaded_protocol_state_projection (equivocator_IM (sub_IM IM equivocating)) subi). + apply (preloaded_valid_state_projection (equivocator_IM (sub_IM IM equivocating)) subi). revert Hs. - apply VLSM_incl_protocol_state. + apply VLSM_incl_valid_state. apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages. - apply H. Qed. @@ -1140,12 +1141,12 @@ Qed. The intermediary results above allow us to prove that the [fixed_equivocation_constraint] has the [constraint_has_been_sent_prop]erty. -The core of this result is proving that given a [protocol_state] @s@ of the +The core of this result is proving that given a [valid_state] <> of the composition of equivocators with no message equivocation and fixed state equivocation, a message which [has_been_sent] for that state but not -[has_been_observed] for a projection of that state @sx@, can nevertheless be +[has_been_observed] for a projection of that state <>, can nevertheless be generated by the composition of the nodes allowed to equivocate, pre-loaded with -the messages observed in the state @sx@. +the messages observed in the state <>. To prove that, we consider a trace witness for the mesage having been sent, we use [projection_has_not_been_observed_is_equivocating] to derive that @@ -1153,7 +1154,7 @@ it must have been sent by one of the machines allowed to equivocate, from this we derive that it can be sent by the restriction of the composition of equivocators to just the equivocating nodes, pre-loaded with the messages observed in the projection, then we use -the [seeded_equivocators_protocol_trace_project] result to reach our conclusion. +the [seeded_equivocators_valid_trace_project] result to reach our conclusion. *) Lemma fixed_equivocation_constraint_has_constraint_has_been_sent_prop : constraint_has_been_sent_prop @@ -1166,32 +1167,32 @@ Proof. clear l. unfold no_additional_equivocations in n. - (* Phase I: exhibiting a [protocol_trace] ending in tr s and sending m *) + (* Phase I: exhibiting a [valid_trace] ending in tr s and sending m *) - apply protocol_state_has_trace in Hs. + apply valid_state_has_trace in Hs. destruct Hs as [is [tr Htr]]. (* subst s *) - assert (Htr'pre : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm FreeE) is s tr). - { revert Htr. apply VLSM_incl_finite_protocol_trace_init_to. + assert (Htr'pre : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm FreeE) is s tr). + { revert Htr. apply VLSM_incl_finite_valid_trace_init_to. apply VLSM_incl_trans with (machine FreeE). - apply (constraint_free_incl (equivocator_IM IM) (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating)). - apply vlsm_incl_pre_loaded_with_all_messages_vlsm. } - assert (Hplst : protocol_state_prop (pre_loaded_with_all_messages_vlsm FreeE) s). - { apply ptrace_last_pstate in Htr'pre; assumption. } + assert (Hplst : valid_state_prop (pre_loaded_with_all_messages_vlsm FreeE) s). + { apply valid_trace_last_pstate in Htr'pre; assumption. } apply proper_sent in Hm; [|assumption]. clear Hplst. specialize (Hm is tr Htr'pre). (* Phase II (a): The restriction of tr to the equivocators allowed to - state-equivocate is protocol for the corresponding composition + state-equivocate is valid for the corresponding composition pre-loaded with the messages observed in the projection sX of s. *) specialize - (finite_protocol_trace_sub_projection (equivocator_IM IM) equivocating + (finite_valid_trace_sub_projection (equivocator_IM IM) equivocating (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating) (equivocator_Hbs IM Hbs) finite_index @@ -1200,13 +1201,13 @@ Proof. spec Hproject is tr. spec Hproject. { subst sX. - rewrite <- (ptrace_get_last Htr) in Hdescriptors |- *. + rewrite <- (valid_trace_get_last Htr) in Hdescriptors |- *. apply (equivocators_trace_sub_item_input_is_seeded_or_sub_previously_sent - _ _ (ptrace_forget_last Htr) descriptors Hdescriptors + _ _ (valid_trace_forget_last Htr) descriptors Hdescriptors ). } - spec Hproject (ptrace_forget_last Htr). + spec Hproject (valid_trace_forget_last Htr). rewrite HeqsX in n. clear HeqsX. @@ -1214,7 +1215,7 @@ Proof. Obtain a projection trXm of tr outputing m using a final_descriptor_m *) - apply (equivocators_trace_project_output_reflecting_iff _ _ _ (proj1 (ptrace_forget_last Htr'pre))) in Hm. + apply (equivocators_trace_project_output_reflecting_iff _ _ _ (proj1 (valid_trace_forget_last Htr'pre))) in Hm. destruct Hm as [final_descriptors_m [initial_descriptors_m [trXm [Hfinal_descriptors_m [Hproject_trXm Hex]]]]]. (* Identify the item outputing m in trXm an its corresponding item in tr. @@ -1242,9 +1243,9 @@ Proof. (* show that that item must be specifying a transition for an equivocating node*) - rewrite <- (ptrace_get_last Htr) in Hdescriptors, n. + rewrite <- (valid_trace_get_last Htr) in Hdescriptors, n. specialize - (projection_has_not_been_observed_is_equivocating _ _ (ptrace_forget_last Htr) + (projection_has_not_been_observed_is_equivocating _ _ (valid_trace_forget_last Htr) _ Hdescriptors _ n item ) as Hitem_equivocating. @@ -1258,12 +1259,12 @@ Proof. (* Phase III (b): Consider a projection trX' obtained using the final_descriptor_m as above, but first restricting the nodes to just the equivocators allowed to equivocate. - We will show that we can use [seeded_equivocators_protocol_trace_project] + We will show that we can use [seeded_equivocators_valid_trace_project] and leverage the result from Phase II (a) - to derive that the resulting projection is protocol. + to derive that the resulting projection is valid. *) specialize - (seeded_equivocators_protocol_trace_project IM Hbs finite_index equivocating + (seeded_equivocators_valid_trace_project IM Hbs finite_index equivocating (fun m : message => no_additional_equivocations (free_composite_vlsm IM) sX m) @@ -1276,7 +1277,7 @@ specialize spec Hsub_project. { specialize (finite_trace_sub_projection_last_state (equivocator_IM IM) equivocating - (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating) _ _ (proj1 (ptrace_forget_last Htr))) + (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating) _ _ (proj1 (valid_trace_forget_last Htr))) as Heq_lst. simpl in Heq_lst. match goal with @@ -1313,7 +1314,7 @@ specialize (equivocators_state_project (sub_IM IM equivocating) initial_descriptors' (composite_state_sub_projection (equivocator_IM IM) equivocating is) ) as isX. clear HeqisX. - apply can_emit_from_protocol_trace with isX trX'; [assumption|]. + apply can_emit_from_valid_trace with isX trX'; [assumption|]. clear HtrX. subst. @@ -1329,17 +1330,17 @@ specialize Qed. (** -Main result of this section, stating that traces which are protocol for the +Main result of this section, stating that traces which are valid for the equivocator-based definition of fixed equivocation project to traces which are -protocol for the simple-nodes definition of fixed equivocation. +valid for the simple-nodes definition of fixed equivocation. *) -Theorem fixed_equivocators_protocol_trace_project +Theorem fixed_equivocators_valid_trace_project (final_descriptors : equivocator_descriptors IM) (is : composite_state (equivocator_IM IM)) (tr : list (composite_transition_item (equivocator_IM IM))) (final_state := finite_trace_last is tr) (Hproper: proper_fixed_equivocator_descriptors final_descriptors final_state) - (Htr : finite_protocol_trace XE is tr) + (Htr : finite_valid_trace XE is tr) : exists (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors IM) @@ -1348,9 +1349,9 @@ Theorem fixed_equivocators_protocol_trace_project proper_fixed_equivocator_descriptors initial_descriptors is /\ equivocators_trace_project IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project IM final_descriptors final_state = final_stateX /\ - finite_protocol_trace X isX trX. + finite_valid_trace X isX trX. Proof. - apply _equivocators_protocol_trace_project; [assumption | assumption| ..] + apply _equivocators_valid_trace_project; [assumption | assumption| ..] ; intros. - exact I. - apply fixed_equivocation_constraint_has_constraint_has_been_sent_prop. @@ -1363,8 +1364,8 @@ Proof. split; [split|]. - intros s tr sX trX Hpr_tr s_pre pre Hs_lst Hpre_tr. assert - (HPreFree_pre_tr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm FreeE) s_pre (pre ++ tr)). - { revert Hpre_tr. apply VLSM_incl_finite_protocol_trace_from. + (HPreFree_pre_tr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm FreeE) s_pre (pre ++ tr)). + { revert Hpre_tr. apply VLSM_incl_finite_valid_trace_from. apply equivocators_fixed_equivocations_vlsm_incl_PreFree. } clear Hpre_tr. revert s tr sX trX Hpr_tr s_pre pre Hs_lst HPreFree_pre_tr. @@ -1373,10 +1374,10 @@ Proof. destruct (destruct_equivocators_partial_trace_project IM finite_index Hpr_tr) as [Hnot_equiv [initial_descriptors [Htr_project Hs_project]]]. - specialize (finite_ptrace_last_pstate XE _ _ (proj1 Htr)) as Hlst. + specialize (finite_valid_trace_last_pstate XE _ _ (proj1 Htr)) as Hlst. apply not_equivocating_equivocator_descriptors_proper_fixed in Hnot_equiv as Hproper ; [|assumption]. - destruct (fixed_equivocators_protocol_trace_project _ _ _ Hproper Htr) + destruct (fixed_equivocators_valid_trace_project _ _ _ Hproper Htr) as [_trX [_initial_descriptors [_ [_Htr_project [_ HtrX]]]]]. rewrite Htr_project in _Htr_project. inversion _Htr_project. subst. assumption. @@ -1387,14 +1388,14 @@ Lemma fixed_equivocators_vlsm_projection Proof. constructor; [constructor|]; intros. - apply PreFreeE_Free_vlsm_projection_type. - revert H. apply VLSM_incl_finite_protocol_trace_from. + revert H. apply VLSM_incl_finite_valid_trace_from. apply equivocators_fixed_equivocations_vlsm_incl_PreFree. - - assert (Hpre_tr : finite_protocol_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). - { revert H. apply VLSM_incl_finite_protocol_trace. + - assert (Hpre_tr : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). + { revert H. apply VLSM_incl_finite_valid_trace. apply equivocators_fixed_equivocations_vlsm_incl_PreFree. } specialize - (VLSM_partial_projection_finite_protocol_trace (fixed_equivocators_vlsm_partial_projection (zero_descriptor IM)) + (VLSM_partial_projection_finite_valid_trace (fixed_equivocators_vlsm_partial_projection (zero_descriptor IM)) sX trX (equivocators_state_project IM (zero_descriptor IM) sX) (equivocators_total_trace_project IM trX) ) as Hsim. spec Hsim. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocationSimulation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocationSimulation.v index 04f0cdca..e538a4a5 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocationSimulation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocationSimulation.v @@ -17,8 +17,8 @@ In this module we show that the composition of equivocators with no message-equivocation and fixed-set state-equivocation can simulate the fixed-set message-equivocation composition of regular nodes. -The proof is based on [generalized_equivocators_finite_protocol_trace_init_to_rev], -but also reuses [seeded_equivocators_finite_protocol_trace_init_to_rev] to +The proof is based on [generalized_equivocators_finite_valid_trace_init_to_rev], +but also reuses [seeded_equivocators_finite_valid_trace_init_to_rev] to lift a trace of the free subcomposition of equivocating nodes generating a message (given by the fixed message-equivocation constraint) to a trace of the subcomposition of equivocating equivocators with no message-equivocation, @@ -59,30 +59,30 @@ Proof. assumption. Qed. -(** Messages [sent_by_non_equivocating] nodes in the projection of a protocol +(** Messages [sent_by_non_equivocating] nodes in the projection of a valid state for the fixed-set state-equivocation composition of equivocators with no -message-equivocation are protocol for that composition of equivocators. +message-equivocation are valid for that composition of equivocators. *) -Lemma fixed_equivocating_sent_by_non_equivocating_protocol +Lemma fixed_equivocating_messages_sent_by_non_equivocating_are_valid eqv_state_s - (Heqv_state_s : protocol_state_prop XE eqv_state_s) + (Heqv_state_s : valid_state_prop XE eqv_state_s) (s := equivocators_total_state_project IM eqv_state_s) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) m (Hm : sent_by_non_equivocating IM Hbs equivocating s m) - : protocol_message_prop XE m. + : valid_message_prop XE m. Proof. specialize (equivocators_fixed_equivocations_vlsm_incl_PreFree IM Hbs index_listing equivocating) as HinclE. apply sent_by_non_equivocating_are_sent in Hm. specialize (StrongFixed_incl_Preloaded IM Hbs equivocating) as Hincl. - apply (VLSM_incl_protocol_state Hincl) in Hs. + apply (VLSM_incl_valid_state Hincl) in Hs. apply - (composite_sent_protocol (equivocator_IM IM) finite_index (equivocator_Hbs IM Hbs) + (composite_sent_valid (equivocator_IM IM) finite_index (equivocator_Hbs IM Hbs) _ _ Heqv_state_s). revert Hm. - apply (VLSM_incl_protocol_state HinclE) in Heqv_state_s. + apply (VLSM_incl_valid_state HinclE) in Heqv_state_s. by specialize (VLSM_projection_has_been_sent_reflect (preloaded_equivocators_no_equivocations_vlsm_X_vlsm_projection IM Hbs finite_index) @@ -94,14 +94,14 @@ property (for the [equivocators_fixed_equivocations_constraint]). *) Lemma fixed_equivocating_non_equivocating_constraint_lifting eqv_state_s - (Heqv_state_s : protocol_state_prop XE eqv_state_s) + (Heqv_state_s : valid_state_prop XE eqv_state_s) l s om (Hv : - protocol_valid + input_valid (seeded_equivocators_no_equivocation_vlsm IM Hbs equivocating (sent_by_non_equivocating IM Hbs equivocating (equivocators_total_state_project IM eqv_state_s))) l (s, om)) - (Hlift_s : protocol_state_prop XE (lift_equivocators_sub_state_to IM equivocating eqv_state_s s)) + (Hlift_s : valid_state_prop XE (lift_equivocators_sub_state_to IM equivocating eqv_state_s s)) : equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating (lift_equivocators_sub_label_to IM equivocating eqv_state_s l) @@ -110,15 +110,15 @@ Proof. specialize (equivocators_fixed_equivocations_vlsm_incl_PreFree IM Hbs index_listing equivocating) as HinclE. destruct Hv as [Hs [Hom [_ Hc]]]. - apply protocol_has_fixed_equivocation in Hlift_s. + apply valid_state_has_fixed_equivocation in Hlift_s. split. - split; [|exact I]. destruct om as [m|]; [|exact I]. left. apply proj1 in Hc as [Hsent | Hseed]. + revert Hsent. simpl. - apply (VLSM_incl_protocol_state HinclE) in Heqv_state_s. - apply (VLSM_incl_protocol_state (seeded_no_equivocation_incl_preloaded IM Hbs equivocating _)) in Hs. + apply (VLSM_incl_valid_state HinclE) in Heqv_state_s. + apply (VLSM_incl_valid_state (seeded_no_equivocation_incl_preloaded IM Hbs equivocating _)) in Hs. by specialize (VLSM_weak_full_projection_has_been_sent (PreFreeSubE_PreFreeE_weak_full_projection IM _ finite_index equivocating @@ -136,9 +136,9 @@ Proof. apply (VLSM_projection_has_been_sent_reflect (preloaded_equivocator_zero_projection (IM i)) (HbsX := equivocator_HasBeenSentCapability (IM i))). - apply (VLSM_incl_protocol_state HinclE) in Heqv_state_s. + apply (VLSM_incl_valid_state HinclE) in Heqv_state_s. revert Heqv_state_s. - apply (VLSM_projection_protocol_state (preloaded_component_projection (equivocator_IM IM) i)). + apply (VLSM_projection_valid_state (preloaded_component_projection (equivocator_IM IM) i)). - destruct (composite_transition _ _ _) as (s', om') eqn:Ht'. apply (equivocating_transition_preserves_fixed_equivocation @@ -151,24 +151,24 @@ Qed. (** A message that can be generated from a state <> of the free composition of equivocating equivocators pre-loaded with all messages has the [composite_has_been_sent] property for the state obtained upon "appending" - state <> to protocol state for the composition of all equivocators. + state <> to valid state for the composition of all equivocators. This result plays an important role in satisfying the no-message equivocation constraint. *) Lemma fixed_equivocation_replay_has_message eqv_state_s - (Heqv_state_s : protocol_state_prop PreFreeE eqv_state_s) + (Heqv_state_s : valid_state_prop PreFreeE eqv_state_s) im s (Him : - protocol_generated_prop + can_produce (pre_loaded_with_all_messages_vlsm (free_composite_vlsm (equivocator_IM (sub_IM IM equivocating)))) s im) : composite_has_been_sent (equivocator_IM IM) (equivocator_Hbs IM Hbs) (lift_equivocators_sub_state_to IM equivocating eqv_state_s s) im. Proof. - apply non_empty_protocol_trace_from_protocol_generated_prop in Him + apply non_empty_valid_trace_from_can_produce in Him as [im_eis [im_etr [item [Him_etr [Hlast [Heqs Him]]]]]]. specialize (PreFreeSubE_PreFreeE_weak_full_projection IM index_listing finite_index @@ -180,8 +180,8 @@ Proof. (finite_sub_index equivocating finite_index) (sub_has_been_sent_capabilities (equivocator_IM IM) equivocating (equivocator_Hbs IM Hbs)) (free_constraint _)). - apply ptrace_add_default_last in Him_etr. - apply ptrace_last_pstate in Him_etr as Him_etr_lst. + apply valid_trace_add_default_last in Him_etr. + apply valid_trace_last_pstate in Him_etr as Him_etr_lst. specialize (VLSM_weak_full_projection_has_been_sent Hproj _ Him_etr_lst im) as Hsent. unfold has_been_sent in Hsent. simpl in Hsent. @@ -223,28 +223,28 @@ Proof. split; [reflexivity|]. split; [assumption|]. apply sent_by_non_equivocating_are_sent in Hsent. - apply (VLSM_eq_protocol_state HeqXE) in Hstate_protocol. - apply (VLSM_incl_protocol_state HinclE) in Hstate_protocol. + apply (VLSM_eq_valid_state HeqXE) in Hstate_valid. + apply (VLSM_incl_valid_state HinclE) in Hstate_valid. left. revert Hsent. subst. specialize (VLSM_projection_has_been_sent_reflect (preloaded_equivocators_no_equivocations_vlsm_X_vlsm_projection IM Hbs finite_index) - eqv_state_s Hstate_protocol im) as Hsent. + eqv_state_s Hstate_valid im) as Hsent. unfold has_been_sent in Hsent. simpl in Hsent. assumption. - (* If <> can be emitted by the free composition of equivocating nodes seeded with the messages [sent_by_non_equivocating] in <>, then we can - use Lemma [seeded_equivocators_finite_protocol_trace_init_to_rev] to + use Lemma [seeded_equivocators_finite_valid_trace_init_to_rev] to simulate that trace in the equivocator-composition of equivocating nodes with no-messages equivocation. *) apply can_emit_has_trace in Hemitted as [im_is [im_tr [im_item [Him_tr Him_item]]]]. - apply ptrace_add_default_last in Him_tr. + apply valid_trace_add_default_last in Him_tr. rewrite finite_trace_last_is_last in Him_tr. apply - (seeded_equivocators_finite_protocol_trace_init_to_rev + (seeded_equivocators_finite_valid_trace_init_to_rev (sub_IM IM equivocating) (sub_has_been_sent_capabilities IM equivocating Hbs) (finite_sub_index equivocating finite_index) _ no_initial_messages_in_sub_IM) in Him_tr @@ -252,15 +252,15 @@ Proof. rewrite finite_trace_last_output_is_last in Him_output. rewrite Him_item in Him_output. (* We will use Lemma - [sub_preloaded_replayed_trace_from_protocol_equivocating] to replay + [sub_preloaded_replayed_trace_from_valid_equivocating] to replay the trace obtained above on top of the given state, thus ensuring that state-equivocation is only introduced for the equivocating nodes. - We will used Lemmas [fixed_equivocating_sent_by_non_equivocating_protocol] + We will used Lemmas [fixed_equivocating_messages_sent_by_non_equivocating_are_valid] and [fixed_equivocating_non_equivocating_constraint_lifting] to satisfy the hypotheses of the replay lemma. *) specialize - (sub_preloaded_replayed_trace_from_protocol_equivocating + (sub_preloaded_replayed_trace_from_valid_equivocating IM Hbs _ finite_index (sent_by_non_equivocating IM Hbs equivocating s) equivocating (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating) @@ -269,8 +269,8 @@ Proof. spec Hreplay. { clear -finite_index HeqXE. intros i ns s Hi Hs. split; [split; exact I|]. - apply (VLSM_eq_protocol_state HeqXE) in Hs. - apply protocol_has_fixed_equivocation in Hs. + apply (VLSM_eq_valid_state HeqXE) in Hs. + apply valid_state_has_fixed_equivocation in Hs. destruct (composite_transition _ _ _) as (s', om') eqn:Ht. apply (equivocating_transition_preserves_fixed_equivocation @@ -279,30 +279,30 @@ Proof. } spec Hreplay. { subst s. - apply ptrace_last_pstate in HtrX. - apply (VLSM_eq_protocol_state HeqXE) in Hstate_protocol. - apply (VLSM_eq_protocol_state HeqX) in HtrX. + apply valid_trace_last_pstate in HtrX. + apply (VLSM_eq_valid_state HeqXE) in Hstate_valid. + apply (VLSM_eq_valid_state HeqX) in HtrX. intros m Hsent. - apply (VLSM_incl_protocol_message (VLSM_eq_proj1 HeqXE)); [by left|]. - by apply (fixed_equivocating_sent_by_non_equivocating_protocol eqv_state_s). + apply (VLSM_incl_valid_message (VLSM_eq_proj1 HeqXE)); [by left|]. + by apply (fixed_equivocating_messages_sent_by_non_equivocating_are_valid eqv_state_s). } spec Hreplay eqv_state_s. spec Hreplay. - { by apply (VLSM_eq_protocol_state HeqXE), (VLSM_eq_protocol_state HeqXE). } + { by apply (VLSM_eq_valid_state HeqXE), (VLSM_eq_valid_state HeqXE). } spec Hreplay. { subst s. - clear -finite_index FreeE_Hbs SubFreeE_Hbs HeqXE Hstate_protocol. + clear -finite_index FreeE_Hbs SubFreeE_Hbs HeqXE Hstate_valid. intros l s om Hv Hlift_s. - apply (VLSM_eq_protocol_state HeqXE) in Hstate_protocol. - apply (VLSM_eq_protocol_state HeqXE) in Hlift_s. + apply (VLSM_eq_valid_state HeqXE) in Hstate_valid. + apply (VLSM_eq_valid_state HeqXE) in Hlift_s. by apply fixed_equivocating_non_equivocating_constraint_lifting. } - apply ptrace_get_last in Him_etr as Him_etr_lst. - apply ptrace_forget_last in Him_etr. + apply valid_trace_get_last in Him_etr as Him_etr_lst. + apply valid_trace_forget_last in Him_etr. specialize (Hreplay _ _ Him_etr). - apply ptrace_add_default_last in Hreplay. + apply valid_trace_add_default_last in Hreplay. eexists _,_; split; [exact Hreplay|]. - (* Having verified the protocol-ness part of the conclusion, now we only + (* Having verified the validity part of the conclusion, now we only need to show two projection properties, and the no message-equivocation constraint for which we employ Lemma [fixed_equivocation_replay_has_message]. *) @@ -317,8 +317,8 @@ Proof. IM index_listing equivocating eqv_state_s im_eis im_etr). + - apply (VLSM_eq_protocol_state HeqXE) in Hstate_protocol. - apply (VLSM_incl_protocol_state HinclE) in Hstate_protocol as Hstate_pre. + apply (VLSM_eq_valid_state HeqXE) in Hstate_valid. + apply (VLSM_incl_valid_state HinclE) in Hstate_valid as Hstate_pre. specialize (NoEquivocation.seeded_no_equivocation_incl_preloaded (equivocator_IM (sub_IM IM equivocating)) (free_constraint _) @@ -326,7 +326,7 @@ Proof. (sent_by_non_equivocating IM Hbs equivocating s) ) as Hsub_incl. - apply (VLSM_incl_finite_protocol_trace Hsub_incl) in Him_etr. + apply (VLSM_incl_finite_valid_trace Hsub_incl) in Him_etr. left. specialize (replayed_trace_from_finite_trace_last IM _ finite_index equivocating eqv_state_s im_eis im_etr (proj2 Him_etr)). @@ -338,7 +338,7 @@ Proof. replace (im_ert' ++ [item]) with (im_ert' ++ [item] ++ []) in Him_etr by (rewrite app_assoc; apply app_nil_r). specialize - (protocol_transition_to + (input_valid_transition_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm (equivocator_IM (sub_IM IM equivocating)))) _ _ _ _ _ Him_etr eq_refl) as Ht. @@ -352,17 +352,17 @@ Qed. The main result, showing that fixed-set message-equivocation traces can be simulated by fixed-set state-equivocation traces is obtained by instantiating -Lemma [generalized_equivocators_finite_protocol_trace_init_to_rev], discharing +Lemma [generalized_equivocators_finite_valid_trace_init_to_rev], discharing the most complex assumption of the lemma (about [replayable_message_prop]) with Lemma [fixed_equivocation_has_replayable_message_prop]. *) -Lemma fixed_equivocators_finite_protocol_trace_init_to_rev +Lemma fixed_equivocators_finite_valid_trace_init_to_rev isX sX trX - (HtrX : finite_protocol_trace_init_to X isX sX trX) + (HtrX : finite_valid_trace_init_to X isX sX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = sX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to XE is s tr /\ + finite_valid_trace_init_to XE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. (* @@ -376,9 +376,9 @@ Proof. } specialize (vlsm_is_pre_loaded_with_False X) as HeqX. specialize (vlsm_is_pre_loaded_with_False XE) as HeqXE. - apply (VLSM_eq_finite_protocol_trace_init_to HeqX) in HtrX. + apply (VLSM_eq_finite_valid_trace_init_to HeqX) in HtrX. apply - (generalized_equivocators_finite_protocol_trace_init_to_rev + (generalized_equivocators_finite_valid_trace_init_to_rev IM Hbs finite_index _ _ (equivocators_fixed_equivocations_constraint IM Hbs index_listing equivocating)) in HtrX @@ -387,15 +387,15 @@ Proof. exists s. split; [assumption|]. exists tr. split; [assumption|]. split; [|assumption]. - apply (VLSM_eq_finite_protocol_trace_init_to HeqXE). + apply (VLSM_eq_finite_valid_trace_init_to HeqXE). assumption. - intro; intros. - apply (VLSM_eq_protocol_state HeqXE) in Hes. + apply (VLSM_eq_valid_state HeqXE) in Hes. split. + split; [|exact I]. destruct om as [im|]; [|exact I]. destruct Hom as [Hom|Hinitial]; [left; assumption|exfalso]. apply no_initial_messages_in_XE in Hinitial. contradiction. - + apply protocol_has_fixed_equivocation in Hes. + + apply valid_state_has_fixed_equivocation in Hes. destruct (composite_transition _ _ _) as (es', om') eqn:Het. simpl. apply @@ -437,14 +437,14 @@ Context {message : Type} Existing Instance Free_Hbs. Existing Instance FreeE_Hbs. -Lemma no_equivocating_equivocators_finite_protocol_trace_init_to_rev +Lemma no_equivocating_equivocators_finite_valid_trace_init_to_rev (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) isX sX trX - (HtrX : finite_protocol_trace_init_to X isX sX trX) + (HtrX : finite_valid_trace_init_to X isX sX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = sX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to XE is s tr /\ + finite_valid_trace_init_to XE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. assert (no_initial_messages_in_XE : forall m, ~vinitial_message_prop (pre_loaded_vlsm XE (fun _ => False)) m). @@ -455,9 +455,9 @@ Proof. } specialize (vlsm_is_pre_loaded_with_False X) as HeqX. specialize (vlsm_is_pre_loaded_with_False XE) as HeqXE. - apply (VLSM_eq_finite_protocol_trace_init_to HeqX) in HtrX. + apply (VLSM_eq_finite_valid_trace_init_to HeqX) in HtrX. apply - (generalized_equivocators_finite_protocol_trace_init_to_rev + (generalized_equivocators_finite_valid_trace_init_to_rev IM Hbs finite_index _ _ (equivocators_fixed_equivocations_constraint IM Hbs index_listing [])) in HtrX @@ -466,15 +466,15 @@ Proof. exists s. split; [assumption|]. exists tr. split; [assumption|]. split; [|assumption]. - apply (VLSM_eq_finite_protocol_trace_init_to HeqXE). + apply (VLSM_eq_finite_valid_trace_init_to HeqXE). assumption. - intro; intros. - apply (VLSM_eq_protocol_state HeqXE) in Hes. + apply (VLSM_eq_valid_state HeqXE) in Hes. split. + split; [|exact I]. destruct om as [im|]; [|exact I]. destruct Hom as [Hom|Hinitial]; [left; assumption|exfalso]. apply no_initial_messages_in_XE in Hinitial. contradiction. - + apply protocol_has_fixed_equivocation in Hes. + + apply valid_state_has_fixed_equivocation in Hes. destruct (composite_transition _ _ _) as (es', om') eqn:Het. simpl. apply @@ -496,14 +496,14 @@ Proof. split; [constructor; assumption|]. split; [reflexivity|]. split; [assumption|]. - apply (VLSM_eq_protocol_state HeqXE) in Hstate_protocol. - apply (VLSM_incl_protocol_state HinclE) in Hstate_protocol. + apply (VLSM_eq_valid_state HeqXE) in Hstate_valid. + apply (VLSM_incl_valid_state HinclE) in Hstate_valid. left. revert Hsent. subst. specialize (VLSM_projection_has_been_sent_reflect (preloaded_equivocators_no_equivocations_vlsm_X_vlsm_projection IM Hbs finite_index) - eqv_state_s Hstate_protocol im) as Hsent. + eqv_state_s Hstate_valid im) as Hsent. unfold has_been_sent in Hsent. simpl in Hsent. assumption. Qed. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v index ef83a5bf..e97d2f34 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v @@ -122,15 +122,15 @@ Proof. intros l (s,om) Hv. apply Hv. Qed. -(** A protocol state for a VLSM satisfying the limited equivocation assumption +(** A valid state for a VLSM satisfying the limited equivocation assumption has limited equivocation. *) -Lemma protocol_state_limited_equivocation +Lemma valid_state_limited_equivocation (s : composite_state equivocator_IM) - (Hs : protocol_state_prop equivocators_limited_equivocations_vlsm s) + (Hs : valid_state_prop equivocators_limited_equivocations_vlsm s) : not_heavy s. Proof. - apply protocol_state_prop_iff in Hs. + apply valid_state_prop_iff in Hs. destruct Hs as [[(is, His) Heq_s] | [l [(s0, oim) [oom' [[_ [_ [_ [_ Hlimited]]]] Ht]]]]]. - subst s. simpl. unfold not_heavy, equivocation_fault. replace (equivocating_validators is) with (@nil index). @@ -144,15 +144,15 @@ Proof. simpl in *. rewrite Ht. reflexivity. Qed. -(** A valid protocol trace for the composition of equivocators with limited -state-equivocation and no message-equivocation is also a valid protocol trace +(** A valid valid trace for the composition of equivocators with limited +state-equivocation and no message-equivocation is also a valid valid trace for the composition of equivocators with no message-equivocation and fixed-set state-equivocation, where the fixed set is given by the state-equivocators measured for the final state of the trace. *) -Lemma equivocators_limited_protocol_trace_is_fixed is s tr - : finite_protocol_trace_init_to equivocators_limited_equivocations_vlsm is s tr -> - finite_protocol_trace_init_to +Lemma equivocators_limited_valid_trace_is_fixed is s tr + : finite_valid_trace_init_to equivocators_limited_equivocations_vlsm is s tr -> + finite_valid_trace_init_to (equivocators_fixed_equivocations_vlsm IM Hbs index_listing (equivocating_validators s)) is s tr. Proof. @@ -160,16 +160,16 @@ Proof. split; [| apply H]. cut (forall equivocating, equivocating_validators s ⊆ equivocating -> - finite_protocol_trace_from_to (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating) is s tr). + finite_valid_trace_from_to (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating) is s tr). { intros H'. apply H'. reflexivity. } - induction H using finite_protocol_trace_init_to_rev_ind; intros equivocating Hincl. - - apply (finite_ptrace_from_to_empty (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating)). - apply initial_is_protocol. assumption. + induction H using finite_valid_trace_init_to_rev_ind; intros equivocating Hincl. + - apply (finite_valid_trace_from_to_empty (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating)). + apply initial_state_is_valid. assumption. - specialize (equivocating_indices_equivocating_validators IM _ finite_index _ reachable_threshold) as Heq. destruct (Heq sf) as [_ Hsf_incl]. - specialize (IHfinite_protocol_trace_init_to equivocating). - spec IHfinite_protocol_trace_init_to. + specialize (IHfinite_valid_trace_init_to equivocating). + spec IHfinite_valid_trace_init_to. { apply proj2 in Ht. specialize (equivocators_transition_preserves_equivocating_indices IM index_listing _ _ _ _ _ Ht) as Hincl'. @@ -180,34 +180,34 @@ Proof. transitivity (equivocating_indices IM index_listing s); assumption. } apply - (finite_protocol_trace_from_to_app + (finite_valid_trace_from_to_app (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating)) with s; [assumption|]. - apply ptrace_add_last; [|reflexivity]. - apply (finite_ptrace_singleton (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating)). - apply ptrace_last_pstate in IHfinite_protocol_trace_init_to. + apply valid_trace_add_last; [|reflexivity]. + apply (finite_valid_trace_singleton (equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating)). + apply valid_trace_last_pstate in IHfinite_valid_trace_init_to. destruct Ht as [[_ [_ [Hv [[Hno_equiv _] Hno_heavy]]]] Ht]. repeat split; [assumption| |assumption|assumption| |assumption]. - + destruct iom as [m|]; [|apply option_protocol_message_None]. + + destruct iom as [m|]; [|apply option_valid_message_None]. destruct Hno_equiv as [Hsent | Hfalse]; [|contradiction]. simpl in Hsent. - apply composite_sent_protocol with index_listing (equivocator_Hbs IM Hbs) s; assumption. + apply composite_sent_valid with index_listing (equivocator_Hbs IM Hbs) s; assumption. + replace (composite_transition _ _ _) with (sf, oom). unfold state_has_fixed_equivocation. transitivity (equivocating_validators sf); assumption. Qed. -(** Projections of valid protocol traces for the composition of equivocators +(** Projections of valid traces for the composition of equivocators with limited state-equivocation and no message-equivocation have the [fixed_limited_equivocation_prop]erty. *) -Lemma equivocators_limited_protocol_trace_projects_to_fixed_limited_equivocation +Lemma equivocators_limited_valid_trace_projects_to_fixed_limited_equivocation (final_descriptors : equivocator_descriptors) (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) (final_state := finite_trace_last is tr) (Hproper: not_equivocating_equivocator_descriptors IM final_descriptors final_state) - (Htr : finite_protocol_trace equivocators_limited_equivocations_vlsm is tr) + (Htr : finite_valid_trace equivocators_limited_equivocations_vlsm is tr) : exists (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors) @@ -218,12 +218,12 @@ Lemma equivocators_limited_protocol_trace_projects_to_fixed_limited_equivocation equivocators_state_project final_descriptors final_state = final_stateX /\ fixed_limited_equivocation_prop IM Hbs Hbr isX trX. Proof. - apply ptrace_add_default_last in Htr as Hfixed_tr. - apply equivocators_limited_protocol_trace_is_fixed in Hfixed_tr. - apply ptrace_last_pstate in Hfixed_tr as Hfixed_last. - apply ptrace_forget_last in Hfixed_tr. + apply valid_trace_add_default_last in Htr as Hfixed_tr. + apply equivocators_limited_valid_trace_is_fixed in Hfixed_tr. + apply valid_trace_last_pstate in Hfixed_tr as Hfixed_last. + apply valid_trace_forget_last in Hfixed_tr. specialize - (fixed_equivocators_protocol_trace_project IM Hbs Hbr (equivocating_validators (finite_trace_last is tr)) + (fixed_equivocators_valid_trace_project IM Hbs Hbr (equivocating_validators (finite_trace_last is tr)) finite_index final_descriptors is tr) as Hpr. feed specialize Hpr; [| assumption |]. - eapply not_equivocating_equivocator_descriptors_proper_fixed; eassumption. @@ -233,7 +233,7 @@ Proof. + apply Hinitial_descriptors. + exists (equivocating_validators (finite_trace_last is tr)). split; [| assumption]. - apply ptrace_add_default_last, ptrace_last_pstate, protocol_state_limited_equivocation in Htr. + apply valid_trace_add_default_last, valid_trace_last_pstate, valid_state_limited_equivocation in Htr. unfold not_heavy in Htr. transitivity (equivocation_fault (finite_trace_last is tr)); [|assumption]. unfold equivocation_fault. @@ -255,18 +255,18 @@ Context . (** If each of the nodes satisfy the [message_dependencies_full_node_condition_prop]erty, -then projections of valid protocol traces for the composition of equivocators +then projections of valid traces for the composition of equivocators with limited state-equivocation and no message-equivocation are also valid -protocol traces for the composition of regular nodes with limited +traces for the composition of regular nodes with limited message-equivocation. *) -Lemma limited_equivocators_protocol_trace_project +Lemma limited_equivocators_valid_trace_project (final_descriptors : equivocator_descriptors) (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) (final_state := finite_trace_last is tr) (Hproper: not_equivocating_equivocator_descriptors IM final_descriptors final_state) - (Htr : finite_protocol_trace equivocators_limited_equivocations_vlsm is tr) + (Htr : finite_valid_trace equivocators_limited_equivocations_vlsm is tr) : exists (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors) @@ -275,15 +275,15 @@ Lemma limited_equivocators_protocol_trace_project proper_equivocator_descriptors initial_descriptors is /\ equivocators_trace_project IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project final_descriptors final_state = final_stateX /\ - finite_protocol_trace Limited isX trX. + finite_valid_trace Limited isX trX. Proof. specialize - (equivocators_limited_protocol_trace_projects_to_fixed_limited_equivocation + (equivocators_limited_valid_trace_projects_to_fixed_limited_equivocation final_descriptors is tr Hproper Htr) as [trX [initial_descriptors [Hinitial_descriptors [Hpr [Hlst_pr Hpr_limited]]]]]. exists trX, initial_descriptors. repeat split; [assumption..| |]. - - eapply trace_exhibits_limited_equivocation_protocol; eassumption. + - eapply traces_exhibiting_limited_equivocation_are_valid; eassumption. - destruct Hpr_limited as [equivs Hpr_limited]. apply Hpr_limited. Qed. @@ -298,8 +298,8 @@ Proof. split; [split|]. - intros s tr sX trX Hpr_tr s_pre pre Hs_lst Hpre_tr. assert - (HPreFree_pre_tr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm FreeE) s_pre (pre ++ tr)). - { revert Hpre_tr. apply VLSM_incl_finite_protocol_trace_from. + (HPreFree_pre_tr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm FreeE) s_pre (pre ++ tr)). + { revert Hpre_tr. apply VLSM_incl_finite_valid_trace_from. apply equivocators_limited_equivocations_vlsm_incl_preloaded_free. } clear Hpre_tr. revert s tr sX trX Hpr_tr s_pre pre Hs_lst HPreFree_pre_tr. @@ -308,7 +308,7 @@ Proof. destruct (destruct_equivocators_partial_trace_project IM finite_index Hpr_tr) as [Hnot_equiv [initial_descriptors [Htr_project Hs_project]]]. - destruct (limited_equivocators_protocol_trace_project _ _ _ Hnot_equiv Htr) + destruct (limited_equivocators_valid_trace_project _ _ _ Hnot_equiv Htr) as [_trX [_initial_descriptors [_ [_Htr_project [_ HtrX]]]]]. rewrite Htr_project in _Htr_project. inversion _Htr_project. subst. assumption. @@ -323,14 +323,14 @@ Lemma limited_equivocators_vlsm_projection Proof. constructor; [constructor|]; intros. - apply PreFreeE_Free_vlsm_projection_type. - revert H. apply VLSM_incl_finite_protocol_trace_from. + revert H. apply VLSM_incl_finite_valid_trace_from. apply equivocators_limited_equivocations_vlsm_incl_preloaded_free. - - assert (Hpre_tr : finite_protocol_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). - { revert H. apply VLSM_incl_finite_protocol_trace. + - assert (Hpre_tr : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). + { revert H. apply VLSM_incl_finite_valid_trace. apply equivocators_limited_equivocations_vlsm_incl_preloaded_free. } specialize - (VLSM_partial_projection_finite_protocol_trace (limited_equivocators_vlsm_partial_projection (zero_descriptor IM)) + (VLSM_partial_projection_finite_valid_trace (limited_equivocators_vlsm_partial_projection (zero_descriptor IM)) sX trX (equivocators_state_project (zero_descriptor IM) sX) (equivocators_total_trace_project IM trX)) as Hsim. spec Hsim. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v index 8f07c931..3dcfe774 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v @@ -16,9 +16,9 @@ In this module we show that the composition of equivocators with no-message equivocation and limited state-equivocation can simulate all traces with the [fixed_limited_equivocation_prop]erty. -As a corollary we show that any state which is protocol for the composition +As a corollary we show that any state which is valid for the composition of regular nodes using a [limited_equivocation_constraint] can be -obtained as the projection of a protocol state for the composition of +obtained as the projection of a valid state for the composition of equivocators with no message equivocation and limited state equivocation. *) @@ -41,8 +41,8 @@ Context . (** If the total weight of the equivocators allowed to state-equivocate is less -than the [threshold], then traces of equivocators which are protocol w.r.t -the fixed state-equivocation constraint are also protocol w.r.t. the +than the [threshold], then traces of equivocators which are valid w.r.t +the fixed state-equivocation constraint are also valid w.r.t. the limited state-equivocation constraint. *) Lemma equivocators_Fixed_incl_Limited @@ -90,24 +90,24 @@ Context simulated by the composition of equivocators with no message-equivocation and limited state-equivocation. *) -Lemma limited_equivocators_finite_protocol_trace_init_to_rev +Lemma limited_equivocators_finite_valid_trace_init_to_rev (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) isX trX (HtrX : fixed_limited_equivocation_prop IM _ _ isX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = finite_trace_last isX trX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to XE is s tr /\ + finite_valid_trace_init_to XE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. destruct HtrX as [equivocating [Hlimited HtrX]]. specialize (Fixed_eq_StrongFixed IM Hbs Hbr finite_index equivocating) as Heq. - apply (VLSM_eq_finite_protocol_trace Heq) in HtrX. + apply (VLSM_eq_finite_valid_trace Heq) in HtrX. clear Heq. - apply ptrace_add_default_last in HtrX. + apply valid_trace_add_default_last in HtrX. specialize - (fixed_equivocators_finite_protocol_trace_init_to_rev IM Hbs finite_index _ + (fixed_equivocators_finite_valid_trace_init_to_rev IM Hbs finite_index _ no_initial_messages_in_IM _ _ _ HtrX) as [is [His [s [Hs [tr [Htr [Hptr Houtput]]]]]]]. exists is. split; [assumption|]. @@ -115,7 +115,7 @@ Proof. exists tr. split; [assumption|]. split; [|assumption]. revert Hptr. - apply VLSM_incl_finite_protocol_trace_init_to. + apply VLSM_incl_finite_valid_trace_init_to. apply equivocators_Fixed_incl_Limited. assumption. Qed. @@ -127,34 +127,34 @@ Context (message_dependencies : message -> set message) . -(** Any protocol state for the composition of reqular nodes under a limited -message-equivocation constraint is the projection of a protocol state of +(** Any valid state for the composition of reqular nodes under a limited +message-equivocation constraint is the projection of a valid state of the composition of equivocators under a no message-equivocation and limited state-equivocation constraint. *) -Lemma limited_equivocators_protocol_state_rev +Lemma limited_equivocators_valid_state_rev (Hwitnessed_equivocation : WitnessedEquivocationCapability IM Datatypes.id sender finite_index) (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) (can_emit_signed : channel_authentication_prop IM Datatypes.id sender) - : forall sX, protocol_state_prop Limited sX -> + : forall sX, valid_state_prop Limited sX -> exists s, equivocators_total_state_project IM s = sX /\ - protocol_state_prop XE s. + valid_state_prop XE s. Proof. intros sX HsX. apply - (limited_protocol_state_has_trace_exhibiting_limited_equivocation IM Hbs Hbr finite_index + (limited_valid_state_has_trace_exhibiting_limited_equivocation IM Hbs Hbr finite_index sender message_dependencies Hwitnessed_equivocation HMsgDep Hfull no_initial_messages_in_IM can_emit_signed) in HsX as [isX [trX [HsX HtrX]]]. - apply limited_equivocators_finite_protocol_trace_init_to_rev in HtrX + apply limited_equivocators_finite_valid_trace_init_to_rev in HtrX as [is [_ [s [Hpr_s [tr [_ [Htr _]]]]]]] ; [|assumption]. exists s. subst. split; [assumption|]. - apply ptrace_last_pstate in Htr. + apply valid_trace_last_pstate in Htr. assumption. Qed. diff --git a/theories/VLSM/Core/Equivocators/Composition/Projections.v b/theories/VLSM/Core/Equivocators/Composition/Projections.v index 7e8c34bf..048df4ce 100644 --- a/theories/VLSM/Core/Equivocators/Composition/Projections.v +++ b/theories/VLSM/Core/Equivocators/Composition/Projections.v @@ -737,7 +737,7 @@ Lemma equivocators_trace_project_preserves_equivocating_indices (tr : list (composite_transition_item equivocator_IM)) (trX : list (composite_transition_item IM)) (is s : composite_state equivocator_IM) - (Htr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) + (Htr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) (Hdescriptors : proper_equivocator_descriptors descriptors s) (Hproject_tr : equivocators_trace_project descriptors tr = Some (trX, idescriptors)) : @@ -745,7 +745,7 @@ Lemma equivocators_trace_project_preserves_equivocating_indices (set_union (equivocating_indices IM index_listing s) (newmachine_descriptors_list IM index_listing descriptors)). Proof. generalize dependent trX. generalize dependent descriptors. - induction Htr using finite_protocol_trace_from_to_rev_ind. + induction Htr using finite_valid_trace_from_to_rev_ind. - intros. inversion Hproject_tr. intros eqv Heqv. assumption. - set (x:={|l:=l|}). intros. @@ -773,14 +773,14 @@ Proof. Qed. (** The state and descriptors obtained after applying [equivocators_trace_project] -on a pre-loaded protocol trace satisfy the [previous_state_descriptor_prop]erty. +on a pre-loaded valid trace satisfy the [previous_state_descriptor_prop]erty. *) Lemma equivocators_trace_project_from_state_descriptors (descriptors idescriptors : equivocator_descriptors) (tr : list (composite_transition_item equivocator_IM)) (trX : list (composite_transition_item IM)) (is s : composite_state equivocator_IM) - (Htr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) + (Htr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) (Hdescriptors : proper_equivocator_descriptors descriptors s) (Hproject_tr : equivocators_trace_project descriptors tr = Some (trX, idescriptors)) : forall eqv, previous_state_descriptor_prop (IM eqv) (descriptors eqv) (is eqv) (idescriptors eqv). @@ -790,9 +790,9 @@ Proof. generalize dependent s. induction tr using rev_ind; intros. - inversion Hproject_tr. subst. destruct (idescriptors eqv); simpl; [reflexivity|lia]. - - apply finite_protocol_trace_from_to_last in Htr as Heq_s. + - apply finite_valid_trace_from_to_last in Htr as Heq_s. rewrite finite_trace_last_is_last in Heq_s. subst s. - apply finite_protocol_trace_from_to_app_split in Htr. + apply finite_valid_trace_from_to_app_split in Htr. destruct Htr as [Htr Hx]. specialize (equivocators_pre_trace_cannot_decrease_state_size IM _ _ _ Htr) as His_tr. specialize (equivocators_pre_trace_cannot_decrease_state_size IM _ _ _ Hx) as Htr_x. @@ -835,7 +835,7 @@ Lemma equivocators_trace_project_preserves_equivocating_indices_final (tr : list (composite_transition_item equivocator_IM)) (trX : list (composite_transition_item IM)) (is s : composite_state equivocator_IM) - (Htr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) + (Htr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is s tr ) (Hdescriptors : not_equivocating_equivocator_descriptors IM descriptors s) (Hproject_tr : equivocators_trace_project descriptors tr = Some (trX, idescriptors)) : @@ -970,7 +970,7 @@ to full (valid) traces Lemma equivocators_trace_project_preserves_zero_descriptors (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) - (Htr : finite_protocol_trace_from PreFreeE is tr) + (Htr : finite_valid_trace_from PreFreeE is tr) (descriptors : equivocator_descriptors) (idescriptors : equivocator_descriptors) (trX : list (composite_transition_item IM)) @@ -978,7 +978,7 @@ Lemma equivocators_trace_project_preserves_zero_descriptors : forall i, descriptors i = Existing 0 -> idescriptors i = Existing 0. Proof. generalize dependent trX. generalize dependent descriptors. - induction Htr using finite_protocol_trace_from_rev_ind. + induction Htr using finite_valid_trace_from_rev_ind. - intros. inversion HtrX. subst. assumption. - intros. apply equivocators_trace_project_app_iff in HtrX. @@ -997,13 +997,13 @@ Proof. apply (IHHtr _ _ Hproject_tr). assumption. Qed. -Lemma preloaded_equivocators_protocol_trace_from_project +Lemma preloaded_equivocators_valid_trace_from_project (final_descriptors : equivocator_descriptors) (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) (final_state := finite_trace_last is tr) (Hproper : proper_equivocator_descriptors final_descriptors final_state) - (Htr : finite_protocol_trace_from PreFreeE is tr) + (Htr : finite_valid_trace_from PreFreeE is tr) : exists (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors), @@ -1016,7 +1016,7 @@ Proof. generalize dependent is. induction tr using rev_ind; intros. - exists []. simpl. exists final_descriptors. repeat split; assumption. - - apply finite_protocol_trace_from_app_iff in Htr. + - apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hx]. specialize (IHtr _ Htr). specialize (equivocators_transition_item_project_proper_characterization final_descriptors x) as Hproperx. @@ -1056,12 +1056,12 @@ Qed. Lemma equivocators_trace_project_zero_descriptors (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) - (Htr : finite_protocol_trace_from PreFreeE is tr) + (Htr : finite_valid_trace_from PreFreeE is tr) : exists (trX : list (composite_transition_item IM)), equivocators_trace_project (zero_descriptor IM) tr = Some (trX, (zero_descriptor IM)). Proof. specialize - (preloaded_equivocators_protocol_trace_from_project + (preloaded_equivocators_valid_trace_from_project (zero_descriptor IM) is tr ) as Hproject. simpl in Hproject. spec Hproject. { apply zero_descriptor_proper. } @@ -1074,12 +1074,12 @@ Proof. reflexivity. Qed. -Lemma preloaded_equivocators_protocol_trace_project_inv +Lemma preloaded_equivocators_valid_trace_project_inv (final_descriptors : equivocator_descriptors) (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) (final_state := finite_trace_last is tr) - (Htr : finite_protocol_trace PreFreeE is tr) + (Htr : finite_valid_trace PreFreeE is tr) (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors) (Hproject: equivocators_trace_project final_descriptors tr = Some (trX, initial_descriptors)) @@ -1089,7 +1089,7 @@ Proof. revert Hproject. revert trX Htr final_descriptors. induction tr using rev_ind; intros; [inversion Hproject; assumption|]. destruct Htr as [Htr Hinit]. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hx]. unfold equivocators_trace_project in Hproject. rewrite fold_right_app in Hproject. @@ -1170,14 +1170,14 @@ Proof. Qed. (** -A corrollary of [preloaded_equivocators_protocol_trace_from_project] selecting +A corrollary of [preloaded_equivocators_valid_trace_from_project] selecting only the [proper_equivocator_descriptors] property. *) -Lemma preloaded_equivocators_protocol_trace_project_proper_initial +Lemma preloaded_equivocators_valid_trace_project_proper_initial (is : composite_state equivocator_IM) (tr : list (composite_transition_item equivocator_IM)) (final_state := finite_trace_last is tr) - (Htr : finite_protocol_trace_from PreFreeE is tr) + (Htr : finite_valid_trace_from PreFreeE is tr) (final_descriptors : equivocator_descriptors) (trX : list (composite_transition_item IM)) (initial_descriptors : equivocator_descriptors) @@ -1186,7 +1186,7 @@ Lemma preloaded_equivocators_protocol_trace_project_proper_initial : proper_equivocator_descriptors initial_descriptors is. Proof. destruct - (preloaded_equivocators_protocol_trace_from_project + (preloaded_equivocators_valid_trace_from_project final_descriptors is tr Hproper Htr ) as [_trX [_initial_descriptors [_Hproject [Hiproper _]]]]. @@ -1198,7 +1198,7 @@ Qed. Lemma equivocators_trace_project_output_reflecting_inv (is: composite_state equivocator_IM) (tr: list (composite_transition_item equivocator_IM)) - (Htr: finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is tr) + (Htr: finite_valid_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is tr) (m : message) (Hbbs : Exists (field_selector output m) tr) : exists @@ -1213,7 +1213,7 @@ Proof. apply (finite_trace_projection_list_in equivocator_IM) in Hitem. destruct item. simpl in *. destruct l as (i, li). simpl in *. specialize - (preloaded_finite_ptrace_projection equivocator_IM i _ _ Htr) + (preloaded_finite_valid_trace_projection equivocator_IM i _ _ Htr) as Htri. specialize (equivocator_vlsm_trace_project_output_reflecting_inv (IM i) _ _ Htri m) as Hex. @@ -1245,7 +1245,7 @@ Proof. subst final. assert (Hfinal_descriptors_proper : proper_equivocator_descriptors final_descriptors (finite_trace_last is tr)). { apply not_equivocating_equivocator_descriptors_proper. assumption. } - destruct (preloaded_equivocators_protocol_trace_from_project _ _ _ Hfinal_descriptors_proper Htr) + destruct (preloaded_equivocators_valid_trace_from_project _ _ _ Hfinal_descriptors_proper Htr) as [trX [initial_descriptors [Hproject_tr _]]]. exists initial_descriptors, trX. split; [assumption|]. split; [assumption|]. specialize @@ -1268,7 +1268,7 @@ Qed. Lemma equivocators_trace_project_output_reflecting_iff (is: composite_state equivocator_IM) (tr: list (composite_transition_item equivocator_IM)) - (Htr: finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is tr) + (Htr: finite_valid_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm equivocator_IM)) is tr) (m : message) : Exists (field_selector output m) tr <-> exists @@ -1294,14 +1294,14 @@ Proof. simpl in *. congruence. Qed. -(** Projecting a pre-loaded protocol trace of the composition of equivocators -using [proper_equivocator_descriptors] one obtains a pre-loaded protocol trace +(** Projecting a pre-loaded valid trace of the composition of equivocators +using [proper_equivocator_descriptors] one obtains a pre-loaded valid trace of the free composition of nodes. *) -Lemma pre_equivocators_protocol_trace_project +Lemma pre_equivocators_valid_trace_project (is final_state : vstate equivocators_no_equivocations_vlsm) (tr : list (composite_transition_item equivocator_IM)) - (Htr : finite_protocol_trace_init_to PreFreeE is final_state tr) + (Htr : finite_valid_trace_init_to PreFreeE is final_state tr) (final_descriptors : equivocator_descriptors) (Hproper : proper_equivocator_descriptors final_descriptors final_state) : exists @@ -1312,27 +1312,27 @@ Lemma pre_equivocators_protocol_trace_project (final_stateX := equivocators_state_project final_descriptors final_state) (trX : list (composite_transition_item IM)), equivocators_trace_project final_descriptors tr = Some (trX, initial_descriptors) /\ - finite_protocol_trace_init_to PreFree isX final_stateX trX. + finite_valid_trace_init_to PreFree isX final_stateX trX. Proof. generalize dependent final_descriptors. generalize dependent final_state. induction tr using rev_ind; intros. - - apply ptrace_get_last in Htr as Hfinal_state_eq. + - apply valid_trace_get_last in Htr as Hfinal_state_eq. subst. exists final_descriptors. split; [assumption|]. exists []. repeat (split; [reflexivity|]). cut (vinitial_state_prop (free_composite_vlsm IM) (equivocators_state_project final_descriptors is)). { intro Hinit. split; [|assumption]. constructor. - apply initial_is_protocol. assumption. + apply initial_state_is_valid. assumption. } apply (equivocators_initial_state_project IM); [|assumption]. apply Htr. - destruct Htr as [Htr Hinit]. - apply finite_protocol_trace_from_to_app_split in Htr. + apply finite_valid_trace_from_to_app_split in Htr. destruct Htr as [Htr Hx]. specialize (IHtr _ (conj Htr Hinit)). - apply finite_protocol_trace_from_to_last in Hx as Hfinal_state_eq. + apply finite_valid_trace_from_to_last in Hx as Hfinal_state_eq. change [x] with ([] ++ [x]) in Hfinal_state_eq. rewrite finite_trace_last_is_last in Hfinal_state_eq. subst. @@ -1360,7 +1360,7 @@ Proof. repeat split. assumption. * apply - (finite_protocol_trace_from_to_app PreFree + (finite_valid_trace_from_to_app PreFree (equivocators_state_project final_descriptors' (finite_trace_last is tr))) ; [assumption|]. specialize (Hchar2 _ eq_refl). @@ -1372,10 +1372,10 @@ Proof. simpl_existT. subst lix input output. destruct Hchar2 as [Hvx_pr Htx_pr]. rewrite Hpr_s. - apply finite_ptrace_from_to_singleton. + apply finite_valid_trace_from_to_singleton. repeat split - ; [|apply any_message_is_protocol_in_preloaded|assumption|assumption]. - apply finite_protocol_trace_from_to_last_pstate in HtrX. assumption. + ; [|apply any_message_is_valid_in_preloaded|assumption|assumption]. + apply finite_valid_trace_from_to_last_pstate in HtrX. assumption. + exists trX. clear Hchar1. rewrite Hchar2. split; [|assumption]. apply (Hpr_app trX). @@ -1455,7 +1455,7 @@ Lemma equivocators_partial_trace_project_extends_left partial_trace_project (sX, trX) = Some (sY, trY) -> forall s'X preX, finite_trace_last s'X preX = sX -> - finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm X) s'X (preX ++ trX) -> + finite_valid_trace_from (pre_loaded_with_all_messages_vlsm X) s'X (preX ++ trX) -> exists s'Y preY, partial_trace_project (s'X, preX ++ trX) = Some (s'Y, preY ++ trY) /\ finite_trace_last s'Y preY = sY. @@ -1463,15 +1463,15 @@ Proof. intros s tr sX trX Hpr_tr s_pre pre Hs_lst Hpre_tr. destruct (destruct_equivocators_partial_trace_project Hpr_tr) as [Hnot_equiv [initial_descriptors [Htr_project Hs_project]]]. - apply (finite_protocol_trace_from_app_iff PreFreeE) in Hpre_tr. + apply (finite_valid_trace_from_app_iff PreFreeE) in Hpre_tr. destruct Hpre_tr as [Hpre Htr]. subst s sX. apply not_equivocating_equivocator_descriptors_proper in Hnot_equiv as Hproper. specialize - (preloaded_equivocators_protocol_trace_project_proper_initial _ _ Htr + (preloaded_equivocators_valid_trace_project_proper_initial _ _ Htr _ _ _ Htr_project Hproper ) as Hinitial_descriptors. destruct - (preloaded_equivocators_protocol_trace_from_project + (preloaded_equivocators_valid_trace_from_project _ _ _ Hinitial_descriptors Hpre ) as [preX [pre_descriptors [Hpre_project [Hpre_desciptors Hs_project]]]]. exists (equivocators_state_project pre_descriptors s_pre), preX. @@ -1503,7 +1503,7 @@ which is guaranteed to always succeed. *) Lemma equivocators_total_trace_project_characterization {s tr} - (Hpre_tr : finite_protocol_trace_from PreFreeE s tr) + (Hpre_tr : finite_valid_trace_from PreFreeE s tr) : equivocators_trace_project (zero_descriptor IM) tr = Some (equivocators_total_trace_project tr, zero_descriptor IM). Proof. unfold equivocators_total_trace_project. @@ -1517,7 +1517,7 @@ Lemma equivocators_total_trace_project_app (X := FreeE) (trace_project := equivocators_total_trace_project) : forall tr1X tr2X, - (exists sX, finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm X) sX (tr1X ++ tr2X)) -> + (exists sX, finite_valid_trace_from (pre_loaded_with_all_messages_vlsm X) sX (tr1X ++ tr2X)) -> trace_project (tr1X ++ tr2X) = trace_project tr1X ++ trace_project tr2X. Proof. intros. destruct H as [sX Hpre_tr]. @@ -1525,7 +1525,7 @@ Proof. apply equivocators_trace_project_app_iff in Htr12_pr. destruct Htr12_pr as [tr1Y [tr2Y [descriptors [Htr2_pr [Htr1_pr Htr12_eq]]]]]. - apply (finite_protocol_trace_from_app_iff PreFreeE) in Hpre_tr. + apply (finite_valid_trace_from_app_iff PreFreeE) in Hpre_tr. destruct Hpre_tr as [Hpre_tr1 Hpre_tr2]. rewrite (equivocators_total_trace_project_characterization Hpre_tr2) in Htr2_pr. inversion Htr2_pr. subst. clear Htr2_pr. @@ -1535,14 +1535,14 @@ Qed. Lemma equivocators_total_VLSM_projection_trace_project {s tr} - (Hpre_tr : finite_protocol_trace_from PreFreeE s tr) + (Hpre_tr : finite_valid_trace_from PreFreeE s tr) : @pre_VLSM_projection_trace_project _ (type PreFreeE) _ equivocators_total_label_project equivocators_total_state_project tr = equivocators_total_trace_project tr. Proof. induction tr using rev_ind; [reflexivity|]. rewrite equivocators_total_trace_project_app by (eexists; exact Hpre_tr). rewrite @pre_VLSM_projection_trace_project_app. - apply finite_protocol_trace_from_app_iff in Hpre_tr as [Hpre_tr Hpre_x]. + apply finite_valid_trace_from_app_iff in Hpre_tr as [Hpre_tr Hpre_x]. specialize (IHtr Hpre_tr). rewrite IHtr. f_equal. @@ -1573,13 +1573,13 @@ Lemma equivocators_total_trace_project_final_state (state_project := equivocators_total_state_project) (trace_project := equivocators_total_trace_project) : forall sX trX, - finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm X) sX trX -> + finite_valid_trace_from (pre_loaded_with_all_messages_vlsm X) sX trX -> state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (trace_project trX). Proof. intros sX trX Hpre_tr. specialize (equivocators_total_trace_project_characterization Hpre_tr) as Htr_pr. specialize - (preloaded_equivocators_protocol_trace_from_project (zero_descriptor IM) sX trX) + (preloaded_equivocators_valid_trace_from_project (zero_descriptor IM) sX trX) as Hproject. simpl in Hproject. spec Hproject. { apply zero_descriptor_proper. } spec Hproject Hpre_tr. @@ -1597,13 +1597,13 @@ Proof. - intros s tr sX trX Hpr_tr Htr. destruct (destruct_equivocators_partial_trace_project Hpr_tr) as [Hnot_equiv [initial_descriptors [Htr_project Hs_project]]]. - apply ptrace_add_default_last in Htr. + apply valid_trace_add_default_last in Htr. apply not_equivocating_equivocator_descriptors_proper in Hnot_equiv as Hproper. - destruct (pre_equivocators_protocol_trace_project _ _ _ Htr _ Hproper) + destruct (pre_equivocators_valid_trace_project _ _ _ Htr _ Hproper) as [_initial_descriptors [_ [_trX [_Htr_project HtrX]]]]. rewrite Htr_project in _Htr_project. inversion _Htr_project. subst. - apply ptrace_forget_last in HtrX. assumption. + apply valid_trace_forget_last in HtrX. assumption. Qed. End equivocators_composition_projections. @@ -1799,7 +1799,7 @@ Proof. apply finite_trace_sub_projection_app. Qed. -Section seeded_equivocators_protocol_trace_project. +Section seeded_equivocators_valid_trace_project. Context (seed : message -> Prop) @@ -1829,10 +1829,10 @@ Proof. apply seeded_no_equivocation_incl_preloaded. Qed. -Lemma seeded_equivocators_protocol_trace_project +Lemma seeded_equivocators_valid_trace_project (is : composite_state sub_equivocator_IM) (tr : list (composite_transition_item sub_equivocator_IM)) - (Htr : finite_protocol_trace SeededXE is tr) + (Htr : finite_valid_trace SeededXE is tr) (final_state := finite_trace_last is tr) (final_descriptors : (equivocator_descriptors sub_IM)) (Hproper : proper_equivocator_descriptors sub_IM final_descriptors final_state) @@ -1844,24 +1844,24 @@ Lemma seeded_equivocators_protocol_trace_project proper_equivocator_descriptors sub_IM initial_descriptors is /\ equivocators_trace_project sub_IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project sub_IM final_descriptors final_state = final_stateX /\ - finite_protocol_trace SeededX isX trX. + finite_valid_trace SeededX isX trX. Proof. - assert (Htr_to : finite_protocol_trace_init_to SeededXE is final_state tr). + assert (Htr_to : finite_valid_trace_init_to SeededXE is final_state tr). { destruct Htr as [Htr Hinit]. split; [|assumption]. - apply finite_protocol_trace_from_add_last; [assumption|reflexivity]. + apply finite_valid_trace_from_add_last; [assumption|reflexivity]. } - assert (Hpre_tr_to : finite_protocol_trace_init_to SubPreFreeE is final_state tr). - { revert Htr_to. apply VLSM_incl_finite_protocol_trace_init_to. + assert (Hpre_tr_to : finite_valid_trace_init_to SubPreFreeE is final_state tr). + { revert Htr_to. apply VLSM_incl_finite_valid_trace_init_to. apply seeded_no_equivocation_incl_preloaded. } - apply pre_equivocators_protocol_trace_project + apply pre_equivocators_valid_trace_project with (final_descriptors0 := final_descriptors) in Hpre_tr_to ; [|assumption..]. destruct Hpre_tr_to as [initial_descriptors [Hproper_initial [trX [Hpr_trX Hpre_trX]]]]. exists trX, initial_descriptors. split; [assumption|]. split; [assumption|]. - apply finite_protocol_trace_init_to_last in Hpre_trX as Hfinal_stateX. + apply finite_valid_trace_init_to_last in Hpre_trX as Hfinal_stateX. symmetry in Hfinal_stateX. split; [assumption|]. clear -SubPreFreeE Hbs_sub finite_index Htr Hproper Hpr_trX. @@ -1875,7 +1875,7 @@ Proof. - clear H. subst. subst final_state. simpl in *. inversion Hpr_trX. subst. cut (vinitial_state_prop SubFree (equivocators_state_project sub_IM initial_descriptors is)). { intro. split; [|assumption]. constructor. - apply protocol_state_prop_iff. left. + apply valid_state_prop_iff. left. exists (exist _ _ H). reflexivity. } apply equivocators_initial_state_project; [|assumption]. @@ -1883,7 +1883,7 @@ Proof. - specialize (H (length tr')) as H'. spec H'. { rewrite app_length. simpl. lia. } destruct Htr as [Htr Hinit]. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hlst]. specialize (H' tr' (conj Htr Hinit) eq_refl). specialize (equivocators_transition_item_project_proper_characterization sub_IM final_descriptors lst) as Hproperx. @@ -1907,12 +1907,12 @@ Proof. destruct Hproperx as [Hproper' [Heq_final_descriptors' [_ [_ Hx]]]]. specialize (H' _ Hproper' _ _ Hpr_trX'). destruct H' as [HtrX' HinitX]. - split; [|assumption]. apply finite_protocol_trace_from_app_iff. + split; [|assumption]. apply finite_valid_trace_from_app_iff. split; [assumption|]. assert (Hlst_trX' : - protocol_state_prop SeededX (finite_trace_last (equivocators_state_project sub_IM initial_descriptors is) trX')). - { apply (finite_ptrace_last_pstate SeededX) in HtrX'. + valid_state_prop SeededX (finite_trace_last (equivocators_state_project sub_IM initial_descriptors is) trX')). + { apply (finite_valid_trace_last_pstate SeededX) in HtrX'. assumption. } destruct oitem as [item|]; inversion _Hprojectx; subst lstX; clear _Hprojectx @@ -1921,16 +1921,16 @@ Proof. specialize (Hx _ eq_refl). destruct Hx as [Hvx Htx]. destruct item. simpl in *. subst. - apply finite_ptrace_singleton. - assert (Htr_to : finite_protocol_trace_init_to SeededXE is (finite_trace_last is tr') tr'). + apply finite_valid_trace_singleton. + assert (Htr_to : finite_valid_trace_init_to SeededXE is (finite_trace_last is tr') tr'). { split; [|assumption]. - apply finite_protocol_trace_from_add_last; [assumption|reflexivity]. + apply finite_valid_trace_from_add_last; [assumption|reflexivity]. } - assert (Hpre_tr_to : finite_protocol_trace_init_to SubPreFreeE is (finite_trace_last is tr') tr'). - { revert Htr_to. apply VLSM_incl_finite_protocol_trace_init_to. + assert (Hpre_tr_to : finite_valid_trace_init_to SubPreFreeE is (finite_trace_last is tr') tr'). + { revert Htr_to. apply VLSM_incl_finite_valid_trace_init_to. apply seeded_no_equivocation_incl_preloaded. } - apply pre_equivocators_protocol_trace_project + apply pre_equivocators_valid_trace_project with (final_descriptors0 := final_descriptors') in Hpre_tr_to as Hpr_tr' ; [|assumption..]. @@ -1938,21 +1938,21 @@ Proof. replace (equivocators_trace_project _ _ _) with (Some (trX', initial_descriptors)) in _Hpr_trX'. inversion _Hpr_trX'. subst _initial_descriptors _trX'. - apply finite_protocol_trace_init_to_last in Heq_final_stateX'. + apply finite_valid_trace_init_to_last in Heq_final_stateX'. simpl in *. rewrite <- Heq_final_stateX' in Htx, Hvx. repeat split; [assumption| |assumption|assumption]. destruct input as [input|] - ; [| apply option_protocol_message_None]. + ; [| apply option_valid_message_None]. apply proj1 in Hc. simpl in Hc. apply or_comm in Hc. destruct Hc as [Hinit_input | Hno_equiv] - ; [ apply initial_message_is_protocol; apply (seeded_equivocators_initial_message input); right; assumption + ; [ apply initial_message_is_valid; apply (seeded_equivocators_initial_message input); right; assumption |]. assert - (Hs_free : protocol_state_prop SubPreFreeE (finite_trace_last is tr')). - { apply proj1, finite_protocol_trace_from_to_last_pstate in Hpre_tr_to. + (Hs_free : valid_state_prop SubPreFreeE (finite_trace_last is tr')). + { apply proj1, finite_valid_trace_from_to_last_pstate in Hpre_tr_to. assumption. } apply @@ -1961,7 +1961,7 @@ Proof. (equivocator_Hbs sub_IM Hbs_sub) _ Hs_free) in Hno_equiv. specialize (Hno_equiv is tr' Hpre_tr_to). - apply finite_protocol_trace_init_to_forget_last in Hpre_tr_to as Hpre_tr. + apply finite_valid_trace_init_to_forget_last in Hpre_tr_to as Hpre_tr. destruct (equivocators_trace_project_output_reflecting_inv _ _ _ (proj1 Hpre_tr) _ Hno_equiv) as [final_descriptors_m [initial_descriptors_m [trXm [Hfinal_descriptors_m [Hproject_trXm Hex]]]]]. specialize (H (length tr')). @@ -1971,7 +1971,7 @@ Proof. by (apply not_equivocating_equivocator_descriptors_proper; assumption). specialize (H final_descriptors_m Hfinal_descriptors_m_proper). - apply pre_equivocators_protocol_trace_project + apply pre_equivocators_valid_trace_project with (final_descriptors0 := final_descriptors_m) in Hpre_tr_to as Hpr_tr' ; [|assumption..]. @@ -1980,8 +1980,8 @@ Proof. specialize (H _ _ Hproject_trXm'). simpl in *. rewrite Hproject_trXm in Hproject_trXm'. inversion Hproject_trXm'. subst trXm' initial_descriptors_m'. clear Hproject_trXm'. - apply option_protocol_message_Some. - apply (protocol_trace_output_is_protocol _ _ _ (proj1 H) _ Hex). + apply option_valid_message_Some. + apply (valid_trace_output_is_valid _ _ _ (proj1 H) _ Hex). Qed. Lemma SeededXE_incl_PreFreeE @@ -2006,8 +2006,8 @@ Proof. split; [split|]. - intros s tr sX trX Hpr_tr s_pre pre Hs_lst Hpre_tr. assert - (HPreFree_pre_tr : finite_protocol_trace_from SubPreFreeE s_pre (pre ++ tr)). - { revert Hpre_tr. apply VLSM_incl_finite_protocol_trace_from. + (HPreFree_pre_tr : finite_valid_trace_from SubPreFreeE s_pre (pre ++ tr)). + { revert Hpre_tr. apply VLSM_incl_finite_valid_trace_from. apply SeededXE_incl_PreFreeE. } clear Hpre_tr. revert s tr sX trX Hpr_tr s_pre pre Hs_lst HPreFree_pre_tr. @@ -2016,13 +2016,13 @@ Proof. destruct (destruct_equivocators_partial_trace_project sub_IM (finite_sub_index selection finite_index) Hpr_tr) as [Hnot_equiv [initial_descriptors [Htr_project Hs_project]]]. apply not_equivocating_equivocator_descriptors_proper in Hnot_equiv as Hproper. - destruct (seeded_equivocators_protocol_trace_project _ _ Htr _ Hproper) + destruct (seeded_equivocators_valid_trace_project _ _ Htr _ Hproper) as [_trX [_initial_descriptors [_ [_Htr_project [_ HtrX]]]]]. rewrite Htr_project in _Htr_project. inversion _Htr_project. subst. assumption. Qed. -End seeded_equivocators_protocol_trace_project. +End seeded_equivocators_valid_trace_project. End equivocators_composition_sub_projections. @@ -2065,8 +2065,8 @@ Proof. split; [split|]. - intros s tr sX trX Hpr_tr s_pre pre Hs_lst Hpre_tr. assert - (HPreFree_pre_tr : finite_protocol_trace_from PreFreeE s_pre (pre ++ tr)). - { revert Hpre_tr. apply VLSM_incl_finite_protocol_trace_from. + (HPreFree_pre_tr : finite_valid_trace_from PreFreeE s_pre (pre ++ tr)). + { revert Hpre_tr. apply VLSM_incl_finite_valid_trace_from. apply equivocators_no_equivocations_vlsm_incl_PreFree. } clear Hpre_tr. revert s tr sX trX Hpr_tr s_pre pre Hs_lst HPreFree_pre_tr. @@ -2078,7 +2078,7 @@ Proof. specialize (sub_composition_all_full_projection equivocator_IM (equivocators_no_equivocations_constraint IM Hbs)) as Hproj. - apply (VLSM_full_projection_finite_protocol_trace Hproj) in Htr. + apply (VLSM_full_projection_finite_valid_trace Hproj) in Htr. specialize (false_composite_no_equivocation_vlsm_with_pre_loaded (SubProjectionTraces.sub_IM equivocator_IM (finite.enum index)) @@ -2086,7 +2086,7 @@ Proof. (equivocator_Hbs sub_IM Hbs_sub)) as Heq. assert (Htr' : - finite_protocol_trace + finite_valid_trace (composite_vlsm (SubProjectionTraces.sub_IM equivocator_IM (finite.enum index)) (no_equivocations_additional_constraint @@ -2097,7 +2097,7 @@ Proof. (VLSM_full_projection_finite_trace_project Hproj tr) ). { revert Htr. - apply VLSM_incl_finite_protocol_trace. + apply VLSM_incl_finite_valid_trace. clear. apply constraint_subsumption_incl. apply preloaded_constraint_subsumption_stronger. @@ -2114,10 +2114,10 @@ Proof. (SubProjectionTraces.free_sub_free_index_obligation_1 eqv)). assumption. } - apply (VLSM_eq_finite_protocol_trace Heq) in Htr'. + apply (VLSM_eq_finite_valid_trace Heq) in Htr'. specialize - (seeded_equivocators_protocol_trace_project IM Hbs finite_index + (seeded_equivocators_valid_trace_project IM Hbs finite_index (finite.enum index) (fun m => False) _ _ Htr' @@ -2170,9 +2170,9 @@ Proof. (vlsm_is_pre_loaded_with_False (free_composite_vlsm (SubProjectionTraces.sub_IM IM (finite.enum index)))) as Heq. - apply (VLSM_eq_finite_protocol_trace Heq) in HtrX. + apply (VLSM_eq_finite_valid_trace Heq) in HtrX. specialize (sub_composition_all_full_projection_rev IM (free_constraint IM)) as Hproj. - assert (HtrX' : finite_protocol_trace (composite_vlsm (SubProjectionTraces.sub_IM IM (finite.enum index)) + assert (HtrX' : finite_valid_trace (composite_vlsm (SubProjectionTraces.sub_IM IM (finite.enum index)) (free_sub_free_constraint IM (free_constraint IM))) (Common.equivocators_state_project (SubProjectionTraces.sub_IM IM (finite.enum index)) @@ -2180,11 +2180,11 @@ Proof. (composite_state_sub_projection equivocator_IM (finite.enum index) s)) (finite_trace_sub_projection IM (finite.enum index) trX)). { revert HtrX. - apply VLSM_incl_finite_protocol_trace. + apply VLSM_incl_finite_valid_trace. apply constraint_subsumption_incl. intro; intros. destruct l. destruct som. exact I. } - apply (VLSM_full_projection_finite_protocol_trace Hproj) in HtrX'. + apply (VLSM_full_projection_finite_valid_trace Hproj) in HtrX'. replace (free_sub_free_state _ _) with (Common.equivocators_state_project IM initial_descriptors s) in HtrX' @@ -2204,12 +2204,12 @@ Proof. + elim H. apply finite.elem_of_enum. Qed. -Lemma equivocators_protocol_trace_from_project +Lemma equivocators_valid_trace_from_project (final_descriptors : equivocator_descriptors IM) (is final_state : vstate equivocators_no_equivocations_vlsm) (tr : list (composite_transition_item equivocator_IM)) (Hproper : not_equivocating_equivocator_descriptors IM final_descriptors final_state) - (Htr : finite_protocol_trace_from_to equivocators_no_equivocations_vlsm is final_state tr) + (Htr : finite_valid_trace_from_to equivocators_no_equivocations_vlsm is final_state tr) : exists isX final_stateX (trX : list (composite_transition_item IM)) @@ -2218,26 +2218,26 @@ Lemma equivocators_protocol_trace_from_project proper_equivocator_descriptors initial_descriptors is /\ equivocators_trace_project IM final_descriptors tr = Some (trX, initial_descriptors) /\ equivocators_state_project final_descriptors final_state = final_stateX /\ - finite_protocol_trace_from_to Free isX final_stateX trX. + finite_valid_trace_from_to Free isX final_stateX trX. Proof. - apply ptrace_get_last in Htr as Hfinal_state. apply ptrace_forget_last in Htr. + apply valid_trace_get_last in Htr as Hfinal_state. apply valid_trace_forget_last in Htr. subst final_state. specialize - (VLSM_partial_projection_finite_protocol_trace_from (equivocators_no_equivocations_vlsm_X_vlsm_partial_projection final_descriptors) + (VLSM_partial_projection_finite_valid_trace_from (equivocators_no_equivocations_vlsm_X_vlsm_partial_projection final_descriptors) is tr ) as Hsim. unfold equivocators_partial_trace_project in Hsim. rewrite decide_True in Hsim by assumption. - assert (HPreFree_tr : finite_protocol_trace_from PreFreeE is tr). - { revert Htr. apply VLSM_incl_finite_protocol_trace_from. apply equivocators_no_equivocations_vlsm_incl_PreFree. } + assert (HPreFree_tr : finite_valid_trace_from PreFreeE is tr). + { revert Htr. apply VLSM_incl_finite_valid_trace_from. apply equivocators_no_equivocations_vlsm_incl_PreFree. } apply not_equivocating_equivocator_descriptors_proper in Hproper. destruct - (preloaded_equivocators_protocol_trace_from_project _ + (preloaded_equivocators_valid_trace_from_project _ _ _ _ Hproper HPreFree_tr ) as [trX [initial_descriptors [Htr_project [Hinitial_desciptors Hfinal_project]]]]. eexists. eexists. eexists. eexists. split; [reflexivity|]. split; [apply Hinitial_desciptors|]. split; [apply Htr_project|]. split; [apply Hfinal_project|]. - apply ptrace_add_default_last. + apply valid_trace_add_default_last. apply Hsim; [|assumption]. rewrite Htr_project. reflexivity. Qed. @@ -2273,15 +2273,15 @@ Proof. constructor; [constructor; intros|]. - apply PreFreeE_Free_vlsm_projection_type. revert H. - apply VLSM_incl_finite_protocol_trace_from. + apply VLSM_incl_finite_valid_trace_from. apply equivocators_no_equivocations_vlsm_incl_PreFree. - intros. - assert (Hpre_tr : finite_protocol_trace PreFreeE sX trX). - { revert H. apply VLSM_incl_finite_protocol_trace. + assert (Hpre_tr : finite_valid_trace PreFreeE sX trX). + { revert H. apply VLSM_incl_finite_valid_trace. apply equivocators_no_equivocations_vlsm_incl_PreFree. } specialize - (VLSM_partial_projection_finite_protocol_trace (equivocators_no_equivocations_vlsm_X_vlsm_partial_projection (zero_descriptor IM)) + (VLSM_partial_projection_finite_valid_trace (equivocators_no_equivocations_vlsm_X_vlsm_partial_projection (zero_descriptor IM)) sX trX (equivocators_total_state_project IM sX) (equivocators_total_trace_project IM trX) ) as Hsim. spec Hsim. @@ -2304,7 +2304,7 @@ Proof. assumption. - intros. specialize - (VLSM_partial_projection_finite_protocol_trace (PreFreeE_PreFree_vlsm_partial_projection IM Hbs finite_index (zero_descriptor IM)) + (VLSM_partial_projection_finite_valid_trace (PreFreeE_PreFree_vlsm_partial_projection IM Hbs finite_index (zero_descriptor IM)) sX trX (equivocators_total_state_project IM sX) (equivocators_total_trace_project IM trX) ) as Hsim. spec Hsim. diff --git a/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/FullReplayTraces.v b/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/FullReplayTraces.v index 594164d2..3f6ee9a2 100644 --- a/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/FullReplayTraces.v +++ b/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/FullReplayTraces.v @@ -53,7 +53,7 @@ Context {message : Type} . Lemma SeededXE_Free_full_projection - (Hseed : forall m, seed m -> protocol_message_prop FreeE m) + (Hseed : forall m, seed m -> valid_message_prop FreeE m) : VLSM_full_projection SeededXE (composite_vlsm equivocator_IM (free_constraint equivocator_IM)) (lift_sub_label equivocator_IM equivocating) (lift_sub_state equivocator_IM equivocating). @@ -64,7 +64,7 @@ Proof. - apply lift_sub_transition. apply H. - apply (lift_sub_state_initial equivocator_IM). assumption. - destruct HmX as [Hinit|Hseeded]; [|apply Hseed; assumption]. - apply initial_message_is_protocol. + apply initial_message_is_valid. destruct Hinit as [i Him]. exists (proj1_sig i). assumption. Qed. @@ -514,8 +514,8 @@ Qed. Section pre_loaded_constrained_projection. (** -By replaying a [protocol_trace] on top of a [protocol_state] we obtain a -[protocol_trace]. We derive this as a more general [VLSM_weak_full_projection] +By replaying a [valid_trace] on top of a [valid_state] we obtain a +[valid_trace]. We derive this as a more general [VLSM_weak_full_projection] result for a class of VLSM parameterized by a constraint having "good" properties and pre-loaded with a seed, to allow deriving the [VLSM_weak_full_projection] result for both the free composition of equivocators @@ -527,25 +527,25 @@ Context (constraint : composite_label equivocator_IM -> composite_state equivocator_IM * option message -> Prop) (seed1 : message -> Prop) (SeededCE := pre_loaded_vlsm (composite_vlsm equivocator_IM constraint) seed1) - (Hconstraint_none : forall i ns s, i ∈ equivocating -> protocol_state_prop SeededCE s -> constraint (existT i (Spawn ns)) (s, None)) - (Hseed : forall m, seed m -> protocol_message_prop SeededCE m) + (Hconstraint_none : forall i ns s, i ∈ equivocating -> valid_state_prop SeededCE s -> constraint (existT i (Spawn ns)) (s, None)) + (Hseed : forall m, seed m -> valid_message_prop SeededCE m) (full_replay_state : composite_state equivocator_IM) - (Hfull_replay_state : protocol_state_prop SeededCE full_replay_state) - (Hsubsumption : forall l s om, protocol_valid SeededXE l (s, om) -> - protocol_state_prop SeededCE (lift_equivocators_sub_state_to full_replay_state s) -> + (Hfull_replay_state : valid_state_prop SeededCE full_replay_state) + (Hsubsumption : forall l s om, input_valid SeededXE l (s, om) -> + valid_state_prop SeededCE (lift_equivocators_sub_state_to full_replay_state s) -> constraint (lift_equivocators_sub_label_to full_replay_state l) (lift_equivocators_sub_state_to full_replay_state s, om)) . -Lemma replayed_initial_state_from_protocol +Lemma replayed_initial_state_from_valid (is : composite_state sub_equivocator_IM) (His : composite_initial_state_prop sub_equivocator_IM is) - : finite_protocol_trace_from SeededCE full_replay_state (replayed_initial_state_from full_replay_state is). + : finite_valid_trace_from SeededCE full_replay_state (replayed_initial_state_from full_replay_state is). Proof. cut (forall l, incl l (@sub_index_listing _ _ equivocating index_listing) -> - finite_protocol_plan_from SeededCE + finite_valid_plan_from SeededCE full_replay_state (map (initial_new_machine_transition_item is) l)). { intros Hplan. specialize (Hplan _ (incl_refl _)). - unfold finite_protocol_plan_from in Hplan. + unfold finite_valid_plan_from in Hplan. assumption. } intro l. @@ -553,17 +553,17 @@ Proof. - constructor. assumption. - spec IHl. { intros i Hi. apply H. apply in_app_iff. left. assumption. } rewrite map_app. - apply finite_protocol_plan_from_app_iff. + apply finite_valid_plan_from_app_iff. split; [assumption|]. - apply finite_ptrace_singleton. simpl. + apply finite_valid_trace_singleton. simpl. repeat split. - + revert IHl. apply apply_plan_last_protocol. - + apply option_protocol_message_None. + + revert IHl. apply apply_plan_last_valid. + + apply option_valid_message_None. + apply His. + apply Hconstraint_none. * clear -x. destruct_dec_sig x i Hi Hx. subst. assumption. - * apply finite_ptrace_last_pstate in IHl. + * apply finite_valid_trace_last_pstate in IHl. remember (finite_trace_last _ _) as lst. replace ((apply_plan _ _ _).2) with lst ; [assumption|]. @@ -572,10 +572,10 @@ Proof. Qed. Lemma lift_initial_message - : forall m, vinitial_message_prop SeededXE m -> protocol_message_prop SeededCE m. + : forall m, vinitial_message_prop SeededXE m -> valid_message_prop SeededCE m. Proof. intros m [Hinit | Hseeded]. - - apply initial_message_is_protocol. destruct Hinit as [[i Hi] Hinit]. + - apply initial_message_is_valid. destruct Hinit as [[i Hi] Hinit]. left. exists i. assumption. - apply Hseed. assumption. Qed. @@ -589,62 +589,62 @@ Proof. + apply Hsubsumption; assumption. - apply lift_equivocators_sub_transition; apply H. - rewrite <- replayed_initial_state_from_lift; [|assumption]. - apply finite_ptrace_last_pstate. - by apply replayed_initial_state_from_protocol. + apply finite_valid_trace_last_pstate. + by apply replayed_initial_state_from_valid. - by apply lift_initial_message. Qed. -Lemma sub_preloaded_replayed_trace_from_protocol_equivocating +Lemma sub_preloaded_replayed_trace_from_valid_equivocating (is : composite_state sub_equivocator_IM) (tr : list (composite_transition_item sub_equivocator_IM)) - (Htr : finite_protocol_trace SeededXE is tr) - : finite_protocol_trace_from SeededCE + (Htr : finite_valid_trace SeededXE is tr) + : finite_valid_trace_from SeededCE full_replay_state (replayed_trace_from full_replay_state is tr). Proof. destruct Htr as [Htr His]. - apply finite_protocol_trace_from_app_iff. - split; [apply replayed_initial_state_from_protocol; assumption|]. + apply finite_valid_trace_from_app_iff. + split; [apply replayed_initial_state_from_valid; assumption|]. rewrite replayed_initial_state_from_lift by assumption. revert Htr. - apply (VLSM_weak_full_projection_finite_protocol_trace_from lift_equivocators_sub_weak_projection). + apply (VLSM_weak_full_projection_finite_valid_trace_from lift_equivocators_sub_weak_projection). Qed. End pre_loaded_constrained_projection. Lemma SeededXE_PreFreeE_weak_full_projection (full_replay_state : composite_state equivocator_IM) - (Hfull_replay_state : protocol_state_prop PreFreeE full_replay_state) + (Hfull_replay_state : valid_state_prop PreFreeE full_replay_state) : VLSM_weak_full_projection SeededXE PreFreeE (lift_equivocators_sub_label_to full_replay_state) (lift_equivocators_sub_state_to full_replay_state). Proof. constructor. intros sX trX HtrX. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True FreeE) as Heq. - apply (VLSM_eq_finite_protocol_trace_from Heq). + apply (VLSM_eq_finite_valid_trace_from Heq). revert sX trX HtrX. apply lift_equivocators_sub_weak_projection. - intros; exact I. - - intros. apply initial_message_is_protocol. right. exact I. - - apply (VLSM_eq_protocol_state Heq) in Hfull_replay_state. assumption. + - intros. apply initial_message_is_valid. right. exact I. + - apply (VLSM_eq_valid_state Heq) in Hfull_replay_state. assumption. - intros. exact I. Qed. Lemma PreFreeSubE_PreFreeE_weak_full_projection (full_replay_state : composite_state equivocator_IM) - (Hfull_replay_state : protocol_state_prop PreFreeE full_replay_state) + (Hfull_replay_state : valid_state_prop PreFreeE full_replay_state) : VLSM_weak_full_projection PreFreeSubE PreFreeE (lift_equivocators_sub_label_to full_replay_state) (lift_equivocators_sub_state_to full_replay_state). Proof. apply basic_VLSM_weak_full_projection; intro; intros. - split; [apply lift_equivocators_sub_valid; apply Hv|exact I]. - apply lift_equivocators_sub_transition; apply H. - rewrite <- replayed_initial_state_from_lift; [|assumption]. - apply finite_ptrace_last_pstate. + apply finite_valid_trace_last_pstate. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True FreeE) as Heq. - apply (VLSM_eq_finite_protocol_trace_from Heq). - apply replayed_initial_state_from_protocol. + apply (VLSM_eq_finite_valid_trace_from Heq). + apply replayed_initial_state_from_valid. + intro; intros. exact I. - + apply (VLSM_eq_protocol_state Heq). assumption. + + apply (VLSM_eq_valid_state Heq). assumption. + assumption. - - apply any_message_is_protocol_in_preloaded. + - apply any_message_is_valid_in_preloaded. Qed. Section seeded_no_equiv. @@ -652,18 +652,18 @@ Section seeded_no_equiv. Context (SeededAllXE : VLSM message := composite_no_equivocation_vlsm_with_pre_loaded equivocator_IM (free_constraint _) (equivocator_Hbs IM Hbs) seed) (full_replay_state : composite_state equivocator_IM) - (Hfull_replay_state : protocol_state_prop SeededAllXE full_replay_state) + (Hfull_replay_state : valid_state_prop SeededAllXE full_replay_state) . Local Lemma SeededNoEquiv_subsumption - : forall l s om, protocol_valid SeededXE l (s, om) -> + : forall l s om, input_valid SeededXE l (s, om) -> no_equivocations_additional_constraint_with_pre_loaded equivocator_IM (free_constraint _) (equivocator_Hbs IM Hbs) seed (lift_equivocators_sub_label_to full_replay_state l) (lift_equivocators_sub_state_to full_replay_state s, om). Proof. split; [|exact I]. destruct om as [m|]; [|exact I]. destruct H as [Hs [_ [_ [Hc1 _]]]]. - apply (VLSM_incl_protocol_state (NoEquivocation.seeded_no_equivocation_incl_preloaded equivocator_IM (free_constraint _) (equivocator_Hbs IM Hbs) seed)) in Hfull_replay_state. - specialize (protocol_state_project_preloaded_to_preloaded _ equivocator_IM (free_constraint _) full_replay_state) + apply (VLSM_incl_valid_state (NoEquivocation.seeded_no_equivocation_incl_preloaded equivocator_IM (free_constraint _) (equivocator_Hbs IM Hbs) seed)) in Hfull_replay_state. + specialize (valid_state_project_preloaded_to_preloaded _ equivocator_IM (free_constraint _) full_replay_state) as Hfull_replay_state_pr. pose (no_equivocations_additional_constraint_with_pre_loaded sub_equivocator_IM (free_constraint sub_equivocator_IM) (equivocator_Hbs sub_IM Hbs_sub) seed) @@ -672,8 +672,8 @@ Proof. (pre_loaded_vlsm_incl_pre_loaded_with_all_messages (composite_vlsm sub_equivocator_IM constraint) seed) as Hincl. - apply (VLSM_incl_protocol_state Hincl) in Hs. - specialize (protocol_state_project_preloaded_to_preloaded _ sub_equivocator_IM constraint s) + apply (VLSM_incl_valid_state Hincl) in Hs. + specialize (valid_state_project_preloaded_to_preloaded _ sub_equivocator_IM constraint s) as Hs_pr. simpl in Hc1. destruct Hc1 as [Hsub_sent | Hseeded]. @@ -688,11 +688,11 @@ Proof. - right. assumption. Qed. -Local Lemma sent_are_protocol - : forall m, seed m -> protocol_message_prop SeededAllXE m. +Local Lemma sent_are_valid + : forall m, seed m -> valid_message_prop SeededAllXE m. Proof. intros m Hm. - apply initial_message_is_protocol. + apply initial_message_is_valid. right. assumption. Qed. @@ -705,28 +705,28 @@ Proof. constructor. apply lift_equivocators_sub_weak_projection; intros. - split; exact I. - - apply sent_are_protocol. assumption. + - apply sent_are_valid. assumption. - assumption. - apply SeededNoEquiv_subsumption; assumption. Qed. -Lemma sub_replayed_trace_from_protocol_equivocating +Lemma sub_replayed_trace_from_valid_equivocating (is : composite_state sub_equivocator_IM) (tr : list (composite_transition_item sub_equivocator_IM)) - (Htr : finite_protocol_trace SeededXE is tr) - : finite_protocol_trace_from SeededAllXE + (Htr : finite_valid_trace SeededXE is tr) + : finite_valid_trace_from SeededAllXE full_replay_state (replayed_trace_from full_replay_state is tr). Proof. unfold composite_no_equivocation_vlsm_with_pre_loaded in SeededAllXE. specialize - (sub_preloaded_replayed_trace_from_protocol_equivocating + (sub_preloaded_replayed_trace_from_valid_equivocating (no_equivocations_additional_constraint_with_pre_loaded equivocator_IM (free_constraint _) (equivocator_Hbs IM Hbs) seed) seed) - as Hprotocol. - spec Hprotocol; [split; exact I|]. - spec Hprotocol; [apply sent_are_protocol|]. - specialize (Hprotocol _ Hfull_replay_state). - apply Hprotocol; [|assumption]. + as Hvalid. + spec Hvalid; [split; exact I|]. + spec Hvalid; [apply sent_are_valid|]. + specialize (Hvalid _ Hfull_replay_state). + apply Hvalid; [|assumption]. intros. apply SeededNoEquiv_subsumption. assumption. diff --git a/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/SimulatingFree.v b/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/SimulatingFree.v index 519a799b..6a93018b 100644 --- a/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/SimulatingFree.v +++ b/theories/VLSM/Core/Equivocators/Composition/SimulatingFree/SimulatingFree.v @@ -62,14 +62,14 @@ for the given state. *) Definition zero_descriptor_constraint_lifting_prop : Prop := forall - es (Hes : protocol_state_prop CE es) + es (Hes : valid_state_prop CE es) om (Hom : sent_except_from (equivocator_IM IM) (equivocator_Hbs IM Hbs) (vinitial_message_prop CE) es om) eqv li, constraintE (existT eqv (ContinueWith 0 li)) (es, om). (** The [replayable_message_prop]erty plays an important role in designing a general, abstract, proof for trace simulation (Lemma -[generalized_equivocators_finite_protocol_trace_init_to_rev]), as it specifies +[generalized_equivocators_finite_valid_trace_init_to_rev]), as it specifies that given a message <> received in a regular node-composition state <> for which the constraint <> is satisfied, then any trace of the equivocators-composition (constrained by <>>) producing <> @@ -78,45 +78,45 @@ to <>, with no transitions being performed on the original copies. *) Definition replayable_message_prop : Prop := forall si s tr - (HtrX : finite_protocol_trace_init_to CX si s tr) + (HtrX : finite_valid_trace_init_to CX si s tr) eqv_state_s - (Hstate_protocol: protocol_state_prop CE eqv_state_s) + (Hstate_valid: valid_state_prop CE eqv_state_s) (Hstate_final_project : equivocators_total_state_project IM eqv_state_s = s) eqv_msg_is eqv_msg_s eqv_msg_tr - (Hmsg_trace : finite_protocol_trace_init_to CE eqv_msg_is eqv_msg_s eqv_msg_tr) + (Hmsg_trace : finite_valid_trace_init_to CE eqv_msg_is eqv_msg_s eqv_msg_tr) iom (Hfinal_msg : last_in_trace_except_from (vinitial_message_prop CE) eqv_msg_tr iom) l (HcX: constraintX l (s, iom)), exists eqv_msg_tr lst_msg_tr, - finite_protocol_trace_from_to CE eqv_state_s lst_msg_tr eqv_msg_tr /\ + finite_valid_trace_from_to CE eqv_state_s lst_msg_tr eqv_msg_tr /\ equivocators_total_trace_project IM eqv_msg_tr = [] /\ equivocators_total_state_project IM lst_msg_tr = s /\ sent_except_from (equivocator_IM IM) (equivocator_Hbs IM Hbs) (vinitial_message_prop CE) lst_msg_tr iom. (** The main result of this section, showing that every trace of the composition of regular nodes can be obtained as a [zero_descriptor] projection -of a protocol trace for the composition of equivocators. +of a valid trace for the composition of equivocators. The proof proceeds by replaying the transitions of the trace on the original copy of the machines, replaying traces corresponding to the messages between those transition through equivocation to guarantee the no-message-equivocation constraint. *) -Lemma generalized_equivocators_finite_protocol_trace_init_to_rev +Lemma generalized_equivocators_finite_valid_trace_init_to_rev (Hsubsumption : zero_descriptor_constraint_lifting_prop) (Hreplayable : replayable_message_prop) isX sX trX - (HtrX : finite_protocol_trace_init_to CX isX sX trX) + (HtrX : finite_valid_trace_init_to CX isX sX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = sX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to CE is s tr /\ + finite_valid_trace_init_to CE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. assert (HinclE : VLSM_incl CE PreFreeE) by apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages. - induction HtrX using finite_protocol_trace_init_to_rev_strong_ind. + induction HtrX using finite_valid_trace_init_to_rev_strong_ind. - specialize (lift_initial_to_equivocators_state IM Hbs _ His) as Hs. remember (lift_to_equivocators_state IM is) as s. cut (equivocators_state_project IM (zero_descriptor IM) s = is). @@ -127,7 +127,7 @@ Proof. split; [|reflexivity]. split; [|assumption]. constructor. - apply initial_is_protocol. assumption. + apply initial_state_is_valid. assumption. } apply functional_extensionality_dep_good. subst. @@ -135,10 +135,10 @@ Proof. - destruct IHHtrX1 as [eqv_state_is [Hstate_start_project [eqv_state_s [Hstate_final_project [eqv_state_tr [Hstate_project [Hstate_trace _ ]]]]]]]. destruct IHHtrX2 as [eqv_msg_is [Hmsg_start_project [eqv_msg_s [_ [eqv_msg_tr [Hmsg_project [Hmsg_trace Hfinal_msg ]]]]]]]. exists eqv_state_is. split; [assumption|]. - apply ptrace_last_pstate in Hstate_trace as Hstate_protocol. + apply valid_trace_last_pstate in Hstate_trace as Hstate_valid. destruct Ht as [[Hs [Hiom [Hv Hc]]] Ht]. specialize - (Hreplayable _ _ _ HtrX1 _ Hstate_protocol Hstate_final_project _ _ _ Hmsg_trace iom) + (Hreplayable _ _ _ HtrX1 _ Hstate_valid Hstate_final_project _ _ _ Hmsg_trace iom) as Hreplay. spec Hreplay. { clear -Heqiom Hfinal_msg. @@ -154,7 +154,7 @@ Proof. as [emsg_tr [es [Hmsg_trace_full_replay [Hemsg_tr_pr [Hes_pr Hbs_iom]]]]]. intros . specialize - (finite_protocol_trace_from_to_app CE _ _ _ _ _ (proj1 Hstate_trace) Hmsg_trace_full_replay) + (finite_valid_trace_from_to_app CE _ _ _ _ _ (proj1 Hstate_trace) Hmsg_trace_full_replay) as Happ. specialize (extend_right_finite_trace_from_to CE Happ) as Happ_extend. @@ -165,7 +165,7 @@ Proof. destruct (vtransition CE el (es, iom)) as (es', om') eqn:Hesom'. specialize (Happ_extend el iom es' om'). - apply ptrace_get_last in Happ as Heqes. + apply valid_trace_get_last in Happ as Heqes. assert (Hes_pr_i : forall i, equivocators_total_state_project IM es i = s i) by (rewrite <- Hes_pr; reflexivity). exists es'. @@ -182,8 +182,8 @@ Proof. match type of Happ_extend with | ?H -> _ => cut H end. - { intro Hpt. - spec Happ_extend Hpt. + { intro Hivt. + spec Happ_extend Hivt. match goal with |- ?H /\ _ => assert (Hproject : H) end. @@ -198,12 +198,12 @@ Proof. } split; [assumption|]. eexists; split; [|split;[split; [exact Happ_extend|]|]]. - - apply ptrace_forget_last in Happ_extend. - apply (VLSM_incl_finite_protocol_trace_from HinclE) in Happ_extend. + - apply valid_trace_forget_last in Happ_extend. + apply (VLSM_incl_finite_valid_trace_from HinclE) in Happ_extend. rewrite (equivocators_total_trace_project_app IM) by (eexists; exact Happ_extend). - apply ptrace_forget_last in Happ. - apply (VLSM_incl_finite_protocol_trace_from HinclE) in Happ. + apply valid_trace_forget_last in Happ. + apply (VLSM_incl_finite_valid_trace_from HinclE) in Happ. rewrite (equivocators_total_trace_project_app IM) by (eexists; exact Happ). rewrite Hemsg_tr_pr. @@ -223,13 +223,13 @@ Proof. - rewrite! finite_trace_last_output_is_last. reflexivity. } clear Happ_extend. - apply ptrace_last_pstate in Happ. + apply valid_trace_last_pstate in Happ. repeat split; [assumption|..|assumption]. - + destruct iom as [im|]; [|apply option_protocol_message_None]. + + destruct iom as [im|]; [|apply option_valid_message_None]. destruct Hbs_iom as [Hbs_iom | Hseeded]. - * apply (preloaded_composite_sent_protocol (equivocator_IM IM) finite_index (equivocator_Hbs IM Hbs) _ _ _ Happ _ Hbs_iom). - * apply initial_message_is_protocol. assumption. - + subst el. cbn. rewrite equivocator_state_project_zero. + * apply (preloaded_composite_sent_valid (equivocator_IM IM) finite_index (equivocator_Hbs IM Hbs) _ _ _ Happ _ Hbs_iom). + * apply initial_message_is_valid. assumption. + + subst el. cbn. rewrite equivocator_state_project_zero. rewrite Hes_pr_eqv. assumption. + apply Hsubsumption; assumption. Qed. @@ -284,24 +284,23 @@ Definition all_equivocating_replayed_trace_from (composite_state_sub_projection (equivocator_IM IM) index_listing is) (VLSM_full_projection_finite_trace_project Hproj tr). -Lemma replayed_trace_from_protocol_equivocating +Lemma replayed_trace_from_valid_equivocating (full_replay_state : composite_state (equivocator_IM IM)) - (Hfull_replay_state : protocol_state_prop SeededXE full_replay_state) + (Hfull_replay_state : valid_state_prop SeededXE full_replay_state) (is : composite_state (equivocator_IM IM)) (tr : list (composite_transition_item (equivocator_IM IM))) - (Htr : finite_protocol_trace SeededXE is tr) - : finite_protocol_trace_from SeededXE + (Htr : finite_valid_trace SeededXE is tr) + : finite_valid_trace_from SeededXE full_replay_state (all_equivocating_replayed_trace_from full_replay_state is tr). Proof. - specialize - (sub_replayed_trace_from_protocol_equivocating IM Hbs _ finite_index seed + apply + (sub_replayed_trace_from_valid_equivocating IM Hbs _ finite_index seed index_listing _ Hfull_replay_state - ) as Hprotocol. - apply Hprotocol. clear Hprotocol. + ). pose (Hproj := preloaded_sub_composition_all_full_projection (equivocator_IM IM) (no_equivocations_additional_constraint_with_pre_loaded (equivocator_IM IM) (free_constraint _) (equivocator_Hbs IM Hbs) seed) seed). - apply (VLSM_full_projection_finite_protocol_trace Hproj) in Htr. + apply (VLSM_full_projection_finite_valid_trace Hproj) in Htr. revert Htr. - apply VLSM_incl_finite_protocol_trace. + apply VLSM_incl_finite_valid_trace. apply basic_VLSM_incl_preloaded_with; intro; intros ; [assumption| |assumption..]. destruct H as [Hv Hc]. @@ -317,22 +316,22 @@ Proof. exists sub_i. subst. assumption. Qed. -(** Specializing the [generalized_equivocators_finite_protocol_trace_init_to_rev] +(** Specializing the [generalized_equivocators_finite_valid_trace_init_to_rev] result using the [free_constraint] for the composition of regular nodes and the no-message-equivocation constraint for the composition of equivocators. *) -Lemma seeded_equivocators_finite_protocol_trace_init_to_rev +Lemma seeded_equivocators_finite_valid_trace_init_to_rev (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) isX sX trX - (HtrX : finite_protocol_trace_init_to SeededFree isX sX trX) + (HtrX : finite_valid_trace_init_to SeededFree isX sX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = sX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to SeededXE is s tr /\ + finite_valid_trace_init_to SeededXE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. apply - (generalized_equivocators_finite_protocol_trace_init_to_rev + (generalized_equivocators_finite_valid_trace_init_to_rev IM Hbs finite_index) ; [..|assumption]. - intro; intros. split; [|exact I]. @@ -354,14 +353,14 @@ Proof. } specialize (NoEquivocation.seeded_no_equivocation_incl_preloaded (equivocator_IM IM) (free_constraint _) (equivocator_Hbs IM Hbs) seed) as HinclE. - apply ptrace_forget_last in Hmsg_trace. + apply valid_trace_forget_last in Hmsg_trace. specialize - (replayed_trace_from_protocol_equivocating - _ Hstate_protocol _ _ Hmsg_trace + (replayed_trace_from_valid_equivocating + _ Hstate_valid _ _ Hmsg_trace ) as Hmsg_trace_full_replay. remember (all_equivocating_replayed_trace_from _ _ _ ) as emsg_tr. - apply ptrace_add_default_last in Hmsg_trace_full_replay. + apply valid_trace_add_default_last in Hmsg_trace_full_replay. eexists _,_; split; [exact Hmsg_trace_full_replay|]. subst. unfold all_equivocating_replayed_trace_from. @@ -376,10 +375,10 @@ Proof. repeat split. simpl. destruct Hfinal_msg as [Hfinal_msg | Hinitial]; [|right; assumption]. left. - apply ptrace_first_pstate in Hmsg_trace_full_replay as Hfst. - apply protocol_state_has_trace in Hfst as [is_s [tr_s [Htr_s His_s]]]. - specialize (finite_protocol_trace_from_to_app SeededXE _ _ _ _ _ Htr_s Hmsg_trace_full_replay) as Happ. - apply (VLSM_incl_finite_protocol_trace_from_to HinclE) in Happ. + apply valid_trace_first_pstate in Hmsg_trace_full_replay as Hfst. + apply valid_state_has_trace in Hfst as [is_s [tr_s [Htr_s His_s]]]. + specialize (finite_valid_trace_from_to_app SeededXE _ _ _ _ _ Htr_s Hmsg_trace_full_replay) as Happ. + apply (VLSM_incl_finite_valid_trace_from_to HinclE) in Happ. apply (has_been_sent_examine_one_trace FreeE_Hbs _ _ _ (conj Happ His_s)). @@ -408,23 +407,23 @@ Context {message : Type} (XE : VLSM message := equivocators_no_equivocations_vlsm IM Hbs) . -(** Further specializing [seeded_equivocators_finite_protocol_trace_init_to_rev] +(** Further specializing [seeded_equivocators_finite_valid_trace_init_to_rev] to remove the pre-loading operation. *) -Lemma equivocators_finite_protocol_trace_init_to_rev +Lemma equivocators_finite_valid_trace_init_to_rev (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) isX sX trX - (HtrX : finite_protocol_trace_init_to Free isX sX trX) + (HtrX : finite_valid_trace_init_to Free isX sX trX) : exists is, equivocators_total_state_project IM is = isX /\ exists s, equivocators_total_state_project IM s = sX /\ exists tr, equivocators_total_trace_project IM tr = trX /\ - finite_protocol_trace_init_to XE is s tr /\ + finite_valid_trace_init_to XE is s tr /\ finite_trace_last_output trX = finite_trace_last_output tr. Proof. specialize (vlsm_is_pre_loaded_with_False Free) as Heq. - apply (VLSM_eq_finite_protocol_trace_init_to Heq) in HtrX. + apply (VLSM_eq_finite_valid_trace_init_to Heq) in HtrX. specialize - (seeded_equivocators_finite_protocol_trace_init_to_rev + (seeded_equivocators_finite_valid_trace_init_to_rev IM Hbs finite_index (fun m => False) no_initial_messages_in_IM _ _ _ HtrX) as [is [His [s [Hs [tr [Htr_pr [Htr Houtput]]]]]]]. @@ -438,9 +437,9 @@ Proof. clear Heq. specialize (vlsm_is_pre_loaded_with_False (composite_vlsm (equivocator_IM IM) constraint)) as Heq. - apply (VLSM_eq_finite_protocol_trace_init_to Heq) in Htr. + apply (VLSM_eq_finite_valid_trace_init_to Heq) in Htr. revert Htr. - apply VLSM_incl_finite_protocol_trace_init_to. + apply VLSM_incl_finite_valid_trace_init_to. apply constraint_subsumption_incl. apply preloaded_constraint_subsumption_stronger. apply strong_constraint_subsumption_strongest. diff --git a/theories/VLSM/Core/Equivocators/EquivocatorReplay.v b/theories/VLSM/Core/Equivocators/EquivocatorReplay.v index 66e46750..14bf5e24 100644 --- a/theories/VLSM/Core/Equivocators/EquivocatorReplay.v +++ b/theories/VLSM/Core/Equivocators/EquivocatorReplay.v @@ -60,7 +60,7 @@ Qed. Lemma equivocator_state_append_initial_state_in_futures (seed : message -> Prop) (base_s : equivocator_state X) - (Hbase_s : protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) + (Hbase_s : valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) s : vinitial_state_prop (equivocator_vlsm X) s -> in_futures (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s (equivocator_state_append base_s s). @@ -71,10 +71,10 @@ Proof. None (equivocator_state_append base_s s) None)]. - apply (finite_ptrace_from_to_singleton (pre_loaded_vlsm (equivocator_vlsm X) seed)). + apply (finite_valid_trace_from_to_singleton (pre_loaded_vlsm (equivocator_vlsm X) seed)). repeat split. - assumption. - - apply option_protocol_message_None. + - apply option_valid_message_None. - apply H. - cbn. f_equal. symmetry. apply equivocator_state_append_singleton_is_extend. apply H. Qed. @@ -82,20 +82,20 @@ Qed. Lemma equivocator_state_append_transition_initial_state (seed : message -> Prop) (base_s : equivocator_state X) - (Hbase_s : protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) + (Hbase_s : valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) s : vinitial_state_prop (equivocator_vlsm X) s -> - protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) + valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) (equivocator_state_append base_s s). Proof. - intros Hs. apply in_futures_protocol_snd with base_s. + intros Hs. apply in_futures_valid_snd with base_s. apply equivocator_state_append_initial_state_in_futures; assumption. Qed. Lemma equivocator_state_append_preloaded_with_weak_projection (seed : message -> Prop) (base_s : equivocator_state X) - (Hbase_s : protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) + (Hbase_s : valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) : VLSM_weak_full_projection (pre_loaded_vlsm (equivocator_vlsm X) seed) (pre_loaded_vlsm (equivocator_vlsm X) seed) (equivocator_state_append_label base_s) (equivocator_state_append base_s). Proof. @@ -103,50 +103,50 @@ Proof. - apply equivocator_state_append_valid. apply Hv. - apply equivocator_state_append_transition; apply H. - apply equivocator_state_append_transition_initial_state; assumption. - - apply initial_message_is_protocol; assumption. + - apply initial_message_is_valid; assumption. Qed. Lemma equivocator_state_append_preloaded_weak_projection (base_s : equivocator_state X) - (Hbase_s : protocol_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) + (Hbase_s : valid_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) : VLSM_weak_full_projection (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) (equivocator_state_append_label base_s) (equivocator_state_append base_s). Proof. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (equivocator_vlsm X)) as Heq. constructor. intros sX trX HtrX. - apply (VLSM_eq_finite_protocol_trace_from Heq) in HtrX. - apply (VLSM_eq_finite_protocol_trace_from Heq). + apply (VLSM_eq_finite_valid_trace_from Heq) in HtrX. + apply (VLSM_eq_finite_valid_trace_from Heq). revert sX trX HtrX. apply equivocator_state_append_preloaded_with_weak_projection. - apply (VLSM_eq_protocol_state Heq). assumption. + apply (VLSM_eq_valid_state Heq). assumption. Qed. Lemma equivocator_state_append_weak_projection (base_s : equivocator_state X) - (Hbase_s : protocol_state_prop (equivocator_vlsm X) base_s) + (Hbase_s : valid_state_prop (equivocator_vlsm X) base_s) : VLSM_weak_full_projection (equivocator_vlsm X) (equivocator_vlsm X) (equivocator_state_append_label base_s) (equivocator_state_append base_s). Proof. specialize (vlsm_is_pre_loaded_with_False (equivocator_vlsm X)) as Heq. constructor. intros sX trX HtrX. - apply (VLSM_eq_finite_protocol_trace_from Heq) in HtrX. - apply (VLSM_eq_finite_protocol_trace_from Heq). + apply (VLSM_eq_finite_valid_trace_from Heq) in HtrX. + apply (VLSM_eq_finite_valid_trace_from Heq). revert sX trX HtrX. apply equivocator_state_append_preloaded_with_weak_projection. - apply (VLSM_eq_protocol_state Heq). assumption. + apply (VLSM_eq_valid_state Heq). assumption. Qed. Lemma equivocator_state_append_in_futures (seed : message -> Prop) - base_s (Hbase_s : protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) - s (Hs : protocol_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) s) + base_s (Hbase_s : valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s) + s (Hs : valid_state_prop (pre_loaded_vlsm (equivocator_vlsm X) seed) s) : in_futures (pre_loaded_vlsm (equivocator_vlsm X) seed) base_s (equivocator_state_append base_s s). Proof. - apply protocol_state_has_trace in Hs as [is [tr [Htr His]]]. + apply valid_state_has_trace in Hs as [is [tr [Htr His]]]. specialize (equivocator_state_append_preloaded_with_weak_projection seed _ Hbase_s) as Hproj. - apply (VLSM_weak_full_projection_finite_protocol_trace_from_to Hproj) in Htr. + apply (VLSM_weak_full_projection_finite_valid_trace_from_to Hproj) in Htr. apply in_futures_trans with (equivocator_state_append base_s is). - apply equivocator_state_append_initial_state_in_futures; assumption. - eexists; exact Htr. @@ -154,14 +154,14 @@ Qed. Lemma equivocator_state_append_sent_left {Hbs : HasBeenSentCapability X} - base_s (Hbase_s : protocol_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) - s (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) s) + base_s (Hbase_s : valid_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) + s (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) s) : forall m, equivocator_has_been_sent X base_s m -> equivocator_has_been_sent X (equivocator_state_append base_s s) m. Proof. specialize (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (equivocator_vlsm X)) as Heq. - apply (VLSM_eq_protocol_state Heq) in Hbase_s. - apply (VLSM_eq_protocol_state Heq) in Hs. + apply (VLSM_eq_valid_state Heq) in Hbase_s. + apply (VLSM_eq_valid_state Heq) in Hs. apply (equivocator_state_append_in_futures _ _ Hbase_s) in Hs. apply (VLSM_eq_in_futures Heq) in Hs. apply (in_futures_preserving_oracle_from_stepwise _ (equivocator_vlsm X) (field_selector output) (equivocator_has_been_sent X)) @@ -172,8 +172,8 @@ Qed. Lemma equivocator_state_append_sent_right {Hbs : HasBeenSentCapability X} (HbsE := equivocator_HasBeenSentCapability X) - base_s (Hbase_s : protocol_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) - s (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) s) + base_s (Hbase_s : valid_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) base_s) + s (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm (equivocator_vlsm X)) s) : forall m, equivocator_has_been_sent X s m -> equivocator_has_been_sent X (equivocator_state_append base_s s) m. Proof. diff --git a/theories/VLSM/Core/Equivocators/MessageProperties.v b/theories/VLSM/Core/Equivocators/MessageProperties.v index 751103fd..d3006c64 100644 --- a/theories/VLSM/Core/Equivocators/MessageProperties.v +++ b/theories/VLSM/Core/Equivocators/MessageProperties.v @@ -57,10 +57,10 @@ Proof. + specialize (IHtr i' eq_refl Hjbs). right. assumption. Qed. -Lemma preloaded_equivocator_vlsm_trace_project_protocol_item_new_machine +Lemma preloaded_equivocator_vlsm_trace_project_valid_item_new_machine (bs : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs btr) + (Hbtr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs btr) (bitem : vtransition_item equivocator_vlsm) (Hitem : bitem ∈ btr) (sn : state) @@ -73,14 +73,14 @@ Proof. apply elem_of_list_split in Hitem. destruct Hitem as [bprefix [bsuffix Heq]]. subst btr. - apply (finite_protocol_trace_from_app_iff (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hbtr. + apply (finite_valid_trace_from_app_iff (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hbtr. destruct Hbtr as [Hbprefix Hbsuffix]. remember (finite_trace_last _ _) as lst. inversion Hbsuffix. subst s' tl bitem. destruct Ht as [[_ [_ Hv]] Ht]. simpl. specialize - (equivocator_protocol_transition_item_project_inv5_new_machine + (equivocator_valid_transition_project_inv5_new_machine _ l s lst iom oom Ht) as Hpitem. unfold VLSM.l in *. @@ -94,14 +94,14 @@ Proof. Qed. (** -For any [transition_item] in a protocol trace segment of an +For any [transition_item] in a valid trace segment of an [equivocator_vlsm] there exists a projection of that trace containing the projection of the item. *) -Lemma preloaded_equivocator_vlsm_trace_project_protocol_item +Lemma preloaded_equivocator_vlsm_trace_project_valid_item (bs bf : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs bf btr) + (Hbtr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs bf btr) (bitem : vtransition_item equivocator_vlsm) (Hitem : bitem ∈ btr) (idl : nat) @@ -116,25 +116,25 @@ Lemma preloaded_equivocator_vlsm_trace_project_protocol_item existing_descriptor X dfinal (finite_trace_last bs btr) /\ equivocator_vlsm_trace_project _ btr dfinal = Some (tr, dfirst). Proof. - specialize (preloaded_equivocator_vlsm_protocol_trace_project_inv2 X bs bf btr) as Hinv2. + specialize (preloaded_equivocator_vlsm_valid_trace_project_inv2 X bs bf btr) as Hinv2. spec Hinv2. { intro contra. subst. inversion Hitem. } spec Hinv2 Hbtr. apply elem_of_list_split in Hitem. destruct Hitem as [bprefix [bsuffix Heq]]. subst btr. - apply (finite_protocol_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hbtr. + apply (finite_valid_trace_from_to_app_split (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hbtr. destruct Hbtr as [Hbprefix Hbsuffix]. remember (finite_trace_last _ _) as lst. inversion Hbsuffix; subst s' f bitem tl. destruct Ht as [[_ [_ Hv]] Ht]. specialize - (equivocator_protocol_transition_item_project_inv5 _ l s lst iom oom Hv Ht) as Hpitem. + (equivocator_valid_transition_project_inv5 _ l s lst iom oom Hv Ht) as Hpitem. unfold VLSM.l in *. destruct (Hpitem _ Hlbitem) as [i [si [Hi [itemx Hitemx]]]]. - apply ptrace_forget_last in Htl. + apply valid_trace_forget_last in Htl. destruct - (preloaded_equivocator_vlsm_trace_project_protocol_inv _ _ _ Htl _ _ Hi) + (preloaded_equivocator_vlsm_trace_project_valid_inv _ _ _ Htl _ _ Hi) as [suffix Hsuffix]. exists itemx. split; [eexists; apply Hitemx|]. remember (Existing i) as dsuffix. @@ -147,14 +147,14 @@ Proof. spec Hsuffix' Hsuffix. subst bitem. destruct - (equivocator_protocol_transition_item_project_inv2 _ l lst s iom oom Hv Ht _ _ _ Hitemx) + (equivocator_valid_transition_project_inv2 _ l lst s iom oom Hv Ht _ _ _ Hitemx) as [_i [_Hdsuffix [s_i [_Hi Hd']]]]. inversion _Hdsuffix. subst _i. clear _Hdsuffix. destruct Hd' as [_i' [_Heq [lst_i' [id [Hex [_Hitemx [Hvs' Hts']]]]]]]. inversion _Heq. subst _i'. clear _Heq. subst lst. destruct - (preloaded_equivocator_vlsm_trace_project_protocol _ _ _ _ Hbprefix idl _ id) + (preloaded_equivocator_vlsm_trace_project_valid _ _ _ _ Hbprefix idl _ id) as [prefix [dfirst [Hprefix _]]]. specialize (equivocator_vlsm_trace_project_app_inv _ _ _ _ _ _ _ _ Hprefix Hsuffix') @@ -179,13 +179,13 @@ Proof. Qed. (** -If an [equivocator_vlsm]'s protocol trace segment [output]s a message, then +If an [equivocator_vlsm]'s valid trace segment [output]s a message, then one of its projections must do so too. *) Lemma equivocator_vlsm_trace_project_output_reflecting_inv (is: vstate equivocator_vlsm) (tr: list (vtransition_item equivocator_vlsm)) - (Htr: finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) is tr) + (Htr: finite_valid_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) is tr) (m : message) (Hbbs : Exists (field_selector output m) tr) : exists @@ -203,14 +203,14 @@ Proof. - destruct item. destruct l; inversion Hsndl. subst. simpl in *. specialize - (preloaded_equivocator_vlsm_trace_project_protocol_item_new_machine + (preloaded_equivocator_vlsm_trace_project_valid_item_new_machine _ _ Htr _ Hin _ eq_refl) as Hitem. simpl in Hitem. destruct Hitem as [_ [Hcontra _]]. congruence. - - apply ptrace_add_default_last in Htr. + - apply valid_trace_add_default_last in Htr. specialize - (preloaded_equivocator_vlsm_trace_project_protocol_item + (preloaded_equivocator_vlsm_trace_project_valid_item _ _ _ Htr _ Hin _ Hsndl) as [itemx [[d Hitemx] [trx [Hinx [ifinal [ifirst [Hifirst [Hifinal Htrx]]]]]]]]. exists ifinal. exists ifirst. split; [assumption|]. @@ -315,8 +315,8 @@ Proof. (oracle_step_update l sidesc im sidesc' om'). spec oracle_step_update. { repeat split - ; [..| eexists _; apply (pre_loaded_with_all_messages_message_protocol_prop X)|assumption|assumption]. - apply (preloaded_equivocator_state_project_protocol_state X _ Hs _ _ Hidesc). + ; [..| eexists _; apply (pre_loaded_with_all_messages_message_valid_initial_state_message X)|assumption|assumption]. + apply (preloaded_equivocator_state_project_valid_state X _ Hs _ _ Hidesc). } specialize (existing_false_label_equivocator_state_project_not_same X Ht _ Hidesc) as Hnot_same. @@ -372,8 +372,8 @@ Proof. (oracle_step_update l sidesc im sidesc' om'). spec oracle_step_update. { repeat split - ; [..| eexists _; apply (pre_loaded_with_all_messages_message_protocol_prop X)|assumption|assumption]. - apply (preloaded_equivocator_state_project_protocol_state X _ Hs _ _ Hidesc). + ; [..| eexists _; apply (pre_loaded_with_all_messages_message_valid_initial_state_message X)|assumption|assumption]. + apply (preloaded_equivocator_state_project_valid_state X _ Hs _ _ Hidesc). } specialize (existing_true_label_equivocator_state_project_not_last X Ht _ Hidesc) as Hnot_last. @@ -541,12 +541,12 @@ Definition equivocator_sent_messages_fn *) Lemma equivocator_sent_messages_full (s : vstate equivocator_vlsm) - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm equivocator_vlsm) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm equivocator_vlsm) s) (m : message) : m ∈ (equivocator_sent_messages_fn s) <-> exists (sm : sent_messages equivocator_vlsm s), proj1_sig sm = m. Proof. - specialize (preloaded_equivocator_state_project_protocol_state _ _ Hs) as HpsX. + specialize (preloaded_equivocator_state_project_valid_state _ _ Hs) as HpsX. split. - intro Hin. apply set_union_in_iterated in Hin. apply Exists_exists in Hin as [msgsi [Hmsgsi Hin]]. @@ -558,7 +558,7 @@ Proof. destruct Hin as [[m' Hm] Heq]. simpl in Heq. subst m'. apply (sent_messages_consistency X) in Hm; [|assumption]. destruct Hs as [om Hs]. - apply (protocol_is_trace (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hs. + apply (valid_state_message_has_trace (pre_loaded_with_all_messages_vlsm equivocator_vlsm)) in Hs. destruct Hs as [[Hs _] | [is [tr [Htr _]]]]. + specialize (Hm si []). spec Hm. @@ -567,13 +567,13 @@ Proof. - revert Hsi Hs. apply (equivocator_vlsm_initial_state_preservation_rev X). } inversion Hm. - + apply ptrace_get_last in Htr as Hlst. + + apply valid_trace_get_last in Htr as Hlst. assert (Hbm : selected_message_exists_in_some_preloaded_traces equivocator_vlsm (field_selector output) s m) ; [|exists (exist _ m Hbm); reflexivity]. exists is. exists tr. exists Htr. subst s. destruct - (preloaded_equivocator_vlsm_trace_project_protocol _ _ _ _ (proj1 Htr) _ _ Hsi) + (preloaded_equivocator_vlsm_trace_project_valid _ _ _ _ (proj1 Htr) _ _ Hsi) as [trX [di [Hproject Hdi]]]. destruct di as [sn | id]. * apply equivocator_vlsm_trace_project_output_reflecting with trX (Existing i) (NewMachine sn) @@ -587,7 +587,7 @@ Proof. - intros [[m' Hm] Heq]. simpl in Heq. subst m'. destruct Hm as [is [tr [Htr Hexists]]]. destruct - (equivocator_vlsm_trace_project_output_reflecting_inv _ _ (proj1 (ptrace_forget_last Htr)) _ Hexists) + (equivocator_vlsm_trace_project_output_reflecting_inv _ _ (proj1 (valid_trace_forget_last Htr)) _ Hexists) as [ifinal [istart [_ [_ [trX [Hproject HexistsX]]]]]]. assert (Hntr : tr <> []) by (intro contra; subst; inversion Hexists). destruct ifinal as [sfinal | i] @@ -596,7 +596,7 @@ Proof. ; inversion Hproject; subst; inversion HexistsX |]. specialize - (preloaded_equivocator_vlsm_protocol_trace_project_inv2 _ _ _ _ Hntr (proj1 Htr) _ _ _ Hproject) + (preloaded_equivocator_vlsm_valid_trace_project_inv2 _ _ _ _ Hntr (proj1 Htr) _ _ _ Hproject) as [si [Hi Histart]]. apply set_union_in_iterated. apply Exists_exists. exists (sent_messages_fn X si). diff --git a/theories/VLSM/Core/Equivocators/Projections.v b/theories/VLSM/Core/Equivocators/Projections.v index 7cf4abe0..7d7fa4e1 100644 --- a/theories/VLSM/Core/Equivocators/Projections.v +++ b/theories/VLSM/Core/Equivocators/Projections.v @@ -13,9 +13,9 @@ trace in the original vlsm leading to the <>, the <>th internal state in <>, by extracting a path leading to si. This section is devoting to formalizing this projects studying its -properties. In particular, we show that given a [protocol_trace] for +properties. In particular, we show that given a [valid_trace] for the [equivocator_vlsm], we can always extract such a trace for any valid -index, and, furthermore, that the trace extracted is protocol for the +index, and, furthermore, that the trace extracted is valid for the original machine. *) @@ -248,7 +248,7 @@ applying an equivocator projection (trace, transition_item) function in terms of the input descriptor and the resulting state. It is assumed that the original_descriptor is a proper descriptor -w.r.t. the final state of the trace/transition on which +w.r.t. the final state of the trace/transition on which [equivocator_vlsm_transition_item_project] or [equivocator_vlsm_trace_project] was applied. In particular this makes s_descriptor a proper descriptor for the state s (see the lemmas above and below). @@ -615,7 +615,7 @@ Qed. (** Next we prove some inversion properties for [equivocator_vlsm_transition_item_project]. *) -Lemma equivocator_protocol_transition_item_project_inv2 +Lemma equivocator_valid_transition_project_inv2 (l : vlabel equivocator_vlsm) (s' s: vstate equivocator_vlsm) (iom oom : option message) @@ -658,7 +658,7 @@ Proof. lia. Qed. -Lemma equivocator_protocol_transition_item_project_inv3 +Lemma equivocator_valid_transition_project_inv3 (l : vlabel equivocator_vlsm) (s s' : vstate equivocator_vlsm) (iom oom : option message) @@ -726,7 +726,7 @@ Proof. simpl in Hn. subst. reflexivity. Qed. -Lemma equivocator_protocol_transition_item_project_inv4 +Lemma equivocator_valid_transition_project_inv4 (l : vlabel equivocator_vlsm) (s s' : vstate equivocator_vlsm) (iom oom : option message) @@ -769,7 +769,7 @@ Proof. specialize (equivocator_state_last_n X s). lia. Qed. -Lemma equivocator_protocol_transition_item_project_inv5_new_machine +Lemma equivocator_valid_transition_project_inv5_new_machine (l : vlabel equivocator_vlsm) (s s' : vstate equivocator_vlsm) (iom oom : option message) @@ -791,7 +791,7 @@ Proof. rewrite decide_True; reflexivity. Qed. -Lemma equivocator_protocol_transition_item_project_inv5 +Lemma equivocator_valid_transition_project_inv5 (l : vlabel equivocator_vlsm) (s s' : vstate equivocator_vlsm) (iom oom : option message) @@ -835,14 +835,14 @@ Proof. Qed. (** -The projection of a segment of an [equivocator_vlsm] protocol trace -is defined and a protocol trace segment in the original vlsm. +The projection of a segment of an [equivocator_vlsm] valid trace +is defined and a valid trace segment in the original vlsm. *) -Lemma preloaded_with_equivocator_vlsm_trace_project_protocol +Lemma preloaded_with_equivocator_vlsm_trace_project_valid (seed : message -> Prop) (bs be : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from_to (pre_loaded_vlsm equivocator_vlsm seed) bs be btr) + (Hbtr : finite_valid_trace_from_to (pre_loaded_vlsm equivocator_vlsm seed) bs be btr) (j : nat) ej (Hj : equivocator_state_project be j = Some ej) @@ -853,15 +853,15 @@ Lemma preloaded_with_equivocator_vlsm_trace_project_protocol match di with | NewMachine sn => vinitial_state_prop X sn - /\ finite_protocol_trace_from_to (pre_loaded_vlsm X seed) sn ej tr + /\ finite_valid_trace_from_to (pre_loaded_vlsm X seed) sn ej tr | Existing i => exists s, equivocator_state_project bs i = Some s /\ - finite_protocol_trace_from_to (pre_loaded_vlsm X seed) s ej tr + finite_valid_trace_from_to (pre_loaded_vlsm X seed) s ej tr end. Proof. induction Hbtr; intros. - exists []. eexists; split; [reflexivity|]. eexists; split; [exact Hj|]. - constructor. revert Hj. apply preloaded_with_equivocator_state_project_protocol_state. assumption. + constructor. revert Hj. apply preloaded_with_equivocator_state_project_valid_state. assumption. - remember {| l := l; input := iom; |} as item. destruct Ht as [[Hs' [Hiom Hv]] Ht]. specialize (IHHbtr Hj) as [tlX [di' [Htl_pr Hdi]]]. @@ -895,26 +895,26 @@ Proof. split; [apply Hproper'|]. destruct itemX. simpl in *. rewrite Heqsi in Hitem_pr. subst. - apply (finite_ptrace_from_to_extend (pre_loaded_vlsm X seed)); [assumption|]. + apply (finite_valid_trace_from_to_extend (pre_loaded_vlsm X seed)); [assumption|]. repeat split; [..|assumption|assumption]. - * apply initial_is_protocol. assumption. - * apply preloaded_with_equivocator_state_project_protocol_message. assumption. + * apply initial_state_is_valid. assumption. + * apply preloaded_with_equivocator_state_project_valid_message. assumption. + destruct Hproper' as [s'i' Heqs'i']; rewrite Heqs'i' in *. simpl in Hchar2. specialize (Hchar2 _ eq_refl) as [HvX HtX]. eexists _; split; [reflexivity|]. destruct itemX. simpl in *. subst. - apply (finite_ptrace_from_to_extend (pre_loaded_vlsm X seed)); [assumption|]. + apply (finite_valid_trace_from_to_extend (pre_loaded_vlsm X seed)); [assumption|]. repeat split; [..|assumption|assumption]. - * revert Heqs'i'. apply preloaded_with_equivocator_state_project_protocol_state. assumption. - * apply preloaded_with_equivocator_state_project_protocol_message. assumption. + * revert Heqs'i'. apply preloaded_with_equivocator_state_project_valid_state. assumption. + * apply preloaded_with_equivocator_state_project_valid_message. assumption. Qed. -Lemma equivocator_vlsm_trace_project_protocol +Lemma equivocator_vlsm_trace_project_valid (bs be : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from_to equivocator_vlsm bs be btr) + (Hbtr : finite_valid_trace_from_to equivocator_vlsm bs be btr) (j : nat) ej (Hj : equivocator_state_project be j = Some ej) @@ -925,40 +925,40 @@ Lemma equivocator_vlsm_trace_project_protocol match di with | NewMachine sn => vinitial_state_prop X sn - /\ finite_protocol_trace_from_to X sn ej tr + /\ finite_valid_trace_from_to X sn ej tr | Existing i => exists s, equivocator_state_project bs i = Some s /\ - finite_protocol_trace_from_to X s ej tr + finite_valid_trace_from_to X s ej tr end. Proof. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj1 (vlsm_is_pre_loaded_with_False equivocator_vlsm))) in Hbtr. - specialize (preloaded_with_equivocator_vlsm_trace_project_protocol _ _ _ _ Hbtr _ _ Hj) + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj1 (vlsm_is_pre_loaded_with_False equivocator_vlsm))) in Hbtr. + specialize (preloaded_with_equivocator_vlsm_trace_project_valid _ _ _ _ Hbtr _ _ Hj) as [tr [di [Hbtr_pr Hdi]]]. eexists _,_; split; [exact Hbtr_pr|]. destruct di as [sn|i]. - destruct Hdi as [Hsn Htr]. split; [assumption|]. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Htr. + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Htr. clear -Htr. destruct X as (T, (S, M)). assumption. - destruct Hdi as [s [Hpr_bs_i Htr]]. eexists; split; [exact Hpr_bs_i|]. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Htr. + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj2 (vlsm_is_pre_loaded_with_False X))) in Htr. clear -Htr. destruct X as (T, (S, M)). assumption. Qed. (** -The projection of a segment of a protocol trace from the [pre_loaded_with_all_messages_vlsm] -corresponding to the [equivocator_vlsm] is defined and it is a protocol +The projection of a segment of a valid trace from the [pre_loaded_with_all_messages_vlsm] +corresponding to the [equivocator_vlsm] is defined and it is a valid trace segment in the [pre_loaded_with_all_messages_vlsm] corresponding to the original vlsm. *) -Lemma preloaded_equivocator_vlsm_trace_project_protocol +Lemma preloaded_equivocator_vlsm_trace_project_valid (bs be : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs be btr) + (Hbtr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs be btr) (j : nat) ej (Hj : equivocator_state_project be j = Some ej) @@ -969,26 +969,26 @@ Lemma preloaded_equivocator_vlsm_trace_project_protocol match di with | NewMachine sn => vinitial_state_prop X sn - /\ finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm X) sn ej tr + /\ finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm X) sn ej tr | Existing i => exists s, equivocator_state_project bs i = Some s /\ - finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm X) s ej tr + finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm X) s ej tr end. Proof. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj1 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True equivocator_vlsm))) in Hbtr. - specialize (preloaded_with_equivocator_vlsm_trace_project_protocol _ _ _ _ Hbtr _ _ Hj) + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj1 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True equivocator_vlsm))) in Hbtr. + specialize (preloaded_with_equivocator_vlsm_trace_project_valid _ _ _ _ Hbtr _ _ Hj) as [tr [di [Hbtr_pr Hdi]]]. eexists _,_; split; [exact Hbtr_pr|]. destruct di as [sn|i]. - destruct Hdi as [Hsn Htr]. split; [assumption|]. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj2 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X))) in Htr. + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj2 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X))) in Htr. clear -Htr. destruct X as (T, (S, M)). assumption. - destruct Hdi as [s [Hpr_bs_i Htr]]. eexists; split; [exact Hpr_bs_i|]. - apply (VLSM_incl_finite_protocol_trace_from_to (VLSM_eq_proj2 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X))) in Htr. + apply (VLSM_incl_finite_valid_trace_from_to (VLSM_eq_proj2 (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True X))) in Htr. clear -Htr. destruct X as (T, (S, M)). assumption. @@ -1030,13 +1030,13 @@ Proof. Qed. (** -Projecting a protocol trace segment on an index which is valid for the +Projecting a valid trace segment on an index which is valid for the first state of the trace does not fail and yields the same index. *) -Lemma preloaded_equivocator_vlsm_trace_project_protocol_inv +Lemma preloaded_equivocator_vlsm_trace_project_valid_inv (bs : vstate equivocator_vlsm) (btr : list (vtransition_item equivocator_vlsm)) - (Hbtr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs btr) + (Hbtr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm equivocator_vlsm) bs btr) (i : nat) si (Hi : equivocator_state_project bs i = Some si) @@ -1051,7 +1051,7 @@ Proof. simpl. destruct Ht as [[_ [_ Hv]] Ht]. specialize - (equivocator_protocol_transition_item_project_inv4 l s s' iom oom Hv Ht i) + (equivocator_valid_transition_project_inv4 l s s' iom oom Hv Ht i) as Hitem. replace {| input := iom |} @@ -1067,13 +1067,13 @@ Proof. Qed. (** -An inversion lemma about projections of a protocol trace +An inversion lemma about projections of a valid trace *) -Lemma preloaded_equivocator_vlsm_protocol_trace_project_inv2 +Lemma preloaded_equivocator_vlsm_valid_trace_project_inv2 (is fs: state) (tr: list transition_item) (Hntr : tr <> []) - (Htr: finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) is fs tr) + (Htr: finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm equivocator_vlsm) is fs tr) (j : nat) (di : MachineDescriptor) (trX: list (vtransition_item X)) @@ -1081,11 +1081,11 @@ Lemma preloaded_equivocator_vlsm_protocol_trace_project_inv2 : exists fsj, equivocator_state_project fs j = Some fsj /\ match di with | NewMachine sn => - finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm X) + finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm X) sn fsj trX | Existing i => exists isi, equivocator_state_project is i = Some isi /\ - finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm X) isi fsj trX /\ + finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm X) isi fsj trX /\ (vinitial_state_prop (pre_loaded_with_all_messages_vlsm equivocator_vlsm) is -> vinitial_state_prop (pre_loaded_with_all_messages_vlsm X) isi) end. Proof. @@ -1093,10 +1093,10 @@ Proof. spec Hj. { rewrite HtrX. eexists; reflexivity. } specialize (Hj is) as [fsj Hfsj]. replace (finite_trace_last _ _) with fs in Hfsj - by (symmetry;apply (ptrace_get_last Htr)). + by (symmetry;apply (valid_trace_get_last Htr)). exists fsj. split; [assumption|]. specialize - (preloaded_equivocator_vlsm_trace_project_protocol _ _ _ Htr _ _ Hfsj) + (preloaded_equivocator_vlsm_trace_project_valid _ _ _ Htr _ _ Hfsj) as [trX' [di' [HtrX' Hdi]]]. rewrite HtrX in HtrX'. inversion HtrX'. subst di' trX'. clear HtrX'. @@ -1120,7 +1120,7 @@ Proof. apply basic_VLSM_strong_projection; intro; intros. - destruct lX as [sn| [|i] lX | [|i] lX]; inversion H; subst; assumption. - destruct lX as [sn| [|i] lX | [|i] lX]; inversion H; subst. - cbn in H0. + cbn in H0. rewrite equivocator_state_project_zero in H0. destruct (vtransition _ _ _); inversion_clear H0. reflexivity. - unfold equivocator_label_zero_project in H. @@ -1131,7 +1131,7 @@ Proof. destruct (vtransition _ _ _); inversion_clear H0; reflexivity. + destruct (equivocator_state_project _ _); [destruct (vtransition _ _ _)|]; inversion_clear H0; reflexivity. - apply H. - - apply equivocator_state_project_protocol_message. assumption. + - apply equivocator_state_project_valid_message. assumption. Qed. Lemma preloaded_equivocator_zero_projection diff --git a/theories/VLSM/Core/MessageDependencies.v b/theories/VLSM/Core/MessageDependencies.v index 81258bc6..68917ea4 100644 --- a/theories/VLSM/Core/MessageDependencies.v +++ b/theories/VLSM/Core/MessageDependencies.v @@ -20,19 +20,22 @@ Context . (** -[MessageDependencies] characterize a @message_dependencies@ function +[MessageDependencies] characterize a <> function through two properties: -- Necessity: All dependent messeges for a message @m@m are required to be -observed by any state emitting the message @m@. +- Necessity: All dependent messeges for a message <>m are required to be +observed by any state emitting the message <>. - Sufficiency: A message can be produced by the machine pre-loaded with its dependencies. + +Together with [message_dependencies_full_node_condition_prop], it +constitutes the _full node assumption_. *) Class MessageDependencies := { message_dependencies_are_necessary (m : message) - `(protocol_generated_prop (pre_loaded_with_all_messages_vlsm X) s m) + `(can_produce (pre_loaded_with_all_messages_vlsm X) s m) : forall (dm : message), dm ∈ message_dependencies m -> has_been_observed X s dm @@ -41,9 +44,9 @@ Class MessageDependencies : can_emit (pre_loaded_vlsm X (fun msg => msg ∈ message_dependencies m)) m }. -(** The (local) full node condition for a given @message_dependencies@ function +(** The (local) full node condition for a given <> function requires that a state (receiving the message) has previously -observed all of @m@'s dependencies. +observed all of <>'s dependencies. *) Definition message_dependencies_full_node_condition (s : vstate X) diff --git a/theories/VLSM/Core/Plans.v b/theories/VLSM/Core/Plans.v index 1bab615e..3bfaf884 100644 --- a/theories/VLSM/Core/Plans.v +++ b/theories/VLSM/Core/Plans.v @@ -156,7 +156,7 @@ Section apply_plans. End apply_plans. -Section protocol_plans. +Section valid_plans. Context {message : Type} @@ -188,22 +188,22 @@ Section protocol_plans. : finite_trace_last start (fst after_a) = snd after_a := (@_apply_plan_last _ (type X) (vtransition X) start a). - (** A plan is protocol w.r.t. a state if by applying it to that state we - obtain a protocol trace sequence. + (** A plan is valid w.r.t. a state if by applying it to that state we + obtain a valid trace sequence. *) - Definition finite_protocol_plan_from + Definition finite_valid_plan_from (s : vstate X) (a : plan) : Prop := - finite_protocol_trace_from _ s (fst (apply_plan s a)). + finite_valid_trace_from _ s (fst (apply_plan s a)). - Lemma finite_protocol_plan_from_app_iff + Lemma finite_valid_plan_from_app_iff (s : vstate X) (a b : plan) (s_a := snd (apply_plan s a)) - : finite_protocol_plan_from s a /\ finite_protocol_plan_from s_a b <-> finite_protocol_plan_from s (a ++ b). + : finite_valid_plan_from s a /\ finite_valid_plan_from s_a b <-> finite_valid_plan_from s (a ++ b). Proof. - unfold finite_protocol_plan_from. + unfold finite_valid_plan_from. specialize (apply_plan_app s a b) as Happ. specialize (apply_plan_last s a) as Hlst. destruct (apply_plan s a) as (aitems, afinal) eqn:Ha. @@ -211,41 +211,41 @@ Section protocol_plans. simpl in *. destruct (apply_plan afinal b) as (bitems, bfinal). rewrite Happ. simpl. clear Happ. subst afinal. - apply finite_protocol_trace_from_app_iff. + apply finite_valid_trace_from_app_iff. Qed. - Lemma finite_protocol_plan_empty + Lemma finite_valid_plan_empty (s : vstate X) - (Hpr : protocol_state_prop X s) : - finite_protocol_plan_from s []. + (Hpr : valid_state_prop X s) : + finite_valid_plan_from s []. Proof. - apply finite_ptrace_empty. + apply finite_valid_trace_from_empty. assumption. Qed. - Lemma apply_plan_last_protocol + Lemma apply_plan_last_valid (s : vstate X) (a : plan) - (Hpra : finite_protocol_plan_from s a) + (Hpra : finite_valid_plan_from s a) (after_a := apply_plan s a) : - protocol_state_prop X (snd after_a). + valid_state_prop X (snd after_a). Proof. subst after_a. rewrite <- apply_plan_last. - apply finite_ptrace_last_pstate. + apply finite_valid_trace_last_pstate. assumption. Qed. - (** By extracting a plan from a [protocol_trace] based on a state @s@ - and reapplying the plan to the same state @s@ we obtain the original trace + (** By extracting a plan from a [valid_trace] based on a state <> + and reapplying the plan to the same state <> we obtain the original trace *) Lemma trace_to_plan_to_trace_from_to (s s' : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from_to X s s' tr) + (Htr : finite_valid_trace_from_to X s s' tr) : apply_plan s (trace_to_plan tr) = (tr, s'). Proof. - induction Htr using finite_protocol_trace_from_to_rev_ind + induction Htr using finite_valid_trace_from_to_rev_ind ;[reflexivity|]. unfold trace_to_plan, _trace_to_plan. rewrite map_last, apply_plan_app. @@ -261,33 +261,33 @@ Section protocol_plans. Lemma trace_to_plan_to_trace (s : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from X s tr) + (Htr : finite_valid_trace_from X s tr) : fst (apply_plan s (trace_to_plan tr)) = tr. Proof. - apply ptrace_add_default_last, trace_to_plan_to_trace_from_to in Htr. + apply valid_trace_add_default_last, trace_to_plan_to_trace_from_to in Htr. rewrite Htr. reflexivity. Qed. - (** The plan extracted from a protocol trace is protocol w.r.t. the starting + (** The plan extracted from a valid trace is valid w.r.t. the starting state of the trace. *) - Lemma finite_protocol_trace_from_to_plan + Lemma finite_valid_trace_from_to_plan (s : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from X s tr) - : finite_protocol_plan_from s (trace_to_plan tr). + (Htr : finite_valid_trace_from X s tr) + : finite_valid_plan_from s (trace_to_plan tr). Proof. - unfold finite_protocol_plan_from. + unfold finite_valid_plan_from. rewrite trace_to_plan_to_trace; assumption. Qed. - (** Characterization of protocol plans. *) - Lemma finite_protocol_plan_iff + (** Characterization of valid plans. *) + Lemma finite_valid_plan_iff (s : vstate X) (a : plan) - : finite_protocol_plan_from s a - <-> protocol_state_prop X s - /\ Forall (fun ai => option_protocol_message_prop X (input_a ai)) a + : finite_valid_plan_from s a + <-> valid_state_prop X s + /\ Forall (fun ai => option_valid_message_prop X (input_a ai)) a /\ forall (prefa suffa : plan) (ai : plan_item) @@ -297,7 +297,7 @@ Section protocol_plans. Proof. induction a using rev_ind; repeat split; intros ; try - ( apply finite_protocol_plan_from_app_iff in H + ( apply finite_valid_plan_from_app_iff in H ; destruct H as [Ha Hx]; apply IHa in Ha as Ha'). - inversion H. assumption. - constructor. @@ -307,7 +307,7 @@ Section protocol_plans. assumption. - destruct Ha' as [_ [Hmsgs _]]. apply Forall_app. split; try assumption. - repeat constructor. unfold finite_protocol_plan_from in Hx. + repeat constructor. unfold finite_valid_plan_from in Hx. remember (snd (apply_plan s a)) as lst. unfold apply_plan, _apply_plan in Hx. simpl in Hx. destruct x. @@ -321,7 +321,7 @@ Section protocol_plans. apply app_inj_tail in Heqa. destruct Heqa; subst. unfold lst. clear lst. remember (snd (apply_plan s prefa)) as lst. - unfold finite_protocol_plan_from in Hx. + unfold finite_valid_plan_from in Hx. unfold apply_plan,_apply_plan in Hx. simpl in Hx. destruct ai. destruct ( vtransition X label_a0 (lst, input_a0)) as (dest, out). @@ -334,15 +334,15 @@ Section protocol_plans. specialize (Ha' _ _ _ eq_refl). assumption. - destruct H as [Hs [Hinput Hvalid]]. apply Forall_app in Hinput. destruct Hinput as [Hinput Hinput_ai]. - apply finite_protocol_plan_from_app_iff. - assert (Ha : finite_protocol_plan_from s a); try (split; try assumption) + apply finite_valid_plan_from_app_iff. + assert (Ha : finite_valid_plan_from s a); try (split; try assumption) ; try apply IHa; repeat split; try assumption. + intros. specialize (Hvalid prefa (suffa ++ [x]) ai). repeat rewrite app_assoc in *. subst a. specialize (Hvalid eq_refl). assumption. - + unfold finite_protocol_plan_from. + + unfold finite_valid_plan_from. specialize (Hvalid a [] x). rewrite app_assoc in Hvalid. rewrite app_nil_r in Hvalid. specialize (Hvalid eq_refl). @@ -352,8 +352,8 @@ Section protocol_plans. destruct (vtransition X label_a0 (sa, input_a0)) as (dest, out) eqn:Ht. simpl. apply Forall_inv in Hinput_ai. simpl in Hinput_ai. - unfold finite_protocol_plan_from in Ha. - apply finite_ptrace_last_pstate in Ha. + unfold finite_valid_plan_from in Ha. + apply finite_valid_trace_last_pstate in Ha. specialize (apply_plan_last s a) as Hlst. simpl in Hlst, Ha. setoid_rewrite Hlst in Ha. setoid_rewrite <- Heqsa in Ha. @@ -363,21 +363,21 @@ Section protocol_plans. with (vtransition X label_a0 (sa, input_a0)). destruct Ha as [_oma Hsa]. destruct Hinput_ai as [_s Hinput_a0]. - apply protocol_generated with sa _oma _s input_a0 label_a0; assumption. + apply valid_generated_state_message with sa _oma _s input_a0 label_a0; assumption. Qed. - (** Characterizing a singleton protocol plan as a protocol transition. *) - Lemma finite_protocol_plan_from_one + (** Characterizing a singleton valid plan as a input valid transition. *) + Lemma finite_valid_plan_from_one (s : vstate X) (a : plan_item) : let res := vtransition X (label_a a) (s, input_a a) in - finite_protocol_plan_from s [a] <-> protocol_transition X (label_a a) (s, input_a a) res. + finite_valid_plan_from s [a] <-> input_valid_transition X (label_a a) (s, input_a a) res. Proof. split; intros; destruct a; unfold apply_plan,_apply_plan in *; simpl in *; - unfold finite_protocol_plan_from in *; + unfold finite_valid_plan_from in *; unfold apply_plan, _apply_plan in *; simpl in *. - match type of H with | context[let (_, _) := let (_, _) := ?t in _ in _] => @@ -386,13 +386,13 @@ Section protocol_plans. inversion H. subst. setoid_rewrite eq_trans. assumption. - match type of H with - | protocol_transition _ _ _ ?t => + | input_valid_transition _ _ _ ?t => destruct t as [dest output] eqn : eq_trans end. setoid_rewrite eq_trans. - apply finite_ptrace_extend. - apply finite_ptrace_empty. - apply protocol_transition_destination in H; intuition. + apply finite_valid_trace_from_extend. + apply finite_valid_trace_from_empty. + apply input_valid_transition_destination in H; intuition. assumption. Qed. @@ -401,41 +401,41 @@ Section protocol_plans. (P : vstate X -> Prop) : Prop := forall (s : vstate X), - (P s -> protocol_state_prop X s -> finite_protocol_plan_from s a -> P (snd (apply_plan s a))). + (P s -> valid_state_prop X s -> finite_valid_plan_from s a -> P (snd (apply_plan s a))). Definition ensures (a : plan) (P : vstate X -> Prop) : Prop := forall (s : vstate X), - (protocol_state_prop X s -> P s -> finite_protocol_plan_from s a). + (valid_state_prop X s -> P s -> finite_valid_plan_from s a). - (* If some property of a state guarantees a plan `b` applied to the state is protocol, + (* If some property of a state guarantees a plan `b` applied to the state is valid, and this property is preserved by the application of some other plan `a`, then these two plans can be composed and the application of `a ++ b` will also - be protocol. *) + be valid. *) Lemma plan_independence (a b : plan) (Pb : vstate X -> Prop) (s : state) - (Hpr : protocol_state_prop X s) - (Ha : finite_protocol_plan_from s a) + (Hpr : valid_state_prop X s) + (Ha : finite_valid_plan_from s a) (Hhave : Pb s) (Hensures : ensures b Pb) (Hpreserves : preserves a Pb) : - finite_protocol_plan_from s (a ++ b). + finite_valid_plan_from s (a ++ b). Proof. unfold ensures in *. unfold preserves in *. - apply finite_protocol_plan_from_app_iff. + apply finite_valid_plan_from_app_iff. split. - assumption. - remember (snd (apply_plan s a)) as s'. specialize (Hensures s'). apply Hensures. rewrite Heqs'. - apply apply_plan_last_protocol. + apply apply_plan_last_valid. intuition. intuition. rewrite Heqs'. @@ -443,4 +443,4 @@ Section protocol_plans. all : intuition. Qed. -End protocol_plans. +End valid_plans. diff --git a/theories/VLSM/Core/ProjectionTraces.v b/theories/VLSM/Core/ProjectionTraces.v index 07542cd7..d810adc2 100644 --- a/theories/VLSM/Core/ProjectionTraces.v +++ b/theories/VLSM/Core/ProjectionTraces.v @@ -25,7 +25,7 @@ Let us fix an indexed set of VLSMs <> and their composition <> using <> to component <> is the type of the <>th component of <>. We defined the signature of the projection to be the same as that of the component, with the exception that the [initial_message]s for the projection are defined -to be all [protocol_message]s of <>: +to be all [valid_message]s of <>: *) Definition composite_vlsm_constrained_projection_sig @@ -53,12 +53,12 @@ to be all [protocol_message]s of <>: : VLSMSign (type (IM i)) := {| initial_state_prop := vinitial_state_prop (IM i) - ; initial_message_prop := fun pmi => exists xm : protocol_message X, proj1_sig xm = pmi + ; initial_message_prop := fun pmi => exists xm : valid_message X, proj1_sig xm = pmi ; s0 := @s0 _ _ (sign (IM i)) |}. (** -[projection_valid]ity is defined as the projection of [protocol_valid]ity of <>: +[projection_valid]ity is defined as the projection of [input_valid]ity of <>: *) Definition projection_valid @@ -68,7 +68,7 @@ to be all [protocol_message]s of <>: := let (si, omi) := siomi in exists (s : vstate X), - s i = si /\ protocol_valid X (existT i li) (s, omi). + s i = si /\ input_valid X (existT i li) (s, omi). (** The following two lemmas ([projection_valid_impl_VLSM1_projection_valid] @@ -98,7 +98,7 @@ to be all [protocol_message]s of <>: { subst si. apply Hcvalid. } unfold projected_state_prop. unfold vvalid in Hcvalid. - unfold protocol_state. + unfold valid_state. remember (@composite_label _ index IM) as CL in |-. remember (existT i li) as er. remember (vtransition X er (s,omi)) as sm'. @@ -107,9 +107,9 @@ to be all [protocol_message]s of <>: destruct sm' as [s'' om']. simpl in Heqs'. subst s''. - assert (Hpt : protocol_transition X er (s,omi) (s', om')). + assert (Hivt : input_valid_transition X er (s,omi) (s', om')). { - unfold protocol_transition. + unfold input_valid_transition. split. { apply Hpv. } unfold vtransition in Heqsm'. @@ -117,17 +117,17 @@ to be all [protocol_message]s of <>: apply Heqsm'. } - pose proof (Hps' := protocol_transition_destination X Hpt). + pose proof (Hps' := input_valid_transition_destination X Hivt). split. { exists (exist _ s' Hps'). pose proof (H := @composite_transition_state_eq message index IndEqDec IM constraint er). - specialize (H s s' omi om' Hpt). rewrite Heqer in H. simpl in H. + specialize (H s s' omi om' Hivt). rewrite Heqer in H. simpl in H. rewrite <- Hsi. rewrite <- H. simpl. reflexivity. } - unfold option_protocol_message_prop. + unfold option_valid_message_prop. fold (vstate X). assert (Hveq: snd (vtransition (IM i) li (si, omi)) = snd (vtransition X er (s, omi))). @@ -146,7 +146,7 @@ to be all [protocol_message]s of <>: rewrite Hveq. rewrite <- Heqsm'. exists s'. simpl. - apply protocol_prop_transition_out in Hpt. + apply input_valid_transition_outputs_valid_state_message in Hivt. assumption. Qed. @@ -156,10 +156,10 @@ to be all [protocol_message]s of <>: (siomi : vstate (IM i) * option message) : VLSM1_projection_valid i li siomi -> - (exists s : protocol_state X, + (exists s : valid_state X, proj1_sig s i = (fst siomi) /\ constraint (existT i li) (proj1_sig s, (snd siomi))) -> - option_protocol_message_prop X (snd siomi) -> + option_valid_message_prop X (snd siomi) -> projection_valid i li siomi. Proof. unfold projection_valid. @@ -173,7 +173,7 @@ to be all [protocol_message]s of <>: exists (proj1_sig s'). split. { apply Hs'. } - unfold protocol_valid. + unfold input_valid. split. { exact (proj2_sig s'). } @@ -190,7 +190,7 @@ to be all [protocol_message]s of <>: Qed. (** -Since [projection_valid]ity is derived from [protocol_valid]ity, which in turn +Since [projection_valid]ity is derived from [input_valid]ity, which in turn depends on [valid]ity in the component, it is easy to see that [projection_valid]ity implies [valid]ity in the component. *) @@ -211,7 +211,7 @@ depends on [valid]ity in the component, it is easy to see that (** We define the projection of <> to index <> as the [VLSM] whose signature is the [composite_vlsm_constrained_projection_sig]nature corresponding to <>, -having the same transition function as <>, the <>th component of +having the same transition function as <>, the <>th component of <>. *) Definition composite_vlsm_constrained_projection_machine (i : index) @@ -238,9 +238,8 @@ with all messages (Lemma [preloaded_component_projection]). We then study the extension of these definitions and results to infinite traces. -Finally, we define and prove some properties of the [projection_friendly] -property which guarantees that any protocol trace of the projection can -be "lifted" to a protocol trace of the composition which projects to it. +Finally, we prove some consequences of the [projection_friendly_prop]erty for +the specific case of projecting a trace to a single component. *) Section ProjectionTraces. @@ -283,20 +282,23 @@ Proof. Qed. (** -Since all [protocol_message]s of <> become [initial_message]s in <>, the +Since all [valid_message]s of <> become [initial_message]s in <>, the following result is not surprising. *) -Lemma protocol_message_projection +Lemma valid_message_projection (iom : option message) - (HpmX : option_protocol_message_prop X iom) - : option_protocol_message_prop Xj iom. + (HpmX : option_valid_message_prop X iom) + : option_valid_message_prop Xj iom. Proof. - apply option_initial_message_is_protocol. + apply option_initial_message_is_valid. destruct iom as [m|];[|exact I]. exists (exist _ m HpmX). reflexivity. Qed. +(** The projection on component <> of valid traces from <> is valid +for the <>th projection +*) Lemma component_projection : VLSM_projection X Xj composite_project_label (fun s => s j). Proof. apply basic_VLSM_projection; intro; intros. @@ -322,26 +324,26 @@ Proof. inversion H0. rewrite state_update_neq by congruence. reflexivity. - apply initial_state_projection. assumption. - - apply protocol_message_projection. apply Hv. + - apply valid_message_projection. apply Hv. Qed. -(* The projection of a finite protocol trace remains a protocol trace *) -Lemma finite_ptrace_projection +(* The projection of a finite valid trace remains a valid trace *) +Lemma finite_valid_trace_projection (s : vstate X) (trx : list (vtransition_item X)) - (Htr : finite_protocol_trace_from X s trx) - : finite_protocol_trace_from Xj (s j) (VLSM_projection_trace_project component_projection trx). + (Htr : finite_valid_trace_from X s trx) + : finite_valid_trace_from Xj (s j) (VLSM_projection_trace_project component_projection trx). Proof. revert Htr. - apply (VLSM_projection_finite_protocol_trace_from component_projection). + apply (VLSM_projection_finite_valid_trace_from component_projection). Qed. -Lemma protocol_state_projection +Lemma valid_state_projection (s : vstate X) - (Hps : protocol_state_prop X s) - : protocol_state_prop Xj (s j). + (Hps : valid_state_prop X s) + : valid_state_prop Xj (s j). Proof. - revert Hps. apply (VLSM_projection_protocol_state component_projection). + revert Hps. apply (VLSM_projection_valid_state component_projection). Qed. Lemma in_futures_projection @@ -353,48 +355,6 @@ Proof. apply (VLSM_projection_in_futures component_projection). Qed. -(* We axiomatize projection friendliness as the converse of protocol_trace_projection *) -Definition finite_projection_friendly - := forall - (sj : vstate (IM j)) - (trj : list (vtransition_item (IM j))) - (Htrj : finite_protocol_trace Xj sj trj), - exists (sx : vstate X) (trx : list (vtransition_item X)), - finite_protocol_trace X sx trx - /\ sx j = sj - /\ VLSM_projection_trace_project component_projection trx = trj. - -Lemma projection_friendly_in_futures - (Hfr : finite_projection_friendly) - (s1 s2 : vstate (IM j)) - (Hfuture : in_futures Xj s1 s2) - : exists (sX1 sX2 : vstate X), - sX1 j = s1 /\ sX2 j = s2 /\ in_futures X sX1 sX2. -Proof. - destruct Hfuture as [tr_s2 Hfuture]. - apply finite_protocol_trace_from_to_complete_left in Hfuture - as [is [tr_s1 [Htr Heq_s1]]]. - apply ptrace_get_last in Htr as Heq_s2. - apply ptrace_forget_last in Htr. - apply Hfr in Htr as [isX [trX [Htr [His Htr_pr]]]]. - apply VLSM_projection_trace_project_app_rev in Htr_pr - as [trX_s1 [trX_s2 [HeqtrX [Htr_s1_pr Htr_s2_pr]]]]. - subst. - destruct Htr as [HtrX HisX]. - apply finite_protocol_trace_from_app_iff in HtrX as HtrX12. - destruct HtrX12 as [HtrX1 HtrX2]. - apply ptrace_add_default_last in HtrX2. - exists (finite_trace_last isX trX_s1). - exists (finite_trace_last isX (trX_s1 ++ trX_s2)). - rewrite !(VLSM_projection_finite_trace_last component_projection), - (VLSM_projection_trace_project_app component_projection). - - repeat split. - rewrite finite_trace_last_app. - eexists; exact HtrX2. - - assumption. - - assumption. -Qed. - End ProjectionTraces. Section PreLoadedProjectionTraces. @@ -407,7 +367,7 @@ Context (j : index) . -(* The projection of a preloaded finite protocol trace remains a preloaded protocol trace *) +(* The projection of a preloaded finite valid trace remains a preloaded valid trace *) Lemma preloaded_component_projection : VLSM_projection (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) (pre_loaded_with_all_messages_vlsm (IM j)) (composite_project_label IM j) (fun s => s j). Proof. apply basic_VLSM_projection_preloaded; intro; intros. @@ -491,65 +451,65 @@ Definition finite_trace_projection_list (tr : list (composite_transition_item IM @pre_VLSM_projection_trace_project _ (composite_type IM) _ (composite_project_label IM j) (fun s => s j) tr. -Lemma preloaded_protocol_state_projection +Lemma preloaded_valid_state_projection (s : state) - (Hps : protocol_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s) - : protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM j)) (s j). + (Hps : valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s) + : valid_state_prop (pre_loaded_with_all_messages_vlsm (IM j)) (s j). Proof. - revert Hps. apply (VLSM_projection_protocol_state preloaded_component_projection). + revert Hps. apply (VLSM_projection_valid_state preloaded_component_projection). Qed. -Lemma preloaded_finite_ptrace_projection +Lemma preloaded_finite_valid_trace_projection (s : composite_state IM) (trx : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s trx) - : finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (VLSM_projection_trace_project preloaded_component_projection trx). + (Htr : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s trx) + : finite_valid_trace_from (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (VLSM_projection_trace_project preloaded_component_projection trx). Proof. - revert Htr. apply (VLSM_projection_finite_protocol_trace_from preloaded_component_projection). + revert Htr. apply (VLSM_projection_finite_valid_trace_from preloaded_component_projection). Qed. -Lemma preloaded_finite_ptrace_from_to_projection +Lemma preloaded_finite_valid_trace_from_to_projection (s s' : composite_state IM) (trx : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s s' trx) - : finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (s' j) (VLSM_projection_trace_project preloaded_component_projection trx). + (Htr : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s s' trx) + : finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (s' j) (VLSM_projection_trace_project preloaded_component_projection trx). Proof. - revert Htr. apply (VLSM_projection_finite_protocol_trace_from_to preloaded_component_projection). + revert Htr. apply (VLSM_projection_finite_valid_trace_from_to preloaded_component_projection). Qed. -Lemma preloaded_finite_ptrace_init_to_projection +Lemma preloaded_finite_valid_trace_init_to_projection (s s' : composite_state IM) (trx : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s s' trx) - : finite_protocol_trace_init_to (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (s' j) (VLSM_projection_trace_project preloaded_component_projection trx). + (Htr : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) s s' trx) + : finite_valid_trace_init_to (pre_loaded_with_all_messages_vlsm (IM j)) (s j) (s' j) (VLSM_projection_trace_project preloaded_component_projection trx). Proof. - revert Htr. apply (VLSM_projection_finite_protocol_trace_init_to preloaded_component_projection). + revert Htr. apply (VLSM_projection_finite_valid_trace_init_to preloaded_component_projection). Qed. -Lemma pre_loaded_with_all_messages_projection_protocol_transition_eq +Lemma pre_loaded_with_all_messages_projection_input_valid_transition_eq (s1 s2 : composite_state IM) (om1 om2 : option message) (l : label) - (Ht : protocol_transition (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) l (s1, om1) (s2, om2)) + (Ht : input_valid_transition (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) l (s1, om1) (s2, om2)) (Hl : projT1 l = j) - : protocol_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) (s1 (projT1 l), om1) (s2 (projT1 l), om2). + : input_valid_transition (pre_loaded_with_all_messages_vlsm (IM (projT1 l))) (projT2 l) (s1 (projT1 l), om1) (s2 (projT1 l), om2). Proof. specialize - (VLSM_projection_protocol_transition preloaded_component_projection l) as Hpt. - subst j. specialize (Hpt (projT2 l)). - spec Hpt. + (VLSM_projection_input_valid_transition preloaded_component_projection l) as Hivt. + subst j. specialize (Hivt (projT2 l)). + spec Hivt. { unfold composite_project_label. destruct (decide _); [| elim n; reflexivity]. replace e with (eq_refl (A := index) (x := projT1 l)); [reflexivity|]. apply Eqdep_dec.UIP_dec. assumption. } - apply Hpt in Ht. assumption. + apply Hivt in Ht. assumption. Qed. -Lemma pre_loaded_with_all_messages_projection_protocol_transition_neq +Lemma pre_loaded_with_all_messages_projection_input_valid_transition_neq [s1 s2 : composite_state IM] [om1 om2 : option message] [l : label] - (Ht : protocol_transition (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) l (s1, om1) (s2, om2)) + (Ht : input_valid_transition (pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM)) l (s1, om1) (s2, om2)) [i : index] (Hi : i <> projT1 l) : s1 i = s2 i. @@ -658,10 +618,10 @@ component <>. In this section we establish some basic properties for projections, building up to Lemma [proj_pre_loaded_with_all_messages_incl], which guarantees that all -[protocol_trace]s of <> are also [protocol_trace]s for the +[valid_trace]s of <> are also [valid_trace]s for the [pre_loaded_with_all_messages_vlsm] associated to the component <>. In particular this ensures that the byzantine traces of <> include all -[protocol_trace]s of <> (see Lemma [pre_loaded_with_all_messages_alt_eq]). +[valid_trace]s of <> (see Lemma [pre_loaded_with_all_messages_alt_eq]). *) @@ -670,94 +630,94 @@ Context (Xj := composite_vlsm_constrained_projection IM constraint j) . -Lemma projection_valid_protocol +Lemma projection_valid_input_valid (l : vlabel Xj) (som : vstate Xj * option message) (Hv : vvalid Xj l som) - : protocol_valid Xj l som. + : input_valid Xj l som. Proof. destruct som as (s, om). destruct (id Hv) as [sX [Hsi [Hps [Hopm _]]]]. repeat split. - - subst. apply protocol_state_projection. assumption. - - apply protocol_message_projection. assumption. + - subst. apply valid_state_projection. assumption. + - apply valid_message_projection. assumption. - assumption. Qed. -Lemma projection_valid_implies_composition_protocol_message +Lemma projection_valid_implies_composition_valid_message (l : label) (s : state) (om : option message) (Hv : vvalid Xj l (s, om)) - : option_protocol_message_prop X om. + : option_valid_message_prop X om. Proof. destruct Hv as [sx [Hs [HpsX [HpmX Hv]]]]. assumption. Qed. -Lemma projection_valid_implies_projection_protocol_message +Lemma projection_valid_implies_projection_valid_message (l : label) (s : state) (om : option message) (Hv : vvalid Xj l (s, om)) - : option_protocol_message_prop Xj om. + : option_valid_message_prop Xj om. Proof. - apply protocol_message_projection. + apply valid_message_projection. revert Hv. - apply projection_valid_implies_composition_protocol_message. + apply projection_valid_implies_composition_valid_message. Qed. -Lemma projection_valid_implies_projection_protocol_state +Lemma projection_valid_implies_projection_valid_state (lj : label) (sj : state) (om : option message) (Hv : vvalid Xj lj (sj, om)) - : protocol_state_prop Xj sj. + : valid_state_prop Xj sj. Proof. destruct Hv as [s [Heq_sj [Hs _]]]. - subst sj. revert Hs. apply protocol_state_projection. + subst sj. revert Hs. apply valid_state_projection. Qed. -Lemma projection_valid_implies_destination_projection_protocol_prop +Lemma projection_valid_implies_projection_valid_state_message_outputs (l : label) (s : state) (om : option message) (Hv : vvalid Xj l (s, om)) s' om' (Ht : vtransition (IM j) l (s, om) = (s', om')) - : protocol_prop Xj s' om'. + : valid_state_message_prop Xj s' om'. Proof. - apply projection_valid_implies_projection_protocol_state in Hv as Hs. + apply projection_valid_implies_projection_valid_state in Hv as Hs. destruct Hs as [_om Hs]. - apply projection_valid_implies_projection_protocol_message in Hv as Hom. + apply projection_valid_implies_projection_valid_message in Hv as Hom. destruct Hom as [_s Hom]. - apply (protocol_generated Xj _ _ Hs _ _ Hom _ Hv _ _ Ht). + apply (valid_generated_state_message Xj _ _ Hs _ _ Hom _ Hv _ _ Ht). Qed. -Lemma projection_valid_implies_destination_projection_protocol_state +Lemma projection_valid_implies_destination_projection_valid_state (l : label) (s : state) (om : option message) (Hv : vvalid Xj l (s, om)) s' om' (Ht : vtransition (IM j) l (s, om) = (s', om')) - : protocol_state_prop Xj s'. + : valid_state_prop Xj s'. Proof. - apply projection_valid_implies_destination_projection_protocol_prop + apply projection_valid_implies_projection_valid_state_message_outputs with (s' := s') (om' := om') in Hv; [|assumption]. eexists. apply Hv. Qed. -Lemma projection_valid_implies_destination_projection_protocol_message +Lemma projection_valid_implies_destination_projection_valid_message (l : label) (s : state) (om : option message) (Hv : vvalid Xj l (s, om)) s' om' (Ht : vtransition (IM j) l (s, om) = (s', om')) - : option_protocol_message_prop Xj om'. + : option_valid_message_prop Xj om'. Proof. - apply projection_valid_implies_destination_projection_protocol_prop + apply projection_valid_implies_projection_valid_state_message_outputs with (s' := s') (om' := om') in Hv; [|assumption]. eexists. apply Hv. Qed. @@ -766,12 +726,12 @@ Qed. Interestingly enough, <> cannot produce any additional messages than the initial ones available from <>. *) -Lemma protocol_message_projection_rev +Lemma valid_message_projection_rev (iom : option message) - (Hpmj: option_protocol_message_prop Xj iom) - : option_protocol_message_prop X iom. + (Hpmj: option_valid_message_prop Xj iom) + : option_valid_message_prop X iom. Proof. - destruct iom as [m|];[|apply option_protocol_message_None]. + destruct iom as [m|];[|apply option_valid_message_None]. destruct Hpmj as [sj Hpmj]. inversion Hpmj; subst. - destruct Hom as [pm <-]. apply @proj2_sig. @@ -779,26 +739,26 @@ Proof. subst s. set (lX := existT j l) in Hv. eexists. - apply (protocol_prop_valid_out X _ _ _ Hv). + apply (input_valid_state_message_outputs X _ _ _ Hv). simpl. replace (vtransition (IM j) _ _) with (sj, Some m). reflexivity. Qed. (** As a stepping stone towards proving trace inclusion between <> and the [pre_loaded_with_all_messages_vlsm] associated to <>, we prove that the -[protocol_prop]erty is transferred. +[valid_state_message_prop]erty is transferred. *) -Lemma proj_pre_loaded_with_all_messages_protocol_prop +Lemma proj_pre_loaded_with_all_messages_valid_state_message_preservation (PreLoaded := pre_loaded_with_all_messages_vlsm (IM j)) (s : state) (om : option message) - (Hps : protocol_prop Xj s om) - : protocol_prop PreLoaded s om. + (Hps : valid_state_message_prop Xj s om) + : valid_state_message_prop PreLoaded s om. Proof. induction Hps. - - apply (protocol_initial PreLoaded). + - apply (valid_initial_state_message PreLoaded). assumption. destruct om;exact I. - - apply (protocol_generated PreLoaded) with s _om _s om l; try assumption. + - apply (valid_generated_state_message PreLoaded) with s _om _s om l; try assumption. eapply (projection_valid_implies_valid IM). exact Hv. Qed. @@ -811,7 +771,7 @@ Lemma proj_pre_loaded_with_all_messages_incl Proof. apply (basic_VLSM_incl (machine Xj) (machine PreLoaded)); intro; intros. - assumption. - - apply initial_message_is_protocol; exact I. + - apply initial_message_is_valid; exact I. - eapply (projection_valid_implies_valid IM). apply Hv. - apply H. @@ -831,7 +791,7 @@ Context {message : Type} . -(** ** A sufficient condition for being [projection_friendly]. *) +(** ** A sufficient condition for the [projection_friendly_prop]erty. *) Context (j : index) @@ -839,8 +799,8 @@ Context . (** -This condition states that [protocol_valid]ity in a projection <> -can be lifted to any [protocol_state] in <> which projects to the +This condition states that [input_valid]ity in a projection <> +can be lifted to any [valid_state] in <> which projects to the corresponding <> state. *) @@ -849,29 +809,29 @@ Definition projection_friendliness_sufficient_condition (lj : vlabel (IM j)) (sj : vstate (IM j)) (om : option message) - (Hpv : protocol_valid Xj lj (sj, om)) + (Hiv : input_valid Xj lj (sj, om)) (s : vstate X) - (Hs : protocol_state_prop X s) + (Hs : valid_state_prop X s) (Hsi : s j = sj) , vvalid X (existT j lj) (s, om). -Lemma projection_friendliness_sufficient_condition_protocol_state +Lemma projection_friendliness_sufficient_condition_valid_state (Hfr : projection_friendliness_sufficient_condition) (s : state) - (Hp : protocol_state_prop Xj s) - : protocol_state_prop X (lift_to_composite_state IM j s). + (Hp : valid_state_prop Xj s) + : valid_state_prop X (lift_to_composite_state IM j s). Proof. - induction Hp using protocol_state_prop_ind. - - apply initial_is_protocol. apply (lift_to_composite_state_initial IM j). assumption. + induction Hp using valid_state_prop_ind. + - apply initial_state_is_valid. apply (lift_to_composite_state_initial IM j). assumption. - destruct Ht as [Hvj Ht]. specialize (Hfr _ _ _ Hvj _ IHHp). spec Hfr; [apply state_update_eq|]. exists om'. destruct Hvj as [_ [_ Hvj]]. - apply (projection_valid_implies_composition_protocol_message IM) in Hvj as Hom. + apply (projection_valid_implies_composition_valid_message IM) in Hvj as Hom. destruct IHHp as [_om HsX]. destruct Hom as [_s Hom]. - specialize (protocol_generated X _ _ HsX _ _ Hom _ Hfr) as Hgen. + specialize (valid_generated_state_message X _ _ HsX _ _ Hom _ Hfr) as Hgen. apply Hgen. simpl. unfold lift_to_composite_state at 1. @@ -893,7 +853,7 @@ Lemma projection_friendliness_lift_to_composite_vlsm_full_projection Proof. apply basic_VLSM_full_projection; intro; intros. - apply (Hfr _ _ _ Hv); [|apply state_update_eq]. - apply (projection_friendliness_sufficient_condition_protocol_state Hfr). + apply (projection_friendliness_sufficient_condition_valid_state Hfr). apply Hv. - unfold lift_to_composite_label, vtransition. simpl. unfold lift_to_composite_state at 1. rewrite state_update_eq. diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index 6c90c07e..ea38b81d 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -128,7 +128,7 @@ Proof. reflexivity. Qed. -Lemma composite_state_sub_projection_lift +Lemma composite_state_sub_projection_lift_to (s0 : composite_state IM) (s : composite_state sub_IM) : composite_state_sub_projection (lift_sub_state_to s0 s) = s. @@ -163,7 +163,7 @@ Proof. apply state_update_neq. congruence. Qed. -Section induced_sub_projection. +Section sec_induced_sub_projection. Context (constraint : composite_label IM -> composite_state IM * option message -> Prop) @@ -178,11 +178,40 @@ Definition composite_label_sub_projection_option | _ => None end. +(** By restricting the components of a composition to a subset we obtain a +[projection_induced_vlsm]. +*) +Definition induced_sub_projection : VLSM message := + projection_induced_vlsm X (composite_type sub_IM) + composite_label_sub_projection_option + composite_state_sub_projection + lift_sub_label lift_sub_state. + +Lemma induced_sub_projection_transition_consistency_None + : weak_projection_transition_consistency_None X (composite_type sub_IM) + composite_label_sub_projection_option composite_state_sub_projection. +Proof. + intros lX HlX sX om s'X om' HtX. + extensionality sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst sub_i. + unfold composite_state_sub_projection. simpl. + destruct lX as [x v]. + apply proj2 in HtX. cbn in HtX. + destruct (vtransition _ _ _) as (si', _om'). + inversion_clear HtX. + rewrite state_update_neq; [reflexivity|]. + intros ->. + unfold composite_label_sub_projection_option in HlX. + simpl in HlX. + case_decide; [discriminate|contradiction]. +Qed. + Lemma composite_label_sub_projection_option_lift - (l : composite_label sub_IM) - : composite_label_sub_projection_option (lift_sub_label l) = Some l. + : induced_projection_label_lift_prop (free_composite_vlsm IM) (composite_type sub_IM) + composite_label_sub_projection_option lift_sub_label. Proof. - destruct l as (sub_i, li). + intros (sub_i, li). destruct_dec_sig sub_i i Hi Heqsub_i. subst. unfold lift_sub_label, composite_label_sub_projection_option. @@ -195,20 +224,12 @@ Proof. reflexivity. Qed. -Lemma composite_transition_item_sub_projection_lift - (item : composite_transition_item sub_IM) - : @pre_VLSM_projection_transition_item_project _ (composite_type IM) _ - composite_label_sub_projection_option composite_state_sub_projection - (pre_VLSM_full_projection_trace_item_project _ _ lift_sub_label lift_sub_state item) - = Some item. +Lemma composite_state_sub_projection_lift + : induced_projection_state_lift_prop (free_composite_vlsm IM) (composite_type sub_IM) + composite_state_sub_projection lift_sub_state. Proof. - destruct item. - unfold pre_VLSM_full_projection_trace_item_project, - pre_VLSM_projection_transition_item_project. - simpl. - rewrite composite_label_sub_projection_option_lift. - f_equal. f_equal. - apply composite_state_sub_projection_lift. + intro. + apply composite_state_sub_projection_lift_to. Qed. Lemma composite_trace_sub_projection_lift @@ -218,21 +239,60 @@ Lemma composite_trace_sub_projection_lift (pre_VLSM_full_projection_finite_trace_project _ _ lift_sub_label lift_sub_state tr) = tr. Proof. - induction tr; [reflexivity|]. - simpl. - rewrite composite_transition_item_sub_projection_lift. - f_equal. - assumption. + apply (induced_projection_trace_lift (free_composite_vlsm IM)). + - apply composite_label_sub_projection_option_lift. + - apply composite_state_sub_projection_lift. Qed. -(** By restricting the components of a composition to a subset we obtain a -[projection_induced_vlsm]. -*) -Definition induced_sub_projection : VLSM message := - projection_induced_vlsm X (composite_type sub_IM) - composite_label_sub_projection_option - composite_state_sub_projection - lift_sub_label lift_sub_state. +Lemma induced_sub_projection_transition_consistency_Some + : induced_projection_transition_consistency_Some X (composite_type sub_IM) + composite_label_sub_projection_option composite_state_sub_projection. +Proof. + intros lX1 lX2 lY HlX1_pr HlX2_pr sX1 sX2 HsXeq_pr iom sX1' oom1 Ht1 sX2' oom2 Ht2. + destruct lX1 as (i, lXi). + unfold composite_label_sub_projection_option in HlX1_pr. + simpl in HlX1_pr. case_decide as Hi; [|discriminate]. + apply Some_inj in HlX1_pr. subst lY. + destruct lX2 as (_i, _lXi). + unfold composite_label_sub_projection_option in HlX2_pr. + simpl in HlX2_pr. case_decide as H_i; [|discriminate]. + apply Some_inj in HlX2_pr. + unfold composite_label_sub_projection in HlX2_pr. + simpl in HlX2_pr. + inversion HlX2_pr. + subst _i. + apply + (@dec_sig_sigT_eq_rev _ _ _ sub_index_prop_dec (fun i => vlabel (IM i))) + in HlX2_pr as ->. + apply f_equal_dep with (x := dexist i Hi) in HsXeq_pr as HsXeq_pri. + cbv in HsXeq_pri. + cbn in Ht1, Ht2. + rewrite <- HsXeq_pri in Ht2. + destruct (vtransition _ _ _) as (si', om'). + inversion Ht1. subst. clear Ht1. + inversion Ht2. subst. clear Ht2. + split; [|reflexivity]. + extensionality sub_j. + apply f_equal_dep with (x := sub_j) in HsXeq_pr. + destruct_dec_sig sub_j j Hj Heqsub_j. + subst. + unfold composite_state_sub_projection in HsXeq_pr |- *. + simpl in HsXeq_pr |- *. + destruct (decide (i = j)). + - subst. rewrite !state_update_eq. reflexivity. + - rewrite !state_update_neq by congruence. assumption. +Qed. + +Lemma weak_induced_sub_projection_transition_consistency_Some + : weak_projection_transition_consistency_Some X (composite_type sub_IM) + composite_label_sub_projection_option composite_state_sub_projection + lift_sub_label lift_sub_state. +Proof. + apply basic_weak_projection_transition_consistency_Some. + - apply composite_label_sub_projection_option_lift. + - apply composite_state_sub_projection_lift. + - apply induced_sub_projection_transition_consistency_Some. +Qed. (** The [induced_sub_projection] is actually a [VLSM_projection] of the original composition. @@ -243,53 +303,14 @@ Lemma induced_sub_projection_is_projection composite_state_sub_projection. Proof. apply projection_induced_vlsm_is_projection. - - intros lX HlX sX om s'X om' HtX. - extensionality sub_i. - destruct_dec_sig sub_i i Hi Heqsub_i. - subst sub_i. - unfold composite_state_sub_projection. simpl. - apply proj2 in HtX. destruct lX. cbn in HtX. - destruct (vtransition _ _ _) as (si', _om'). - inversion_clear HtX. - rewrite state_update_neq; [reflexivity|]. - intro. subst. - unfold composite_label_sub_projection_option in HlX. - simpl in HlX. - case_decide; [discriminate|contradiction]. - - intros lX lY Hl s1 om s1' om1' Ht1 s2' om2' Ht2. - apply proj2 in Ht1. - destruct lX as (i, lX). - unfold composite_label_sub_projection_option in Hl. - simpl in Hl. case_decide; [|discriminate]. - inversion Hl. subst lY. clear Hl. - cbn in Ht1, Ht2. - unfold lift_sub_state at 1 in Ht2. - rewrite - (lift_sub_state_to_eq (fun (n : index) => proj1_sig (vs0 (IM n))) - (composite_state_sub_projection s1) i H) in Ht2. - unfold composite_state_sub_projection in Ht2. simpl in Ht2. - destruct (vtransition _ _ _) as (si', om'). - inversion Ht1. subst. - inversion Ht2. subst. - split; [|reflexivity]. - extensionality sub_j. - destruct_dec_sig sub_j j Hj Heqsub_j. - subst. - unfold composite_state_sub_projection. simpl. - destruct (decide (i = j)). - + subst. rewrite! state_update_eq. reflexivity. - + rewrite! state_update_neq by congruence. - unfold lift_sub_state. - rewrite - (lift_sub_state_to_eq (fun (n : index) => proj1_sig (vs0 (IM n))) - (λ subi : sub_index, s1 (`subi)) j Hj). - reflexivity. + - apply induced_sub_projection_transition_consistency_None. + - apply weak_induced_sub_projection_transition_consistency_Some. Qed. Lemma induced_sub_projection_valid_projection l s om (Hv : vvalid induced_sub_projection l (s, om)) : exists i, i ∈ sub_index_list /\ - exists l s, protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, om). + exists l s, input_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, om). Proof. destruct l as (sub_i, li). destruct Hv as [lX [sX [HlX [Heqs [HsX [Hom Hv]]]]]]. @@ -307,13 +328,14 @@ Proof. apply proj1 in Hv. cbn in Hv. exists li, (sX i). - repeat split; [|apply any_message_is_protocol_in_preloaded|assumption]. - apply (VLSM_projection_protocol_state (preloaded_component_projection IM i)). - apply (VLSM_incl_protocol_state (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))). - by apply (VLSM_incl_protocol_state (constraint_free_incl IM constraint)). + repeat split; [|apply any_message_is_valid_in_preloaded|assumption]. + apply (VLSM_projection_valid_state (preloaded_component_projection IM i)). + apply (VLSM_incl_valid_state (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))). + apply (VLSM_incl_valid_state (constraint_free_incl IM constraint)). + assumption. Qed. -End induced_sub_projection. +End sec_induced_sub_projection. Section induced_sub_projection_subsumption. @@ -323,19 +345,14 @@ Context . Lemma induced_sub_projection_constraint_subsumption_incl - (Hsubsumption : protocol_constraint_subsumption IM constraint1 constraint2) + (Hsubsumption : input_valid_constraint_subsumption IM constraint1 constraint2) : VLSM_incl (induced_sub_projection constraint1) (induced_sub_projection constraint2). Proof. - specialize (constraint_subsumption_incl IM _ _ Hsubsumption) as Hincl. - apply basic_VLSM_incl; intro; intros. - - assumption. - - apply initial_message_is_protocol. revert HmX. - apply (VLSM_incl_protocol_message Hincl). - intro; intros. assumption. - - destruct Hv as [Hs [Hom [lX [sX [HlX [Heqs Hv]]]]]]. - apply (VLSM_incl_protocol_valid Hincl) in Hv. - exists lX, sX; intuition. - - apply H. + apply projection_induced_vlsm_incl. + - apply weak_induced_sub_projection_transition_consistency_Some. + - apply weak_induced_sub_projection_transition_consistency_Some. + - apply constraint_subsumption_incl. + assumption. Qed. End induced_sub_projection_subsumption. @@ -473,7 +490,7 @@ Qed. Lemma finite_trace_sub_projection_last_state (start : composite_state IM) (transitions : list (composite_transition_item IM)) - (Htr : finite_protocol_trace_from X start transitions) + (Htr : finite_valid_trace_from X start transitions) (lstx := finite_trace_last start transitions) (lstj := finite_trace_last (composite_state_sub_projection start) @@ -597,23 +614,23 @@ Definition state_sub_item_input_is_seeded_or_sub_previously_sent (s : composite_state IM) : Prop := forall is tr, - finite_protocol_trace_init_to Pre is s tr -> + finite_valid_trace_init_to Pre is s tr -> trace_sub_item_input_is_seeded_or_sub_previously_sent tr. -Lemma finite_protocol_trace_sub_projection +Lemma finite_valid_trace_sub_projection (s : composite_state IM) (tr : list (composite_transition_item IM)) (Hmsg : trace_sub_item_input_is_seeded_or_sub_previously_sent tr) - (Htr : finite_protocol_trace X s tr) - : finite_protocol_trace Xj (composite_state_sub_projection s) (finite_trace_sub_projection tr). + (Htr : finite_valid_trace X s tr) + : finite_valid_trace Xj (composite_state_sub_projection s) (finite_trace_sub_projection tr). Proof. destruct Htr as [Htr His]. apply (composite_initial_state_sub_projection s) in His. split; [|assumption]. - apply (initial_is_protocol Xj) in His as Hisp. + apply (initial_state_is_valid Xj) in His as Hisp. induction tr using rev_ind; simpl ; [constructor; assumption|]. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hx]. spec IHtr. { intros pre item suf m Heq Hin_m Hitem. @@ -626,17 +643,17 @@ Proof. } spec IHtr Htr. rewrite finite_trace_sub_projection_app. - apply finite_protocol_trace_from_app_iff. + apply finite_valid_trace_from_app_iff. split; [assumption|]. match goal with - |- finite_protocol_trace_from _ ?l _ => remember l as lst + |- finite_valid_trace_from _ ?l _ => remember l as lst end. - assert (Hlst : protocol_state_prop Xj lst). - { apply finite_ptrace_last_pstate in IHtr. subst. assumption. } + assert (Hlst : valid_state_prop Xj lst). + { apply finite_valid_trace_last_pstate in IHtr. subst. assumption. } simpl. unfold pre_VLSM_projection_transition_item_project, composite_label_sub_projection_option. case_decide; [|constructor; assumption]. - apply (finite_ptrace_singleton Xj). + apply (finite_valid_trace_singleton Xj). inversion Hx; subst. simpl in *. destruct Ht as [Hv Ht]. specialize (transition_sub_projection _ _ _ _ _ Ht H) @@ -646,8 +663,8 @@ Proof. as Hvj. rewrite <- (finite_trace_sub_projection_last_state s tr Htr) in Htj, Hvj. repeat split; [assumption | | assumption | | assumption]. - - destruct iom as [m|]; [|apply (option_protocol_message_None Xj)]. - apply (option_protocol_message_Some Xj). + - destruct iom as [m|]; [|apply (option_valid_message_None Xj)]. + apply (option_valid_message_Some Xj). clear -Hmsg m H IHtr tr. remember {| input := Some m |} as x. specialize (Hmsg tr x []). @@ -660,8 +677,8 @@ Proof. specialize (Hmsg m eq_refl eq_refl). spec Hmsg. { subst x. assumption. } destruct Hmsg as [Hseed | [item [Hitem [Hout Hsub_item]]]] - ; [apply (initial_message_is_protocol Xj); right; assumption|]. - apply (protocol_trace_output_is_protocol Xj _ _ IHtr). + ; [apply (initial_message_is_valid Xj); right; assumption|]. + apply (valid_trace_output_is_valid Xj _ _ IHtr). apply Exists_exists. specialize (@pre_VLSM_projection_transition_item_project_is_Some _ (composite_type IM) _ @@ -693,16 +710,16 @@ Proof. left. remember (finite_trace_last (composite_state_sub_projection _) _) as lst. specialize (proper_sent Sub_Free lst (HasBeenSentCapability := Sub_Free_HasBeenSentCapability)) as Hproper. - assert (Hlstp : protocol_state_prop (pre_loaded_with_all_messages_vlsm Sub_Free) lst). - { revert Hlst. apply VLSM_incl_protocol_state. apply Xj_incl_Pre_Sub_Free. } + assert (Hlstp : valid_state_prop (pre_loaded_with_all_messages_vlsm Sub_Free) lst). + { revert Hlst. apply VLSM_incl_valid_state. apply Xj_incl_Pre_Sub_Free. } spec Hproper Hlstp. apply Hproper. apply has_been_sent_consistency; [apply Sub_Free_HasBeenSentCapability| assumption| ]. exists (composite_state_sub_projection s), (finite_trace_sub_projection tr). split. { split;[|assumption]. - apply (VLSM_incl_finite_protocol_trace_from_to Xj_incl_Pre_Sub_Free). - apply ptrace_add_last. + apply (VLSM_incl_finite_valid_trace_from_to Xj_incl_Pre_Sub_Free). + apply valid_trace_add_last. assumption. symmetry;assumption. } @@ -723,50 +740,50 @@ Proof. assumption. Qed. -Lemma protocol_state_sub_projection +Lemma valid_state_sub_projection (s : state) (Hs : state_sub_item_input_is_seeded_or_sub_previously_sent s) - (Hps : protocol_state_prop X s) - : protocol_state_prop Xj (composite_state_sub_projection s). + (Hps : valid_state_prop X s) + : valid_state_prop Xj (composite_state_sub_projection s). Proof. - apply protocol_state_has_trace in Hps. + apply valid_state_has_trace in Hps. destruct Hps as [is [tr Htr]]. - specialize (Hs _ _ (VLSM_incl_finite_protocol_trace_init_to X_incl_Pre _ _ _ Htr)). - apply ptrace_get_last in Htr as Hlst. - apply ptrace_forget_last in Htr. + specialize (Hs _ _ (VLSM_incl_finite_valid_trace_init_to X_incl_Pre _ _ _ Htr)). + apply valid_trace_get_last in Htr as Hlst. + apply valid_trace_forget_last in Htr. specialize (finite_trace_sub_projection_last_state _ _ (proj1 Htr)) as Hlst'. - apply (finite_protocol_trace_sub_projection _ _ Hs) in Htr as Hptr. - - destruct Hptr as [Hptr _]. apply finite_ptrace_last_pstate in Hptr. + apply (finite_valid_trace_sub_projection _ _ Hs) in Htr as Hptr. + - destruct Hptr as [Hptr _]. apply finite_valid_trace_last_pstate in Hptr. simpl in *. rewrite Hlst' in Hptr. subst. assumption. Qed. -Lemma finite_protocol_trace_from_sub_projection +Lemma finite_valid_trace_from_sub_projection (s : composite_state IM) (tr : list (composite_transition_item IM)) (lst := finite_trace_last s tr) (Hmsg : state_sub_item_input_is_seeded_or_sub_previously_sent lst) - (Htr : finite_protocol_trace_from X s tr) - : finite_protocol_trace_from Xj (composite_state_sub_projection s) (finite_trace_sub_projection tr). + (Htr : finite_valid_trace_from X s tr) + : finite_valid_trace_from Xj (composite_state_sub_projection s) (finite_trace_sub_projection tr). Proof. - apply finite_protocol_trace_from_complete_left in Htr. + apply finite_valid_trace_from_complete_left in Htr. destruct Htr as [is [pre [Htr Hs]]]. assert (Hpre := proj1 Htr). - apply finite_protocol_trace_from_app_iff in Hpre. + apply finite_valid_trace_from_app_iff in Hpre. destruct Hpre as [Hpre _]. specialize (finite_trace_sub_projection_last_state _ _ Hpre) as Hpre_lst. - apply finite_protocol_trace_sub_projection in Htr. + apply finite_valid_trace_sub_projection in Htr. - destruct Htr as [Htr His]. rewrite finite_trace_sub_projection_app in Htr. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [_ Htr]. subst s. simpl in *. rewrite Hpre_lst in Htr. assumption. - specialize (Hmsg is (pre ++ tr)). apply Hmsg. - apply (VLSM_incl_finite_protocol_trace_init_to X_incl_Pre). - apply ptrace_add_last. + apply (VLSM_incl_finite_valid_trace_init_to X_incl_Pre). + apply valid_trace_add_last. assumption. rewrite finite_trace_last_app. unfold lst. subst. reflexivity. @@ -873,14 +890,14 @@ End sub_composition. (** ** Lifting a trace from a sub-composition to the full composition -In this section we first show that given a protocol state for a composition of +In this section we first show that given a valid state for a composition of all nodes we can reset some of its state-components to initial states for those -components without losing the protocol state property. +components without losing the valid state property. The set of nodes for which the reset operation will happen is <>. We then show that a similar result holds for replacing the equivocator -components with the components corresponding to any protocol state of the +components with the components corresponding to any valid state of the composition of just the equivocators. We proving those results for compositions pre-loaded with all messages @@ -902,7 +919,7 @@ Context (SubFree : VLSM message := free_composite_vlsm equivocating_IM) (PreSubFree := pre_loaded_with_all_messages_vlsm SubFree) (base_s : composite_state IM) - (Hbase_s : protocol_state_prop PreFree base_s) + (Hbase_s : valid_state_prop PreFree base_s) . (** A partial label projection function which only keeps non-equivocating transitions. @@ -993,11 +1010,11 @@ Proof. Qed. (** -Given any protocol trace for the composition of all nodes and an initial state +Given any valid trace for the composition of all nodes and an initial state for the composition of just the equivocators, the trace obtained by resetting the components corresponding to the equivocators to those of the given initial state and removing the transitions corresponding to the equivocators is -still a protocol trace. +still a valid trace. *) Lemma remove_equivocating_transitions_preloaded_projection eqv_is (Heqv_is : composite_initial_state_prop (sub_IM IM equivocators) eqv_is) @@ -1014,11 +1031,11 @@ Qed. Lemma preloaded_lift_sub_state_to_initial_state : weak_full_projection_initial_state_preservation PreSubFree PreFree (lift_sub_state_to IM equivocators base_s). Proof. - apply protocol_state_has_trace in Hbase_s as Htr. + apply valid_state_has_trace in Hbase_s as Htr. destruct Htr as [is [tr Htr]]. intros eqv_is Heqv_is. - apply (VLSM_projection_finite_protocol_trace_init_to (remove_equivocating_transitions_preloaded_projection _ Heqv_is)) in Htr. - apply ptrace_last_pstate in Htr. assumption. + apply (VLSM_projection_finite_valid_trace_init_to (remove_equivocating_transitions_preloaded_projection _ Heqv_is)) in Htr. + apply valid_trace_last_pstate in Htr. assumption. Qed. Lemma lift_sub_to_valid l s om @@ -1058,10 +1075,10 @@ Proof. Qed. (** -Given any protocol state for the composition of all nodes and a protocol trace +Given any valid state for the composition of all nodes and a valid trace for the composition of just the equivocators, the trace obtained by completing the state-components from the trace with the components from the given -protocol state is a protocol trace for the composition of all nodes. +valid state is a valid trace for the composition of all nodes. **) Lemma PreSubFree_PreFree_weak_full_projection : VLSM_weak_full_projection PreSubFree PreFree (lift_sub_label IM equivocators) (lift_sub_state_to IM equivocators base_s). @@ -1071,7 +1088,96 @@ Proof. apply lift_sub_to_valid. apply Hv. - apply lift_sub_to_transition. apply H. - apply preloaded_lift_sub_state_to_initial_state; assumption. - - apply any_message_is_protocol_in_preloaded. + - apply any_message_is_valid_in_preloaded. +Qed. + +(** If the composition constraint only depends on the projection sub-state, +then valid traces of the [induced_sub_projection] can be lifted to valid traces +of the constrained composition. +*) +Lemma induced_sub_projection_lift + (constraint : composite_label IM -> composite_state IM * option message -> Prop) + (Hconstraint_consistency : + forall s1 s2, + composite_state_sub_projection IM equivocators s1 = composite_state_sub_projection IM equivocators s2 -> + forall l om, constraint l (s1, om) -> constraint l (s2, om) + ) + : VLSM_full_projection + (induced_sub_projection IM equivocators constraint) + (composite_vlsm IM constraint) + (lift_sub_label IM equivocators) + (lift_sub_state IM equivocators). +Proof. + apply basic_VLSM_full_projection; intro; intros. + - destruct Hv as [_ [_ [[i li] [sX [Heql [Heqs [HsX [Hom [Hv Hc]]]]]]]]]. + cbn in Hv, Hc. + unfold composite_label_sub_projection_option in Heql. + simpl in Heql. + case_decide; [|congruence]. + inversion Heql. subst l. clear Heql. + cbn. unfold constrained_composite_valid. cbn. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi := H). + subst. + split; [assumption|]. + revert Hc. + apply Hconstraint_consistency. + symmetry. + apply composite_state_sub_projection_lift_to. + - apply proj2 in H. revert H. cbn. + destruct (vtransition _ _ _) as (si', _om'). + inversion 1. subst. clear H. + f_equal. + extensionality i. + destruct l as (sub_j, lj). + destruct_dec_sig sub_j j Hj Heqsub_j. + subst. + simpl. + destruct (decide (i = j)). + + subst. rewrite state_update_eq. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi := Hj). + unfold composite_state_sub_projection. + simpl. + rewrite state_update_eq. + reflexivity. + + rewrite state_update_neq by congruence. + destruct (decide (i ∈ equivocators)). + * unfold lift_sub_state. + rewrite !lift_sub_state_to_eq with (Hi := e). + unfold composite_state_sub_projection. simpl. + rewrite state_update_neq by congruence. + rewrite lift_sub_state_to_eq with (Hi := e). + reflexivity. + * unfold lift_sub_state, lift_sub_state_to. + destruct (decide _); [contradiction|reflexivity]. + - apply (lift_sub_state_initial IM). + destruct H as [sX [<- HsX]]. + intro sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. apply HsX. + - assumption. +Qed. + +(** A specialization of [basic_projection_induces_friendliness] for +[induced_sub_projection]s. +*) +Lemma induced_sub_projection_friendliness + (constraint : composite_label IM -> composite_state IM * option message -> Prop) + (Hlift_proj : VLSM_full_projection + (induced_sub_projection IM equivocators constraint) + (composite_vlsm IM constraint) + (lift_sub_label IM equivocators) + (lift_sub_state IM equivocators)) + : projection_friendly_prop (induced_sub_projection_is_projection IM equivocators constraint). +Proof. + eapply (basic_projection_induces_friendliness (composite_vlsm IM constraint)). + assumption. + Unshelve. + - apply induced_sub_projection_transition_consistency_None. + - apply composite_label_sub_projection_option_lift. + - apply composite_state_sub_projection_lift. + - apply induced_sub_projection_transition_consistency_Some. Qed. End lift_sub_state_to_preloaded. @@ -1260,19 +1366,19 @@ Proof. specialize (PreSubFree_PreFree_weak_full_projection IM indices (proj1_sig (composite_s0 IM))) as Hproj. spec Hproj. - { apply initial_is_protocol. destruct (composite_s0 IM). assumption. } + { apply initial_state_is_valid. destruct (composite_s0 IM). assumption. } apply - (VLSM_incl_protocol_transition + (VLSM_incl_input_valid_transition (pre_loaded_vlsm_incl_pre_loaded_with_all_messages (free_composite_vlsm sub_IM) P)) in Ht. - apply (VLSM_weak_full_projection_protocol_transition Hproj) in Ht. + apply (VLSM_weak_full_projection_input_valid_transition Hproj) in Ht. clear Hproj. specialize (ProjectionTraces.preloaded_component_projection IM i) as Hproj. remember (lift_sub_state_to _ _ _ s) as sX. remember (lift_sub_state_to _ _ _ s') as sX'. remember (lift_sub_label _ _ _) as lX. - specialize (VLSM_projection_protocol_transition Hproj lX li) as Hproj_t. + specialize (VLSM_projection_input_valid_transition Hproj lX li) as Hproj_t. subst lX. unfold lift_sub_label in Hproj_t. simpl in Hproj_t. spec Hproj_t. @@ -1363,14 +1469,14 @@ Context . Lemma sub_IM_has_been_sent_iff_by_sender s - (Hs : protocol_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm sub_IM)) s) + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm (free_composite_vlsm sub_IM)) s) m v (Hsender : sender m = Some v) (Hv : A v ∈ indices) : composite_has_been_sent sub_IM sub_IM_Hbs s m -> @has_been_sent _ _ (sub_IM_Hbs (dexist (A v) Hv)) (s (dexist (A v) Hv)) m. Proof. - apply protocol_state_has_trace in Hs as Htr. + apply valid_state_has_trace in Hs as Htr. destruct Htr as [is [tr Htr]]. specialize (has_been_sent_iff_by_sender sub_IM @@ -1392,7 +1498,7 @@ Proof. by (subst; apply (sub_IM_state_pi IM)). apply has_been_sent_irrelevance. subst. - apply (preloaded_protocol_state_projection sub_IM (dec_exist (sub_index_prop indices) (A v) Hv)) + apply (preloaded_valid_state_projection sub_IM (dec_exist (sub_index_prop indices) (A v) Hv)) in Hs. assumption. Qed. @@ -1482,13 +1588,13 @@ Proof. destruct Hv as [lX [sX [_ [Heqs [_ [Hm [_ Hc]]]]]]]. cbn in Hc |- *. specialize - (composite_no_initial_protocol_messages_have_sender IM A sender + (composite_no_initial_valid_messages_have_sender IM A sender can_emit_signed no_initial_messages_in_IM _ _ Hm) as Hhas_sender. destruct (sender m) as [v|] eqn:Hsender; [|congruence]. clear Hhas_sender. simpl in Hc. - apply (can_emit_protocol_iff (composite_vlsm IM sub_IM_not_equivocating_constraint) m) + apply (emitted_messages_are_valid_iff (composite_vlsm IM sub_IM_not_equivocating_constraint) m) in Hm as [[i [[im Him] Heqm]] | Hemitted]. - exfalso. clear -no_initial_messages_in_IM Him. elim (no_initial_messages_in_IM i im); assumption. @@ -1679,9 +1785,9 @@ Lemma lift_sub_free_full_projection Proof. constructor. intros sX trX HtrX. - apply (VLSM_eq_finite_protocol_trace (vlsm_is_pre_loaded_with_False Free)), - (VLSM_full_projection_finite_protocol_trace (lift_sub_free_preloaded_with_full_projection _)), - (VLSM_eq_finite_protocol_trace (vlsm_is_pre_loaded_with_False SubFree)). + apply (VLSM_eq_finite_valid_trace (vlsm_is_pre_loaded_with_False Free)), + (VLSM_full_projection_finite_valid_trace (lift_sub_free_preloaded_with_full_projection _)), + (VLSM_eq_finite_valid_trace (vlsm_is_pre_loaded_with_False SubFree)). assumption. Qed. @@ -1691,9 +1797,9 @@ Lemma lift_sub_preloaded_free_full_projection Proof. constructor. intros sX trX HtrX. - apply (VLSM_eq_finite_protocol_trace (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True Free)), - (VLSM_full_projection_finite_protocol_trace (lift_sub_free_preloaded_with_full_projection _)), - (VLSM_eq_finite_protocol_trace (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True SubFree)). + apply (VLSM_eq_finite_valid_trace (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True Free)), + (VLSM_full_projection_finite_valid_trace (lift_sub_free_preloaded_with_full_projection _)), + (VLSM_eq_finite_valid_trace (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True SubFree)). assumption. Qed. diff --git a/theories/VLSM/Core/VLSM.v b/theories/VLSM/Core/VLSM.v index 59576965..8c9aa438 100644 --- a/theories/VLSM/Core/VLSM.v +++ b/theories/VLSM/Core/VLSM.v @@ -86,6 +86,12 @@ function and a [valid]ity condition. ; valid := @valid _ _ _ vlsm |}. + (** For technical reasons, e.g., the need to easily talk about VLSMs over + the same set of messages and about VLSMs of the same type (over the same + set of messages, labels and states), the VLSM definition is split into + three parts, [VLSM_Type], [VLSM_Sign], and [VLSMClass], which are + packaged together by the following definition. + *) Definition VLSM (message : Type) := sigT (fun T : VLSMType message => sigT (fun S : VLSMSign T => VLSMClass S)). @@ -187,7 +193,7 @@ In Coq, we can define these objects (which we name [transition_item]s) as consis : Prop := forall m, trace_received_not_sent_before_or_after tr m -> P m. - (** 'proto_run's are used for an alternative definition of 'protocol_prop' which + (** [proto_run] is used for an alternative definition of [valid_state_message_prop], which takes into account transitions. See 'vlsm_run_prop'. *) Record proto_run : Type := mk_proto_run @@ -472,164 +478,170 @@ Existing Instance TypeX. Existing Instance SignX. Existing Instance MachineX. -(** *** Protocol states and messages +(** *** Valid states and messages -We further characterize certain objects as being _protocol_, which means they can +We further characterize certain objects as being _valid_, which means they can be witnessed or experienced during executions of the protocol. For example, -a message is a [protocol_message] if there exists an execution of the protocol +a message is a [valid_message] if there exists an execution of the protocol in which it is produced. -We choose here to define protocol states and messages together as the -[protocol_prop] property, inductively defined over the +We choose here to define valid states and messages together as the +[valid_state_message_prop]erty, inductively defined over the [state * option message] product type, as this definition avoids the need of using a mutually recursive definition. The inductive definition has three cases: -- if <> is a [state] with the [initial_state_prop]erty, then <<(s, None)>> has the [protocol_prop]erty; -- if <> is a <> with the [initial_message_prop]erty, then <<(>>[s0, Some]<< m)>> has the [protocol_prop]erty; -- for all [state]s <>, [option]al <> <>, - and [label] <>: +- if <> is a [state] with the [initial_state_prop]erty, then <<(s, None)>> + has the [valid_state_message_prop]erty; +- if <> is a <> with the [initial_message_prop]erty, + then <<(>>[s0, Some]<< m)>> has the [valid_state_message_prop]erty; +- for all [state]s <>, [option]al <> <>, and [label] <>: - if there is an (optional) <> <<_om>> such that <<(s, _om)>> has the [protocol_prop]erty; + - if there is an (optional) <> <<_om>> such that <<(s, _om)>> + has the [valid_state_message_prop]erty; - and if there is a [state] <<_s>> such that <<(_s, om)>> has the [protocol_prop]erty; + - and if there is a [state] <<_s>> such that <<(_s, om)>> has the + [valid_state_message_prop]erty; - and if <> [valid] <<(s, om)>>, + - and if <> [valid] <<(s, om)>>, - then [transition] <> has the [protocol_prop]erty. + - then [transition] <> has the [valid_state_message_prop]erty. *) - Inductive protocol_prop : state -> option message -> Prop := - | protocol_initial + Inductive valid_state_message_prop : state -> option message -> Prop := + | valid_initial_state_message (s : state) (Hs : initial_state_prop s) (om : option message) (Hom : option_initial_message_prop om) - : protocol_prop s om - | protocol_generated + : valid_state_message_prop s om + | valid_generated_state_message (s : state) (_om : option message) - (Hps : protocol_prop s _om) + (Hps : valid_state_message_prop s _om) (_s : state) (om : option message) - (Hpm : protocol_prop _s om) + (Hpm : valid_state_message_prop _s om) (l : label) (Hv : valid l (s, om)) s' om' (Ht : transition l (s, om) = (s', om')) - : protocol_prop s' om'. + : valid_state_message_prop s' om'. - Definition protocol_initial_state + Definition valid_initial_state [s:state] (Hs: initial_state_prop s) - : protocol_prop s None - := protocol_initial s Hs None I. + : valid_state_message_prop s None + := valid_initial_state_message s Hs None I. (** -The [protocol_state_prop]erty and the [protocol_message_prop]erty are now +The [valid_state_prop]erty and the [valid_message_prop]erty are now definable as simple projections of the above definition. Moreover, we use these derived properties to define the corresponding -dependent types [protocol_state] and [protocol_message]. +dependent types [valid_state] and [valid_message]. *) - Definition protocol_state_prop (s : state) := - exists om : option message, protocol_prop s om. + Definition valid_state_prop (s : state) := + exists om : option message, valid_state_message_prop s om. - Definition protocol_message_prop (m : message) := - exists s : state, protocol_prop s (Some m). + Definition valid_message_prop (m : message) := + exists s : state, valid_state_message_prop s (Some m). - Definition protocol_state : Type := - { s : state | protocol_state_prop s }. + Definition valid_state : Type := + { s : state | valid_state_prop s }. - Definition protocol_message : Type := - { m : message | protocol_message_prop m }. + Definition valid_message : Type := + { m : message | valid_message_prop m }. - Lemma initial_is_protocol + Lemma initial_state_is_valid (s : state) (Hinitial : initial_state_prop s) : - protocol_state_prop s. + valid_state_prop s. Proof. exists None. - apply protocol_initial. + apply valid_initial_state_message. assumption. exact I. Qed. - Lemma initial_message_is_protocol + Lemma initial_message_is_valid (m : message) (Hinitial : initial_message_prop m) : - protocol_message_prop m. + valid_message_prop m. Proof. exists (proj1_sig (vs0 X)). - apply protocol_initial. + apply valid_initial_state_message. apply proj2_sig. assumption. Qed. (** -As often times we work with optional protocol messages, it is convenient -to define a protocol message property for optional messages: +As often times we work with optional valid messages, it is convenient +to define a valid message property for optional messages: *) - Definition option_protocol_message_prop (om : option message) := - exists s : state, protocol_prop s om. + Definition option_valid_message_prop (om : option message) := + exists s : state, valid_state_message_prop s om. - Lemma option_protocol_message_None - : option_protocol_message_prop None. + Lemma option_valid_message_None + : option_valid_message_prop None. Proof. exists (proj1_sig (vs0 X)). - apply protocol_initial. + apply valid_initial_state_message. apply proj2_sig. exact I. Qed. - Lemma option_protocol_message_Some + Lemma option_valid_message_Some (m : message) - (Hpm : protocol_message_prop m) - : option_protocol_message_prop (Some m). + (Hpm : valid_message_prop m) + : option_valid_message_prop (Some m). Proof. destruct Hpm as [s Hpm]. exists s. assumption. Qed. - Lemma option_initial_message_is_protocol + Lemma option_initial_message_is_valid (om : option message) (Hinitial : option_initial_message_prop om) : - option_protocol_message_prop om. + option_valid_message_prop om. Proof. destruct om; - [apply option_protocol_message_Some - |apply option_protocol_message_None]. - apply initial_message_is_protocol;assumption. + [apply option_valid_message_Some + |apply option_valid_message_None]. + apply initial_message_is_valid;assumption. Qed. -(** *** Protocol validity and protocol transitions +(** *** Input validity and input valid transitions -To achieve this, it is useful to further define _protocol_ validity and -_protocol_ transitions: -*) +To specify that a particular (input of a) transition can actually be +encountered as part of a protocol execution, we define the notions of +[input_valid]ity and [input_valid_transition]. - Definition protocol_valid +Input validity requires that the [valid] predicate holds for the +given inputs and that they have a [valid_state] and a [valid_message]. +*) + Definition input_valid (l : label) (som : state * option message) : Prop := let (s, om) := som in - protocol_state_prop s - /\ option_protocol_message_prop om + valid_state_prop s + /\ option_valid_message_prop om /\ valid l (s,om). - - Definition protocol_transition +(** Input valid transitions are transitions with [input_valid] inputs. *) + Definition input_valid_transition (l : label) (som : state * option message) (som' : state * option message) := - protocol_valid l som + input_valid l som /\ transition l som = som'. - Definition protocol_transition_preserving + Definition input_valid_transition_preserving (R : state -> state -> Prop) : Prop := @@ -637,46 +649,46 @@ _protocol_ transitions: (s1 s2 : state) (l : label) (om1 om2 : option message) - (Hprotocol: protocol_transition l (s1, om1) (s2, om2)), + (Hvalid_transition: input_valid_transition l (s1, om1) (s2, om2)), R s1 s2. (** Next three lemmas show the two definitions above are strongly related. *) - Lemma protocol_transition_valid + Lemma input_valid_transition_valid (l : label) (som : state * option message) (som' : state * option message) - (Ht : protocol_transition l som som') - : protocol_valid l som. + (Ht : input_valid_transition l som som') + : input_valid l som. Proof. destruct Ht as [Hpv Ht]. assumption. Qed. - Lemma protocol_valid_transition + Lemma input_valid_can_transition (l : label) (som : state * option message) - (Hv : protocol_valid l som) + (Hv : input_valid l som) : exists (som' : state * option message), - protocol_transition l som som'. + input_valid_transition l som som'. Proof. exists (transition l som). repeat split; assumption. Qed. - Lemma protocol_valid_transition_iff + Lemma input_valid_transition_iff (l : label) (som : state * option message) - : protocol_valid l som + : input_valid l som <-> exists (som' : state * option message), - protocol_transition l som som'. + input_valid_transition l som som'. Proof. split. - - apply protocol_valid_transition. - - intros [som' Hpt]. - apply protocol_transition_valid with som'. + - apply input_valid_can_transition. + - intros [som' Hivt]. + apply input_valid_transition_valid with som'. assumption. Qed. @@ -686,156 +698,140 @@ The next couple of lemmas relate the two definitions above with pre-existing concepts. *) - Lemma protocol_generated_valid - {l : label} - {s : state} - {_om : option message} - {_s : state} - {om : option message} - (Hps : protocol_prop s _om) - (Hpm : protocol_prop _s om) - (Hv : valid l (s, om)) - : protocol_valid l (s, om). - Proof. - repeat split; try assumption. - - exists _om. assumption. - - exists _s. assumption. - Qed. - - Lemma protocol_transition_origin + Lemma input_valid_transition_origin {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s',om')) - : protocol_state_prop s. + (Ht : input_valid_transition l (s, om) (s',om')) + : valid_state_prop s. Proof. destruct Ht as [[[_om Hp] _] _]. exists _om. assumption. Qed. - Lemma protocol_transition_destination + Lemma input_valid_transition_destination {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) - : protocol_state_prop s'. + (Ht : input_valid_transition l (s, om) (s', om')) + : valid_state_prop s'. Proof. exists om'. destruct Ht as [[[_om Hs] [[_s Hom] Hv]] Ht]. - apply protocol_generated with s _om _s om l; assumption. + apply valid_generated_state_message with s _om _s om l; assumption. Qed. - Lemma protocol_transition_in + Lemma input_valid_transition_in {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) - : option_protocol_message_prop om. + (Ht : input_valid_transition l (s, om) (s', om')) + : option_valid_message_prop om. Proof. destruct Ht as [[_ [[_s Hom] _]] _]. exists _s. assumption. Qed. - Lemma protocol_prop_transition_out + Lemma input_valid_transition_outputs_valid_state_message {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) - : protocol_prop s' om'. + (Ht : input_valid_transition l (s, om) (s', om')) + : valid_state_message_prop s' om'. Proof. destruct Ht as [[[_om Hps] [[_s Hpm] Hv]] Ht]. - apply protocol_generated with s _om _s om l; assumption. + apply valid_generated_state_message with s _om _s om l; assumption. Qed. - Lemma protocol_transition_out + Lemma input_valid_transition_out {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) - : option_protocol_message_prop om'. + (Ht : input_valid_transition l (s, om) (s', om')) + : option_valid_message_prop om'. Proof. - apply protocol_prop_transition_out in Ht. + apply input_valid_transition_outputs_valid_state_message in Ht. exists s'. assumption. Qed. - Lemma protocol_transition_is_valid + Lemma input_valid_transition_is_valid {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) + (Ht : input_valid_transition l (s, om) (s', om')) : valid l (s, om). Proof. destruct Ht as [[_ [_ Hv]] _]. assumption. Qed. - Lemma protocol_transition_transition + Lemma input_valid_transition_transition {l : label} {s s' : state} {om om' : option message} - (Ht : protocol_transition l (s, om) (s', om')) + (Ht : input_valid_transition l (s, om) (s', om')) : transition l (s, om) = (s', om'). Proof. destruct Ht as [_ Ht]. assumption. Qed. - Lemma protocol_prop_valid_out + Lemma input_valid_state_message_outputs (l : label) (s : state) (om : option message) - (Hv : protocol_valid l (s, om)) + (Hv : input_valid l (s, om)) s' om' (Ht : transition l (s, om) = (s', om')) - : protocol_prop s' om'. + : valid_state_message_prop s' om'. Proof. destruct Hv as [[_om Hs] [[_s Hom] Hv]]. - apply protocol_generated with s _om _s om l; assumption. + apply valid_generated_state_message with s _om _s om l; assumption. Qed. (** For VLSMs initialized with many initial messages such as the [composite_vlsm_constrained_projection] or the [pre_loaded_with_all_messages_vlsm], the question of whether a [VLSM] [can_emit] a message <> becomes more - useful than that whether <> is a [protocol_message]. + useful than that whether <> is a [valid_message]. *) - Definition option_protocol_generated_prop + Definition option_can_produce (s : state) (om : option message) := exists (som : state * option message) (l : label), - protocol_transition l som (s, om). + input_valid_transition l som (s, om). - Definition protocol_generated_prop + Definition can_produce (s : state) (m : message) - := option_protocol_generated_prop s (Some m). + := option_can_produce s (Some m). - (** Of course, if a VLSM [can_emit] <<(s,m)>>, then <<(s,m)>> is protocol. + (** Of course, if a VLSM [can_emit] <<(s,m)>>, then <<(s,m)>> is valid. *) - Lemma option_protocol_generated_prop_protocol + Lemma option_can_produce_valid (s : state) (om : option message) - (Hm : option_protocol_generated_prop s om) - : protocol_prop s om . + (Hm : option_can_produce s om) + : valid_state_message_prop s om . Proof. destruct Hm as [(s0, om0) [l [[[_om0 Hs0] [[_s0 Hom0] Hv]] Ht]]]. - apply protocol_generated with s0 _om0 _s0 om0 l; assumption. + apply valid_generated_state_message with s0 _om0 _s0 om0 l; assumption. Qed. - Definition protocol_generated_prop_protocol + Definition can_produce_valid (s : state) (m : message) - (Hm : protocol_generated_prop s m) - : protocol_prop s (Some m) - := option_protocol_generated_prop_protocol s (Some m) Hm. + (Hm : can_produce s m) + : valid_state_message_prop s (Some m) + := option_can_produce_valid s (Some m) Hm. - Lemma option_protocol_generated_prop_protocol_iff + Lemma option_can_produce_valid_iff (s : state) (om : option message) - : protocol_prop s om <-> - option_protocol_generated_prop s om \/ initial_state_prop s /\ option_initial_message_prop om. + : valid_state_message_prop s om <-> + option_can_produce s om \/ initial_state_prop s /\ option_initial_message_prop om. Proof. split. - intros Hm; inversion Hm; subst. @@ -845,16 +841,16 @@ pre-existing concepts. repeat split; [..|assumption|assumption] ; eexists; [exact Hps|exact Hpm]. - intros [Hem | Him]. - + apply option_protocol_generated_prop_protocol. assumption. + + apply option_can_produce_valid. assumption. + constructor; apply Him. Qed. - Definition protocol_generated_prop_protocol_iff + Definition can_produce_valid_iff (s : state) (m : message) - : protocol_prop s (Some m) <-> - protocol_generated_prop s m \/ initial_state_prop s /\ initial_message_prop m - := option_protocol_generated_prop_protocol_iff s (Some m). + : valid_state_message_prop s (Some m) <-> + can_produce s m \/ initial_state_prop s /\ initial_message_prop m + := option_can_produce_valid_iff s (Some m). Definition can_emit (m : message) @@ -863,64 +859,64 @@ pre-existing concepts. (som : state * option message) (l : label) (s : state), - protocol_transition l som (s, Some m). + input_valid_transition l som (s, Some m). Lemma can_emit_iff (m : message) - : can_emit m <-> exists s, protocol_generated_prop s m. + : can_emit m <-> exists s, can_produce s m. Proof. split. - intros [som [l [s Ht]]]. exists s, som, l. assumption. - intros [s [som [l Ht]]]. exists som, l, s. assumption. Qed. - (** If a VLSM [can_emit] a message <>, then <> is protocol. + (** If a VLSM [can_emit] a message <>, then <> is valid. *) - Lemma can_emit_protocol + Lemma emitted_messages_are_valid (m : message) (Hm : can_emit m) - : protocol_message_prop m . + : valid_message_prop m . Proof. apply can_emit_iff in Hm. destruct Hm as [s Hm]. - apply protocol_generated_prop_protocol in Hm. + apply can_produce_valid in Hm. exists s. assumption. Qed. - (** A characterization of protocol messages in terms of [can_emit] + (** A characterization of valid messages in terms of [can_emit] *) - Lemma can_emit_protocol_iff + Lemma emitted_messages_are_valid_iff (m : message) - : protocol_message_prop m <-> initial_message_prop m \/ can_emit m. + : valid_message_prop m <-> initial_message_prop m \/ can_emit m. Proof. split. - intros [s Hm]. - apply protocol_generated_prop_protocol_iff in Hm as [Hgen | [_ Him]]. + apply can_produce_valid_iff in Hm as [Hgen | [_ Him]]. + right. apply can_emit_iff. exists s. assumption. + left. assumption. - intros [Him | Hm]. - + apply initial_message_is_protocol. assumption. - + apply can_emit_protocol. assumption. + + apply initial_message_is_valid. assumption. + + apply emitted_messages_are_valid. assumption. Qed. -(** *** Protocol state and protocol message characterization +(** *** valid state and valid message characterization The definition and results below show that the mutually-recursive definitions -for [protocol_state]s and [protocol_message]s can be derived from the +for [valid_state]s and [valid_message]s can be derived from the prior definitions. -The results below offers equivalent characterizations for [protocol_state]s -and [protocol_message]s, similar to their recursive definition. +The results below offers equivalent characterizations for [valid_state]s +and [valid_message]s, similar to their recursive definition. *) - Lemma protocol_state_prop_iff : + Lemma valid_state_prop_iff : forall s' : state, - protocol_state_prop s' + valid_state_prop s' <-> (exists is : initial_state, s' = proj1_sig is) \/ exists (l : label) (som : state * option message) (om' : option message), - protocol_transition l som (s', om'). + input_valid_transition l som (s', om'). Proof. intros; split. - intro Hps'. destruct Hps' as [om' Hs]. @@ -931,49 +927,43 @@ and [protocol_message]s, similar to their recursive definition. + exists _om. assumption. + exists _s. assumption. - intros [[[s His] Heq] | [l [[s om] [om' [[[_om Hps] [[_s Hpm] Hv]] Ht]]]]]; subst. - + exists None. apply protocol_initial; [assumption | exact I]. - + exists om'. apply protocol_generated with s _om _s om l; assumption. + + exists None. apply valid_initial_state_message; [assumption | exact I]. + + exists om'. apply valid_generated_state_message with s _om _s om l; assumption. Qed. - (** A specialized induction principle for [protocol_state_prop]. + (** A specialized induction principle for [valid_state_prop]. - Compared to opening the existential and using [protocol_prop_ind], - this avoids the redundancy of the [protocol_initial_state] case - needing a proof for any [initial_state] and then [protocol_initial_message] - asking for a proof specifically for [s0], and also avoids the - little trouble of needing <> or <> to - handle the pair in the index in [protocol_prop (s,om)]. + Compared to opening the existential and using [valid_state_message_prop_ind], + this expresses the inductive assumptions simply in terms of + [input_valid_transition], rather than than having everything exploded + as [valid_state_message_prop] assumptions over witnesses <<_s>> + and <<_om>>, and a spurious inductive assumption <

>. *) - Lemma protocol_state_prop_ind + Lemma valid_state_prop_ind (P : state -> Prop) (IHinit : forall (s : state) (Hs : initial_state_prop s), P s) (IHgen : forall (s' : state) (l: label) (om om' : option message) (s : state) - (Ht : protocol_transition l (s, om) (s', om')) (Hs : P s), + (Ht : input_valid_transition l (s, om) (s', om')) (Hs : P s), P s' ) - : forall (s : state) (Hs : protocol_state_prop s), P s. + : forall (s : state) (Hs : valid_state_prop s), P s. Proof. intros. destruct Hs as [om Hs]. induction Hs. - apply IHinit. assumption. - - specialize (IHgen s' l0 om om' s). - apply IHgen; try assumption. - repeat split; try assumption. - + exists _om. assumption. - + exists _s. assumption. + - apply (IHgen s' l0 om om' s);firstorder. Qed. + (* valid message characterization - similar to the definition in the report. *) - (* Protocol message characterization - similar to the definition in the report. *) - - Lemma protocol_message_prop_iff : + Lemma valid_message_prop_iff : forall m' : message, - protocol_message_prop m' + valid_message_prop m' <-> (exists im : initial_message, m' = proj1_sig im) \/ exists (l : label) (som : state * option message) (s' : state), - protocol_transition l som (s', Some m'). + input_valid_transition l som (s', Some m'). Proof. intros; split. - intros [s' Hpm']. @@ -982,9 +972,9 @@ and [protocol_message]s, similar to their recursive definition. + right. exists l0. exists (s, om). exists s'. firstorder. - intros [[[s His] Heq] | [l [[s om] [s' [[[_om Hps] [[_s Hpm] Hv]] Ht]]]]]; subst. - + apply initial_message_is_protocol. assumption. + + apply initial_message_is_valid. assumption. + exists s'. - apply protocol_generated with s _om _s om l; assumption. + apply valid_generated_state_message with s _om _s om l; assumption. Qed. (** ** Trace Properties @@ -999,17 +989,17 @@ We will now split our groundwork for defining traces into the finite case and the infinite case. *) -(** *** Finite [protocol_trace]s +(** *** Finite [valid_trace]s -A [finite_protocol_trace_from] a [state] <> is a pair <<(start, steps)>> where <> +A [finite_valid_trace_from] a [state] <> is a pair <<(start, steps)>> where <> is a list of [transition_item]s, and is inductively defined by: -- <<(s, [])>> is a [finite_protocol_trace_from] <> -- if there is a [protocol_transition] <> +- <<(s, [])>> is a [finite_valid_trace_from] <> if <> is valid +- if there is an [input_valid_transition] <> - and if <<(s,steps)>> is a [protocol_trace_from] <> + and if <<(s,steps)>> is a [valid_trace_from] <> then <<(s', ({| l := l; input := iom; destination := s; output := oom |} :: steps)>> - is a [protocol_transition_from] <>. + is a [finite_valid_trace_from] <>. Note that the definition is given such that it extends an existing trace by adding a transition to its front. @@ -1017,34 +1007,34 @@ The reason for this choice is to have this definition be similar to the one for infinite traces, which can only be extended at the front. *) - Inductive finite_protocol_trace_from : state -> list transition_item -> Prop := - | finite_ptrace_empty : forall (s : state) - (Hs : protocol_state_prop s), - finite_protocol_trace_from s [] - | finite_ptrace_extend : forall (s : state) (tl : list transition_item) - (Htl : finite_protocol_trace_from s tl) + Inductive finite_valid_trace_from : state -> list transition_item -> Prop := + | finite_valid_trace_from_empty : forall (s : state) + (Hs : valid_state_prop s), + finite_valid_trace_from s [] + | finite_valid_trace_from_extend : forall (s : state) (tl : list transition_item) + (Htl : finite_valid_trace_from s tl) (s' : state) (iom oom : option message) (l : label) - (Ht : protocol_transition l (s', iom) (s, oom)), - finite_protocol_trace_from s' ({| l := l; input := iom; destination := s; output := oom |} :: tl). + (Ht : input_valid_transition l (s', iom) (s, oom)), + finite_valid_trace_from s' ({| l := l; input := iom; destination := s; output := oom |} :: tl). - Definition finite_ptrace_singleton : + Definition finite_valid_trace_singleton : forall {l : label} {s s': state} {iom oom : option message}, - protocol_transition l (s, iom) (s', oom) -> - finite_protocol_trace_from s ({| l := l; input := iom; destination := s'; output := oom |} :: []) + input_valid_transition l (s, iom) (s', oom) -> + finite_valid_trace_from s ({| l := l; input := iom; destination := s'; output := oom |} :: []) := fun l s s' iom oom Hptrans => - finite_ptrace_extend s' [] - (finite_ptrace_empty s' (protocol_transition_destination Hptrans)) + finite_valid_trace_from_extend s' [] + (finite_valid_trace_from_empty s' (input_valid_transition_destination Hptrans)) _ _ _ _ Hptrans. (** -To complete our definition of a finite protocol trace, we must also guarantee that <> is an +To complete our definition of a finite valid trace, we must also guarantee that <> is an initial state according to the protocol. *) - Definition finite_protocol_trace (s : state) (ls : list transition_item) : Prop := - finite_protocol_trace_from s ls /\ initial_state_prop s. + Definition finite_valid_trace (s : state) (ls : list transition_item) : Prop := + finite_valid_trace_from s ls /\ initial_state_prop s. - Opaque finite_protocol_trace. + Opaque finite_valid_trace. (** In the remainder of the section we provide various results allowing us to @@ -1053,58 +1043,58 @@ prove or decompose the above properties in proofs. (** This is a bit more useful than the small proof suggests, because applying it always leaves just one subgoal. - The tactical <> + The tactical <> only works if the assumption is available, which may require an <> and writing out the full VLSM and state expressions as part of the proof script. *) - Lemma finite_protocol_trace_empty (s : state): + Lemma finite_valid_trace_empty (s : state): vinitial_state_prop X s -> - finite_protocol_trace s []. + finite_valid_trace s []. Proof. - split;[constructor;apply initial_is_protocol|];assumption. + split;[constructor;apply initial_state_is_valid|];assumption. Qed. - Lemma finite_ptrace_first_valid_transition + Lemma finite_valid_trace_first_valid_transition (s : state) (tr : list transition_item) (te : transition_item) - (Htr : finite_protocol_trace_from s (te :: tr)) - : protocol_transition (l te) (s, input te) (destination te, output te). + (Htr : finite_valid_trace_from s (te :: tr)) + : input_valid_transition (l te) (s, input te) (destination te, output te). Proof. inversion Htr. assumption. Qed. - Lemma finite_ptrace_first_pstate + Lemma finite_valid_trace_first_pstate (s : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from s tr) - : protocol_state_prop s. + (Htr : finite_valid_trace_from s tr) + : valid_state_prop s. Proof. inversion Htr; subst; [assumption|]. apply Ht. Qed. - Lemma finite_ptrace_tail + Lemma finite_valid_trace_tail (s : state) (tr : list transition_item) (te : transition_item) - (Htr : finite_protocol_trace_from s (te :: tr)) - : finite_protocol_trace_from (destination te) tr. + (Htr : finite_valid_trace_from s (te :: tr)) + : finite_valid_trace_from (destination te) tr. Proof. inversion Htr. assumption. Qed. - Lemma finite_ptrace_last_pstate + Lemma finite_valid_trace_last_pstate (s : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from s tr) - : protocol_state_prop (finite_trace_last s tr). + (Htr : finite_valid_trace_from s tr) + : valid_state_prop (finite_trace_last s tr). Proof. generalize dependent s. induction tr; intros. - - simpl. apply finite_ptrace_first_pstate with []. assumption. - - apply finite_ptrace_tail in Htr. + - simpl. apply finite_valid_trace_first_pstate with []. assumption. + - apply finite_valid_trace_tail in Htr. apply IHtr in Htr. replace (finite_trace_last s (a :: tr)) @@ -1116,15 +1106,15 @@ prove or decompose the above properties in proofs. reflexivity. Qed. - Lemma protocol_transition_to + Lemma input_valid_transition_to (s : state) (tr : list transition_item) (tr1 tr2 : list transition_item) (te : transition_item) - (Htr : finite_protocol_trace_from s tr) + (Htr : finite_valid_trace_from s tr) (Heq : tr = tr1 ++ [te] ++ tr2) (lst1 := finite_trace_last s tr1) - : protocol_transition (l te) (lst1, input te) (destination te, output te). + : input_valid_transition (l te) (lst1, input te) (destination te, output te). Proof. generalize dependent s. generalize dependent tr. induction tr1. @@ -1135,76 +1125,76 @@ prove or decompose the above properties in proofs. apply IHtr1. assumption. Qed. - Lemma finite_ptrace_consecutive_valid_transition + Lemma finite_valid_trace_consecutive_valid_transition (s : state) (tr : list transition_item) (tr1 tr2 : list transition_item) (te1 te2 : transition_item) - (Htr : finite_protocol_trace_from s tr) + (Htr : finite_valid_trace_from s tr) (Heq : tr = tr1 ++ [te1; te2] ++ tr2) - : protocol_transition (l te2) (destination te1, input te2) (destination te2, output te2). + : input_valid_transition (l te2) (destination te1, input te2) (destination te2, output te2). Proof. change ([te1; te2] ++ tr2) with ([te1] ++ [te2] ++ tr2) in Heq. rewrite app_assoc in Heq. - specialize (protocol_transition_to s tr (tr1 ++ [te1]) tr2 te2 Htr Heq) + specialize (input_valid_transition_to s tr (tr1 ++ [te1]) tr2 te2 Htr Heq) as Ht. rewrite finite_trace_last_is_last in Ht. assumption. Qed. - Lemma protocol_trace_output_is_protocol + Lemma valid_trace_output_is_valid (is : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from is tr) + (Htr : finite_valid_trace_from is tr) (m : message) (Houtput : trace_has_message (field_selector output) m tr) - : protocol_message_prop m. + : valid_message_prop m. Proof. revert is Htr. induction Houtput as [item tr' Hm| item tr'];intros;inversion Htr; subst. - - simpl in Hm. + - simpl in Hm. subst. - apply protocol_transition_out in Ht. + apply input_valid_transition_out in Ht. assumption. - apply (IHHoutput s). assumption. Qed. - Lemma protocol_trace_input_is_protocol + Lemma valid_trace_input_is_valid (is : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from is tr) + (Htr : finite_valid_trace_from is tr) (m : message) (Hinput : trace_has_message (field_selector input) m tr) - : protocol_message_prop m. + : valid_message_prop m. Proof. revert is Htr. induction Hinput as [item tr' Hm| item tr'];intros ;inversion Htr as [|s _tr' Htr' _is iom oom l Ht ];subst. - simpl in Hm. subst. - apply protocol_transition_in in Ht. + apply input_valid_transition_in in Ht. assumption. - apply (IHHinput s). assumption. Qed. - Lemma protocol_trace_observed_is_protocol + Lemma valid_trace_observed_is_valid (is : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from is tr) + (Htr : finite_valid_trace_from is tr) (m : message) (Hobserved : trace_has_message item_sends_or_receives m tr) - : protocol_message_prop m. + : valid_message_prop m. Proof. apply trace_has_message_observed_iff in Hobserved as [Hm |Hm] - ; revert Htr m Hm; [apply protocol_trace_input_is_protocol | apply protocol_trace_output_is_protocol]. + ; revert Htr m Hm; [apply valid_trace_input_is_valid | apply valid_trace_output_is_valid]. Qed. Lemma first_transition_valid (s : state) (te : transition_item) - : finite_protocol_trace_from s [te] <-> protocol_transition (l te) (s, input te) (destination te, output te). + : finite_valid_trace_from s [te] <-> input_valid_transition (l te) (s, input te) (destination te, output te). Proof. split. @@ -1212,26 +1202,26 @@ prove or decompose the above properties in proofs. inversion Htr. assumption. - destruct te. simpl. intro Ht. - apply protocol_transition_destination in Ht as Hdestination0. + apply input_valid_transition_destination in Ht as Hdestination0. constructor; [|assumption]. constructor. assumption. Qed. Lemma extend_right_finite_trace_from (s1 : state) (ts : list transition_item) - (Ht12 : finite_protocol_trace_from s1 ts) + (Ht12 : finite_valid_trace_from s1 ts) (l3 : label) (s2 := finite_trace_last s1 ts) (iom3 : option message) (s3 : state) (oom3 : option message) - (Hv23 : protocol_transition l3 (s2, iom3) (s3, oom3)) - : finite_protocol_trace_from s1 (ts ++ [{| l := l3; destination := s3; input := iom3; output := oom3 |}]). + (Hv23 : input_valid_transition l3 (s2, iom3) (s3, oom3)) + : finite_valid_trace_from s1 (ts ++ [{| l := l3; destination := s3; input := iom3; output := oom3 |}]). Proof. induction Ht12. - - simpl. apply finite_ptrace_singleton;assumption. + - simpl. apply finite_valid_trace_singleton;assumption. - rewrite <- app_comm_cons. - apply finite_ptrace_extend; try assumption. + apply finite_valid_trace_from_extend; try assumption. simpl in IHHt12. apply IHHt12. unfold s2 in *; clear s2. rewrite finite_trace_last_cons in Hv23. @@ -1239,22 +1229,22 @@ prove or decompose the above properties in proofs. Qed. (** -We can now prove several general properties of [finite_protocol_trace]s. For example, +We can now prove several general properties of [finite_valid_trace]s. For example, the following lemma states that given two such traces, such that the latter's starting state is equal to the former's last state, it is possible to _concatenate_ them into a single -[finite_protocol_trace]. +[finite_valid_trace]. *) - Lemma finite_protocol_trace_from_app_iff (s : state) (ls ls' : list transition_item) (s' := finite_trace_last s ls) - : finite_protocol_trace_from s ls /\ finite_protocol_trace_from s' ls' + Lemma finite_valid_trace_from_app_iff (s : state) (ls ls' : list transition_item) (s' := finite_trace_last s ls) + : finite_valid_trace_from s ls /\ finite_valid_trace_from s' ls' <-> - finite_protocol_trace_from s (ls ++ ls'). + finite_valid_trace_from s (ls ++ ls'). Proof. subst s'. revert s. induction ls;intro s. - rewrite finite_trace_last_nil. simpl. - intuition (eauto using finite_ptrace_first_pstate, finite_ptrace_empty). + intuition (eauto using finite_valid_trace_first_pstate, finite_valid_trace_from_empty). - rewrite finite_trace_last_cons. simpl. specialize (IHls (destination a)). split. @@ -1266,46 +1256,46 @@ is equal to the former's last state, it is possible to _concatenate_ them into a split;[constructor|];assumption. Qed. - Lemma finite_protocol_trace_from_rev_ind + Lemma finite_valid_trace_from_rev_ind (P : state -> list transition_item -> Prop) (Hempty: forall s, - protocol_state_prop s -> P s nil) + valid_state_prop s -> P s nil) (Hextend : forall s tr, - finite_protocol_trace_from s tr -> + finite_valid_trace_from s tr -> P s tr -> forall sf iom oom l - (Hx: protocol_transition l (finite_trace_last s tr,iom) (sf,oom)), + (Hx: input_valid_transition l (finite_trace_last s tr,iom) (sf,oom)), let x:= {|l:=l; input:=iom; destination:=sf; output:=oom|} in P s (tr++[x])): forall s tr, - finite_protocol_trace_from s tr -> + finite_valid_trace_from s tr -> P s tr. Proof. induction tr using rev_ind; intro Htr. - inversion Htr. apply Hempty. congruence. - - apply finite_protocol_trace_from_app_iff in Htr. + - apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Htr Hx]. destruct x; apply (Hextend _ _ Htr (IHtr Htr)). inversion Hx; congruence. Qed. - Lemma finite_protocol_trace_rev_ind + Lemma finite_valid_trace_rev_ind (P : state -> list transition_item -> Prop) (Hempty: forall si, initial_state_prop si -> P si nil) (Hextend : forall si tr, - finite_protocol_trace si tr -> + finite_valid_trace si tr -> P si tr -> forall sf iom oom l - (Hx: protocol_transition l (finite_trace_last si tr,iom) (sf,oom)), + (Hx: input_valid_transition l (finite_trace_last si tr,iom) (sf,oom)), let x:= {|l:=l; input:=iom; destination:=sf; output:=oom|} in P si (tr++[x])): forall si tr, - finite_protocol_trace si tr -> + finite_valid_trace si tr -> P si tr. Proof. intros si tr [Htr Hinit]. - induction Htr using finite_protocol_trace_from_rev_ind. + induction Htr using finite_valid_trace_from_rev_ind. - apply Hempty;auto. - apply Hextend;[split|..];auto. Qed. @@ -1314,31 +1304,31 @@ is equal to the former's last state, it is possible to _concatenate_ them into a traces. *) - Lemma finite_protocol_trace_from_prefix + Lemma finite_valid_trace_from_prefix (s : state) (ls : list transition_item) - (Htr : finite_protocol_trace_from s ls) + (Htr : finite_valid_trace_from s ls) (n : nat) - : finite_protocol_trace_from s (list_prefix ls n). + : finite_valid_trace_from s (list_prefix ls n). Proof. specialize (list_prefix_suffix ls n); intro Hdecompose. rewrite <- Hdecompose in Htr. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [Hpr _]. assumption. Qed. - Lemma finite_protocol_trace_from_suffix + Lemma finite_valid_trace_from_suffix (s : state) (ls : list transition_item) - (Htr : finite_protocol_trace_from s ls) + (Htr : finite_valid_trace_from s ls) (n : nat) (nth : state) (Hnth : finite_trace_nth s ls n = Some nth) - : finite_protocol_trace_from nth (list_suffix ls n). + : finite_valid_trace_from nth (list_suffix ls n). Proof. rewrite <- (list_prefix_suffix ls n) in Htr. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [_ Htr]. replace (finite_trace_last s (list_prefix ls n)) with nth in Htr;[assumption|]. { @@ -1352,18 +1342,18 @@ traces. } Qed. - Lemma finite_protocol_trace_from_segment + Lemma finite_valid_trace_from_segment (s : state) (ls : list transition_item) - (Htr : finite_protocol_trace_from s ls) + (Htr : finite_valid_trace_from s ls) (n1 n2 : nat) (Hle : n1 <= n2) (n1th : state) (Hnth : finite_trace_nth s ls n1 = Some n1th) - : finite_protocol_trace_from n1th (list_segment ls n1 n2). + : finite_valid_trace_from n1th (list_segment ls n1 n2). Proof. - apply finite_protocol_trace_from_suffix with s. - - apply finite_protocol_trace_from_prefix. assumption. + apply finite_valid_trace_from_suffix with s. + - apply finite_valid_trace_from_prefix. assumption. - destruct n1;[assumption|]. unfold finite_trace_nth in Hnth |- *. simpl in Hnth |- *. @@ -1373,11 +1363,11 @@ traces. (* begin hide *) - Lemma can_emit_from_protocol_trace + Lemma can_emit_from_valid_trace (si : state) (m : message) (tr : list transition_item) - (Hprotocol: finite_protocol_trace si tr) + (Htr: finite_valid_trace si tr) (Hm : trace_has_message (field_selector output) m tr) : can_emit m. Proof. @@ -1386,8 +1376,8 @@ traces. apply elem_of_list_split in Hin. destruct Hin as [l1 [l2 Hconcat]]. unfold can_emit. - destruct Hprotocol as [Hprotocol _]. - specialize (protocol_transition_to _ _ _ _ _ Hprotocol Hconcat). + destruct Htr as [Htr _]. + specialize (input_valid_transition_to _ _ _ _ _ Htr Hconcat). intros Ht. simpl in Houtput, Ht. rewrite Houtput in Ht. @@ -1397,54 +1387,54 @@ traces. (* End Hide *) (** -** Finite [protocol_trace]s with a final state +** Finite [valid_trace]s with a final state *) (** -It is often necessary to refer to know ending state of a [finite_protocol_trace_from]. +It is often necessary to refer to know ending state of a [finite_valid_trace_from]. This is either the [destination] of the [last] [transition_item] in the trace, or the starting state. To avoid repeating reasoning about [last], we define variants of -[finite_protocol_trace_from] and [finite_protocol_trace] +[finite_valid_trace_from] and [finite_valid_trace] that include the final state, and give appropriate induction principles. *) -(** The final state of a finite portion of a protocol trace. - This is defined over [finite_protocol_trace_from] because +(** The final state of a finite portion of a valid trace. + This is defined over [finite_valid_trace_from] because an initial state is necessary in case <> is empty, and this allows the definition to have only one non-implicit parameter. *) - Inductive finite_protocol_trace_from_to : state -> state -> list transition_item -> Prop := - | finite_ptrace_from_to_empty : forall (s : state) - (Hs : protocol_state_prop s), - finite_protocol_trace_from_to s s [] - | finite_ptrace_from_to_extend : forall (s f : state) (tl : list transition_item) - (Htl : finite_protocol_trace_from_to s f tl) + Inductive finite_valid_trace_from_to : state -> state -> list transition_item -> Prop := + | finite_valid_trace_from_to_empty : forall (s : state) + (Hs : valid_state_prop s), + finite_valid_trace_from_to s s [] + | finite_valid_trace_from_to_extend : forall (s f : state) (tl : list transition_item) + (Htl : finite_valid_trace_from_to s f tl) (s' : state) (iom oom : option message) (l : label) - (Ht : protocol_transition l (s', iom) (s, oom)), - finite_protocol_trace_from_to s' f ({| l := l; input := iom; destination := s; output := oom |} :: tl). + (Ht : input_valid_transition l (s', iom) (s, oom)), + finite_valid_trace_from_to s' f ({| l := l; input := iom; destination := s; output := oom |} :: tl). - Lemma finite_ptrace_from_to_singleton s s' iom oom l - : protocol_transition l (s, iom) (s', oom) -> - finite_protocol_trace_from_to s s' [{| l := l; input := iom; destination := s'; output := oom |}]. + Lemma finite_valid_trace_from_to_singleton s s' iom oom l + : input_valid_transition l (s, iom) (s', oom) -> + finite_valid_trace_from_to s s' [{| l := l; input := iom; destination := s'; output := oom |}]. Proof. intro Ht. constructor;[|assumption]. constructor. - apply protocol_transition_destination in Ht. + apply input_valid_transition_destination in Ht. assumption. Qed. - Lemma finite_protocol_trace_from_to_forget_last - s f tr : finite_protocol_trace_from_to s f tr -> finite_protocol_trace_from s tr. + Lemma finite_valid_trace_from_to_forget_last + s f tr : finite_valid_trace_from_to s f tr -> finite_valid_trace_from s tr. Proof. induction 1;constructor;auto. Qed. - Lemma finite_protocol_trace_from_to_last - s f tr : finite_protocol_trace_from_to s f tr -> finite_trace_last s tr = f. + Lemma finite_valid_trace_from_to_last + s f tr : finite_valid_trace_from_to s f tr -> finite_trace_last s tr = f. Proof. induction 1. - apply finite_trace_last_nil. @@ -1452,11 +1442,11 @@ that include the final state, and give appropriate induction principles. Qed. - Lemma finite_protocol_trace_from_add_last + Lemma finite_valid_trace_from_add_last s f tr : - finite_protocol_trace_from s tr -> + finite_valid_trace_from s tr -> finite_trace_last s tr = f -> - finite_protocol_trace_from_to s f tr. + finite_valid_trace_from_to s f tr. Proof. intro Hfrom. induction Hfrom. @@ -1466,113 +1456,113 @@ that include the final state, and give appropriate induction principles. constructor;auto. Qed. - Lemma finite_protocol_trace_from_to_first_pstate - s f tr : finite_protocol_trace_from_to s f tr -> protocol_state_prop s. + Lemma finite_valid_trace_from_to_first_pstate + s f tr : finite_valid_trace_from_to s f tr -> valid_state_prop s. Proof. intro Htr. - apply finite_protocol_trace_from_to_forget_last in Htr. - apply finite_ptrace_first_pstate in Htr. + apply finite_valid_trace_from_to_forget_last in Htr. + apply finite_valid_trace_first_pstate in Htr. assumption. Qed. - Lemma finite_protocol_trace_from_to_last_pstate - s f tr : finite_protocol_trace_from_to s f tr -> protocol_state_prop f. + Lemma finite_valid_trace_from_to_last_pstate + s f tr : finite_valid_trace_from_to s f tr -> valid_state_prop f. Proof. intro Htr. - rewrite <- (finite_protocol_trace_from_to_last _ _ _ Htr). - apply finite_ptrace_last_pstate. - apply finite_protocol_trace_from_to_forget_last in Htr. + rewrite <- (finite_valid_trace_from_to_last _ _ _ Htr). + apply finite_valid_trace_last_pstate. + apply finite_valid_trace_from_to_forget_last in Htr. assumption. Qed. - Lemma finite_protocol_trace_from_to_app + Lemma finite_valid_trace_from_to_app (m s f: state) (ls ls' : list transition_item) - : finite_protocol_trace_from_to s m ls - -> finite_protocol_trace_from_to m f ls' - -> finite_protocol_trace_from_to s f (ls ++ ls'). + : finite_valid_trace_from_to s m ls + -> finite_valid_trace_from_to m f ls' + -> finite_valid_trace_from_to s f (ls ++ ls'). Proof. intros Hl Hl';induction Hl;simpl. - trivial. - constructor;auto. Qed. - Lemma finite_protocol_trace_from_to_app_split + Lemma finite_valid_trace_from_to_app_split (s f: state) (ls ls' : list transition_item) - : finite_protocol_trace_from_to s f (ls ++ ls') -> + : finite_valid_trace_from_to s f (ls ++ ls') -> let m := finite_trace_last s ls in - finite_protocol_trace_from_to s m ls - /\ finite_protocol_trace_from_to m f ls'. + finite_valid_trace_from_to s m ls + /\ finite_valid_trace_from_to m f ls'. Proof. revert s;induction ls;intros s;simpl. - rewrite finite_trace_last_nil. intro Htr. split;[|assumption]. - apply finite_protocol_trace_from_to_first_pstate in Htr. + apply finite_valid_trace_from_to_first_pstate in Htr. constructor;assumption. - rewrite finite_trace_last_cons. inversion 1; subst; simpl in *. apply IHls in Htl as []. - auto using finite_ptrace_from_to_extend. + auto using finite_valid_trace_from_to_extend. Qed. - Definition finite_protocol_trace_init_to si sf tr : Prop - := finite_protocol_trace_from_to si sf tr + Definition finite_valid_trace_init_to si sf tr : Prop + := finite_valid_trace_from_to si sf tr /\ initial_state_prop si. - Lemma finite_protocol_trace_init_add_last si sf tr: - finite_protocol_trace si tr -> + Lemma finite_valid_trace_init_add_last si sf tr: + finite_valid_trace si tr -> finite_trace_last si tr = sf -> - finite_protocol_trace_init_to si sf tr. + finite_valid_trace_init_to si sf tr. Proof. intros [Htr Hinit] Hf. - split;eauto using finite_protocol_trace_from_add_last. + split;eauto using finite_valid_trace_from_add_last. Qed. - Lemma finite_protocol_trace_init_to_forget_last si sf tr: - finite_protocol_trace_init_to si sf tr -> - finite_protocol_trace si tr. + Lemma finite_valid_trace_init_to_forget_last si sf tr: + finite_valid_trace_init_to si sf tr -> + finite_valid_trace si tr. Proof. intros [Hinit Htr]. - split;eauto using finite_protocol_trace_from_to_forget_last. + split;eauto using finite_valid_trace_from_to_forget_last. Qed. - Lemma finite_protocol_trace_init_to_last si sf tr: - finite_protocol_trace_init_to si sf tr -> + Lemma finite_valid_trace_init_to_last si sf tr: + finite_valid_trace_init_to si sf tr -> finite_trace_last si tr = sf. Proof. intros [Htr _]. - eauto using finite_protocol_trace_from_to_last. + eauto using finite_valid_trace_from_to_last. Qed. Lemma extend_right_finite_trace_from_to (s1 s2 : state) (ts : list transition_item) - (Ht12 : finite_protocol_trace_from_to s1 s2 ts) + (Ht12 : finite_valid_trace_from_to s1 s2 ts) (l3 : label) (iom3 : option message) (s3 : state) (oom3 : option message) - (Hv23 : protocol_transition l3 (s2, iom3) (s3, oom3)) - : finite_protocol_trace_from_to s1 s3 (ts ++ [{| l := l3; destination := s3; input := iom3; output := oom3 |}]). + (Hv23 : input_valid_transition l3 (s2, iom3) (s3, oom3)) + : finite_valid_trace_from_to s1 s3 (ts ++ [{| l := l3; destination := s3; input := iom3; output := oom3 |}]). Proof. induction Ht12. - - simpl. apply finite_ptrace_from_to_singleton;assumption. + - simpl. apply finite_valid_trace_from_to_singleton;assumption. - rewrite <- app_comm_cons. - apply finite_ptrace_from_to_extend; auto. + apply finite_valid_trace_from_to_extend; auto. Qed. - Lemma finite_protocol_trace_from_to_rev_ind + Lemma finite_valid_trace_from_to_rev_ind (P : state -> state -> list transition_item -> Prop) (Hempty: forall si - (Hsi : protocol_state_prop si), + (Hsi : valid_state_prop si), P si si nil) (Hextend : forall si s tr (IHtr : P si s tr) - (Htr : finite_protocol_trace_from_to si s tr) + (Htr : finite_valid_trace_from_to si s tr) sf iom oom l - (Ht : protocol_transition l (s,iom) (sf,oom)), + (Ht : input_valid_transition l (s,iom) (sf,oom)), P si sf (tr++[{|l:=l; input:=iom; destination:=sf; output:=oom|}])): forall si sf tr, - finite_protocol_trace_from_to si sf tr -> + finite_valid_trace_from_to si sf tr -> P si sf tr. Proof. intros si sf tr Htr. @@ -1580,7 +1570,7 @@ that include the final state, and give appropriate induction principles. induction tr using rev_ind; intros sf Htr. - inversion Htr;subst. apply Hempty;assumption. - - apply finite_protocol_trace_from_to_app_split in Htr. + - apply finite_valid_trace_from_to_app_split in Htr. destruct Htr as [Htr Hstep]. inversion Hstep;subst. inversion Htl;subst. @@ -1590,24 +1580,24 @@ that include the final state, and give appropriate induction principles. assumption. Qed. - Lemma finite_protocol_trace_init_to_rev_ind + Lemma finite_valid_trace_init_to_rev_ind (P : state -> state -> list transition_item -> Prop) (Hempty: forall si (Hsi : initial_state_prop si), P si si nil) (Hextend : forall si s tr (IHtr : P si s tr) - (Htr : finite_protocol_trace_init_to si s tr) + (Htr : finite_valid_trace_init_to si s tr) sf iom oom l - (Ht : protocol_transition l (s,iom) (sf,oom)), + (Ht : input_valid_transition l (s,iom) (sf,oom)), P si sf (tr++[{|l:=l; input:=iom; destination:=sf; output:=oom|}])): forall si sf tr, - finite_protocol_trace_init_to si sf tr -> + finite_valid_trace_init_to si sf tr -> P si sf tr. Proof. intros si sf tr Htr. destruct Htr as [Htr Hinit]. - induction Htr using finite_protocol_trace_from_to_rev_ind. + induction Htr using finite_valid_trace_from_to_rev_ind. - apply Hempty. assumption. - apply Hextend with s. + apply IHHtr. assumption. @@ -1615,51 +1605,51 @@ that include the final state, and give appropriate induction principles. + assumption. Qed. -(** An inductive protocol trace property which also identifies the final message. +(** An inductive valid trace property which also identifies the final message. -As shown by the [finite_protocol_trace_init_to_emit_protocol] and -[finite_protocol_trace_init_to_emit_protocol_rev] lemmas below, this definition -is the trace-equivalent of the [protocol_prop]erty. +As shown by the [finite_valid_trace_init_to_emit_valid_state_message] and +[finite_valid_trace_init_to_emit_valid_state_message_rev] lemmas below, this definition +is the trace-equivalent of the [valid_state_message_prop]erty. -This inductive property is reflecting that fact that a that @protocol_prop (s,om)@ -holds only if @s@ and @om@ are the final state and output of an initial protocol +This inductive property is reflecting that fact that a that <> +holds only if <> and <> are the final state and output of an initial valid trace, or a pair of an initial state and option-initial message. -It follows the inductive structure of @protocol_prop@, but augments every node of +It follows the inductive structure of <>, but augments every node of the tree with such an exhibiting trace. Its main benefit is that when performing induction over it, one can also use the induction hypothesis for the (trace generating the) message being received. Although this definition could be used directly, we prefer to use it to derive -a stronger induction principle ([finite_protocol_trace_init_to_rev_strong_ind]) -over [finite_protocol_trace_init_to] traces. +a stronger induction principle ([finite_valid_trace_init_to_rev_strong_ind]) +over [finite_valid_trace_init_to] traces. *) - Inductive finite_protocol_trace_init_to_emit : state -> state -> option message -> list transition_item -> Prop := - | finite_ptrace_init_to_emit_empty : forall (is : state) (om : option message) + Inductive finite_valid_trace_init_to_emit : state -> state -> option message -> list transition_item -> Prop := + | finite_valid_trace_init_to_emit_empty : forall (is : state) (om : option message) (His : initial_state_prop is) (Him : option_initial_message_prop om), - finite_protocol_trace_init_to_emit is is om [] - | finite_ptrace_init_to_emit_extend + finite_valid_trace_init_to_emit is is om [] + | finite_valid_trace_init_to_emit_extend : forall (is s : state) (_om : option message) (tl : list transition_item) - (Hs : finite_protocol_trace_init_to_emit is s _om tl) + (Hs : finite_valid_trace_init_to_emit is s _om tl) (iom_is iom_s : state) (iom : option message) (iom_tl : list transition_item) - (Hiom : finite_protocol_trace_init_to_emit iom_is iom_s iom iom_tl) + (Hiom : finite_valid_trace_init_to_emit iom_is iom_s iom iom_tl) (l : label) (Hv : valid l (s, iom)) (s' : state) (oom : option message) (Ht : transition l (s, iom) = (s', oom)), - finite_protocol_trace_init_to_emit is s' oom (tl ++ [{| l := l; input := iom; destination := s'; output := oom |}]). + finite_valid_trace_init_to_emit is s' oom (tl ++ [{| l := l; input := iom; destination := s'; output := oom |}]). - Lemma finite_protocol_trace_init_to_emit_initial_state + Lemma finite_valid_trace_init_to_emit_initial_state (is f : state) (om : option message) (tl : list transition_item) - (Htl : finite_protocol_trace_init_to_emit is f om tl) + (Htl : finite_valid_trace_init_to_emit is f om tl) : initial_state_prop is. Proof. induction Htl; assumption. Qed. -(** A property characterizing the "emit" message of [finite_protocol_trace_init_to_emit]. +(** A property characterizing the "emit" message of [finite_valid_trace_init_to_emit]. There are two cases: (1) either the trace is empty, and then we require the message to be initial; or (2) the trace is not empty, and the message @@ -1672,9 +1662,9 @@ is the output of the last transition. - exact (option_initial_message_prop om). Defined. - Lemma finite_protocol_trace_init_to_emit_output + Lemma finite_valid_trace_init_to_emit_output (s f : state) (om : option message) (tl : list transition_item) - (Htl : finite_protocol_trace_init_to_emit s f om tl) + (Htl : finite_valid_trace_init_to_emit s f om tl) : empty_initial_message_or_final_output tl om. Proof. unfold empty_initial_message_or_final_output. @@ -1687,21 +1677,21 @@ is the output of the last transition. apply app_inj_tail, proj2 in Heqtl. subst. reflexivity. Qed. - Lemma finite_protocol_trace_init_to_emit_protocol + Lemma finite_valid_trace_init_to_emit_valid_state_message (s f : state) (om : option message) (tl : list transition_item) - (Htl : finite_protocol_trace_init_to_emit s f om tl) - : protocol_prop f om. + (Htl : finite_valid_trace_init_to_emit s f om tl) + : valid_state_message_prop f om. Proof. induction Htl. - - apply protocol_initial; assumption. - - apply protocol_generated with s _om iom_s iom l0; assumption. + - apply valid_initial_state_message; assumption. + - apply valid_generated_state_message with s _om iom_s iom l0; assumption. Qed. - Lemma finite_protocol_trace_init_to_emit_protocol_rev + Lemma finite_valid_trace_init_to_emit_valid_state_message_rev f om - (Hp : protocol_prop f om) + (Hp : valid_state_message_prop f om) : exists (s : state) (tl : list transition_item), - finite_protocol_trace_init_to_emit s f om tl. + finite_valid_trace_init_to_emit s f om tl. Proof. induction Hp. - exists s, []. constructor; assumption. @@ -1709,110 +1699,110 @@ is the output of the last transition. destruct IHHp2 as [om_is [om_tl Hom]]. eexists; eexists. apply - (finite_ptrace_init_to_emit_extend _ _ _ _ Hs _ _ _ _ Hom _ Hv _ _ Ht). + (finite_valid_trace_init_to_emit_extend _ _ _ _ Hs _ _ _ _ Hom _ Hv _ _ Ht). Qed. - Lemma finite_protocol_trace_init_to_emit_forget_emit + Lemma finite_valid_trace_init_to_emit_forget_emit (s f : state) (_om : option message) (tl : list transition_item) - (Htl : finite_protocol_trace_init_to_emit s f _om tl) - : finite_protocol_trace_init_to s f tl. + (Htl : finite_valid_trace_init_to_emit s f _om tl) + : finite_valid_trace_init_to s f tl. Proof. - apply finite_protocol_trace_init_to_emit_initial_state in Htl as Hinit. + apply finite_valid_trace_init_to_emit_initial_state in Htl as Hinit. split; [|assumption]. clear Hinit. induction Htl. - - constructor. apply initial_is_protocol. assumption. - - apply finite_protocol_trace_from_to_app with s; [assumption|]. - apply finite_ptrace_from_to_singleton. - apply finite_protocol_trace_init_to_emit_protocol in Htl1. - apply finite_protocol_trace_init_to_emit_protocol in Htl2. + - constructor. apply initial_state_is_valid. assumption. + - apply finite_valid_trace_from_to_app with s; [assumption|]. + apply finite_valid_trace_from_to_singleton. + apply finite_valid_trace_init_to_emit_valid_state_message in Htl1. + apply finite_valid_trace_init_to_emit_valid_state_message in Htl2. repeat split; [..|assumption|assumption]; eexists; [exact Htl1 | exact Htl2]. Qed. - Lemma finite_protocol_trace_init_to_add_emit + Lemma finite_valid_trace_init_to_add_emit (s f : state) (tl : list transition_item) - (Htl : finite_protocol_trace_init_to s f tl) - : finite_protocol_trace_init_to_emit s f (finite_trace_last_output tl) tl. + (Htl : finite_valid_trace_init_to s f tl) + : finite_valid_trace_init_to_emit s f (finite_trace_last_output tl) tl. Proof. - induction Htl using finite_protocol_trace_init_to_rev_ind. + induction Htl using finite_valid_trace_init_to_rev_ind. - constructor; [assumption | exact I]. - rewrite finite_trace_last_output_is_last. simpl. destruct Ht as [[_ [[_s Hiom] Hv]] Ht]. - specialize (finite_protocol_trace_init_to_emit_protocol_rev _ _ Hiom) as [iom_s [iom_tr Hiom_tr]]. - apply (finite_ptrace_init_to_emit_extend _ _ _ _ IHHtl _ _ _ _ Hiom_tr _ Hv _ _ Ht). + specialize (finite_valid_trace_init_to_emit_valid_state_message_rev _ _ Hiom) as [iom_s [iom_tr Hiom_tr]]. + apply (finite_valid_trace_init_to_emit_extend _ _ _ _ IHHtl _ _ _ _ Hiom_tr _ Hv _ _ Ht). Qed. -(** Inspired by [finite_protocol_trace_init_to_add_emit], we can derive -an induction principle for [finite_protocol_trace_init_to] stronger than -[finite_protocol_trace_init_to_rev_ind], which allows the induction hypothesis +(** Inspired by [finite_valid_trace_init_to_add_emit], we can derive +an induction principle for [finite_valid_trace_init_to] stronger than +[finite_valid_trace_init_to_rev_ind], which allows the induction hypothesis to be used for the trace generating the message received in the last transition. *) - Lemma finite_protocol_trace_init_to_rev_strong_ind + Lemma finite_valid_trace_init_to_rev_strong_ind (P : state -> state -> list transition_item -> Prop) (Hempty: forall is (His : initial_state_prop is), P is is nil) (Hextend : forall is s tr (IHs : P is s tr) - (Hs : finite_protocol_trace_init_to is s tr) + (Hs : finite_valid_trace_init_to is s tr) iom iom_si iom_s iom_tr (Heqiom : empty_initial_message_or_final_output iom_tr iom) (IHiom : P iom_si iom_s iom_tr) - (Hiom : finite_protocol_trace_init_to iom_si iom_s iom_tr) + (Hiom : finite_valid_trace_init_to iom_si iom_s iom_tr) sf oom l - (Ht : protocol_transition l (s,iom) (sf,oom)), + (Ht : input_valid_transition l (s,iom) (sf,oom)), P is sf (tr++[{|l:=l; input:=iom; destination:=sf; output:=oom|}])) : forall si sf tr, - finite_protocol_trace_init_to si sf tr -> + finite_valid_trace_init_to si sf tr -> P si sf tr. Proof. intros is sf tr Htr. - apply finite_protocol_trace_init_to_add_emit in Htr. + apply finite_valid_trace_init_to_add_emit in Htr. remember (finite_trace_last_output tr) as om. clear Heqom. induction Htr. - apply Hempty. assumption. - - assert (Hpt : protocol_transition l0 (s, iom) (s', oom)). - { apply finite_protocol_trace_init_to_emit_protocol in Htr1. - apply finite_protocol_trace_init_to_emit_protocol in Htr2. + - assert (Hivt : input_valid_transition l0 (s, iom) (s', oom)). + { apply finite_valid_trace_init_to_emit_valid_state_message in Htr1. + apply finite_valid_trace_init_to_emit_valid_state_message in Htr2. repeat split; [..|assumption|assumption]; eexists; [exact Htr1 | exact Htr2]. } - apply finite_protocol_trace_init_to_emit_output in Htr2 as Houtput. - apply finite_protocol_trace_init_to_emit_forget_emit in Htr1. - apply finite_protocol_trace_init_to_emit_forget_emit in Htr2. - apply (Hextend _ _ _ IHHtr1 Htr1 _ _ _ _ Houtput IHHtr2 Htr2 _ _ _ Hpt). + apply finite_valid_trace_init_to_emit_output in Htr2 as Houtput. + apply finite_valid_trace_init_to_emit_forget_emit in Htr1. + apply finite_valid_trace_init_to_emit_forget_emit in Htr2. + apply (Hextend _ _ _ IHHtr1 Htr1 _ _ _ _ Houtput IHHtr2 Htr2 _ _ _ Hivt). Qed. (** *** Infinite [protcol_trace]s *) -(** We now define [infinite_protocol_trace]s. The definitions +(** We now define [infinite_valid_trace]s. The definitions resemble their finite counterparts, adapted to the technical necessities of defining infinite objects. Notably, <> is stored as a stream, as opposed to a list. *) - CoInductive infinite_protocol_trace_from : + CoInductive infinite_valid_trace_from : state -> Stream transition_item -> Prop := - | infinite_protocol_trace_extend : forall (s : state) (tl : Stream transition_item) - (Htl : infinite_protocol_trace_from s tl) + | infinite_valid_trace_from_extend : forall (s : state) (tl : Stream transition_item) + (Htl : infinite_valid_trace_from s tl) (s' : state) (iom oom : option message) (l : label) - (Ht : protocol_transition l (s', iom) (s, oom)), - infinite_protocol_trace_from s' (Cons {| l := l; input := iom; destination := s; output := oom |} tl). + (Ht : input_valid_transition l (s', iom) (s, oom)), + infinite_valid_trace_from s' (Cons {| l := l; input := iom; destination := s; output := oom |} tl). - Definition infinite_protocol_trace (s : state) (st : Stream transition_item) - := infinite_protocol_trace_from s st /\ initial_state_prop s. + Definition infinite_valid_trace (s : state) (st : Stream transition_item) + := infinite_valid_trace_from s st /\ initial_state_prop s. (** As for the finite case, the following lemmas help decompose teh above definitions, mostly reducing them to properties about their finite segments. *) - Lemma infinite_protocol_trace_consecutive_valid_transition + Lemma infinite_valid_trace_consecutive_valid_transition (is : state) (tr tr2 : Stream transition_item) (tr1 : list transition_item) (te1 te2 : transition_item) - (Htr : infinite_protocol_trace_from is tr) + (Htr : infinite_valid_trace_from is tr) (Heq : tr = stream_app (tr1 ++ [te1; te2]) tr2) - : protocol_transition (l te2) (destination te1, input te2) (destination te2, output te2). + : input_valid_transition (l te2) (destination te1, input te2) (destination te2, output te2). Proof. generalize dependent is. generalize dependent tr. induction tr1. @@ -1822,22 +1812,22 @@ definitions, mostly reducing them to properties about their finite segments. specialize (IHtr1 s Htl). assumption. Qed. - Lemma infinite_protocol_trace_from_app_iff + Lemma infinite_valid_trace_from_app_iff (s : state) (ls : list transition_item) (ls' : Stream transition_item) (s' := finite_trace_last s ls) - : finite_protocol_trace_from s ls /\ infinite_protocol_trace_from s' ls' + : finite_valid_trace_from s ls /\ infinite_valid_trace_from s' ls' <-> - infinite_protocol_trace_from s (stream_app ls ls'). + infinite_valid_trace_from s (stream_app ls ls'). Proof. intros. generalize dependent ls'. generalize dependent s. induction ls; intros; split. - destruct 1. assumption. - simpl; intros Hls'; split; try assumption. constructor. inversion Hls'; try assumption. - apply (protocol_transition_origin Ht). + apply (input_valid_transition_origin Ht). - simpl. intros [Htr Htr']. - destruct a. apply infinite_protocol_trace_extend. + destruct a. apply infinite_valid_trace_from_extend. + apply IHls. inversion Htr. split. apply Htl. unfold s' in Htr'. rewrite finite_trace_last_cons in Htr'. @@ -1849,25 +1839,25 @@ definitions, mostly reducing them to properties about their finite segments. + unfold s'. rewrite finite_trace_last_cons. assumption. Qed. - Lemma infinite_protocol_trace_from_prefix + Lemma infinite_valid_trace_from_prefix (s : state) (ls : Stream transition_item) - (Htr : infinite_protocol_trace_from s ls) + (Htr : infinite_valid_trace_from s ls) (n : nat) - : finite_protocol_trace_from s (stream_prefix ls n). + : finite_valid_trace_from s (stream_prefix ls n). Proof. specialize (stream_prefix_suffix ls n); intro Hdecompose. rewrite <- Hdecompose in Htr. - apply infinite_protocol_trace_from_app_iff in Htr. + apply infinite_valid_trace_from_app_iff in Htr. destruct Htr as [Hpr _]. assumption. Qed. - Lemma infinite_protocol_trace_from_prefix_rev + Lemma infinite_valid_trace_from_prefix_rev (s : state) (ls : Stream transition_item) - (Hpref: forall n : nat, finite_protocol_trace_from s (stream_prefix ls n)) - : infinite_protocol_trace_from s ls. + (Hpref: forall n : nat, finite_valid_trace_from s (stream_prefix ls n)) + : infinite_valid_trace_from s ls. Proof. revert s ls Hpref. cofix Hls. @@ -1883,28 +1873,28 @@ definitions, mostly reducing them to properties about their finite segments. assumption. Qed. - Lemma infinite_protocol_trace_from_EqSt : - forall s tl1 tl2, EqSt tl1 tl2 -> infinite_protocol_trace_from s tl1 -> infinite_protocol_trace_from s tl2. + Lemma infinite_valid_trace_from_EqSt : + forall s tl1 tl2, EqSt tl1 tl2 -> infinite_valid_trace_from s tl1 -> infinite_valid_trace_from s tl2. Proof. intros s tl1 tl2 Heq Htl1. - apply infinite_protocol_trace_from_prefix_rev. + apply infinite_valid_trace_from_prefix_rev. intro n. apply stream_prefix_EqSt with (n0 := n) in Heq. - apply infinite_protocol_trace_from_prefix with (n := n) in Htl1. + apply infinite_valid_trace_from_prefix with (n := n) in Htl1. rewrite <- Heq. assumption. Qed. - Lemma infinite_protocol_trace_from_segment + Lemma infinite_valid_trace_from_segment (s : state) (ls : Stream transition_item) - (Htr : infinite_protocol_trace_from s ls) + (Htr : infinite_valid_trace_from s ls) (n1 n2 : nat) (Hle : n1 <= n2) (n1th := Str_nth n1 (Cons s (Streams.map destination ls))) - : finite_protocol_trace_from n1th (stream_segment ls n1 n2). + : finite_valid_trace_from n1th (stream_segment ls n1 n2). Proof. - apply finite_protocol_trace_from_suffix with s. - - apply infinite_protocol_trace_from_prefix. assumption. + apply finite_valid_trace_from_suffix with s. + - apply infinite_valid_trace_from_prefix. assumption. - destruct n1; try reflexivity. unfold n1th. clear n1th. unfold finite_trace_nth. @@ -1914,118 +1904,118 @@ definitions, mostly reducing them to properties about their finite segments. reflexivity. Qed. -(** *** Protocol traces +(** *** valid traces Finally, we define [Trace] as a sum-type of its finite/infinite variants. It inherits some previously introduced definitions, culminating with the -[protocol_trace]. +[valid_trace]. *) - Definition ptrace_from_prop (tr : Trace) : Prop := + Definition valid_trace_from_prop (tr : Trace) : Prop := match tr with - | Finite s ls => finite_protocol_trace_from s ls - | Infinite s sm => infinite_protocol_trace_from s sm + | Finite s ls => finite_valid_trace_from s ls + | Infinite s sm => infinite_valid_trace_from s sm end. - Definition protocol_trace_prop (tr : Trace) : Prop := + Definition valid_trace_prop (tr : Trace) : Prop := match tr with - | Finite s ls => finite_protocol_trace s ls - | Infinite s sm => infinite_protocol_trace s sm + | Finite s ls => finite_valid_trace s ls + | Infinite s sm => infinite_valid_trace s sm end. - Definition protocol_trace : Type := - { tr : Trace | protocol_trace_prop tr}. + Definition valid_trace : Type := + { tr : Trace | valid_trace_prop tr}. - Lemma protocol_trace_from + Lemma valid_trace_from (tr : Trace) - (Htr : protocol_trace_prop tr) - : ptrace_from_prop tr. + (Htr : valid_trace_prop tr) + : valid_trace_from_prop tr. Proof. destruct tr; simpl; destruct Htr as [Htr Hinit]; assumption. Qed. - Lemma protocol_trace_initial + Lemma valid_trace_initial (tr : Trace) - (Htr : protocol_trace_prop tr) + (Htr : valid_trace_prop tr) : initial_state_prop (trace_first tr). Proof. destruct tr; simpl; destruct Htr as [Htr Hinit]; assumption. Qed. - Lemma protocol_trace_from_iff + Lemma valid_trace_from_iff (tr : Trace) - : protocol_trace_prop tr - <-> ptrace_from_prop tr /\ initial_state_prop (trace_first tr). + : valid_trace_prop tr + <-> valid_trace_from_prop tr /\ initial_state_prop (trace_first tr). Proof. split. - intro Htr; split. - + apply protocol_trace_from; assumption. - + apply protocol_trace_initial; assumption. + + apply valid_trace_from; assumption. + + apply valid_trace_initial; assumption. - destruct tr; simpl; intros [Htr Hinit]; split; assumption. Qed. -(** Having defined [protocol_trace]s, we now connect them to protocol states +(** Having defined [valid_trace]s, we now connect them to valid states and messages, in the following sense: for each state-message pair (<>, <>) -that has the [protocol_prop]erty, there exists a [protocol_trace] which ends +that has the [valid_state_message_prop]erty, there exists a [valid_trace] which ends in <> by outputting <> *) - Lemma protocol_is_trace + Lemma valid_state_message_has_trace (s : state) (om : option message) - (Hp : protocol_prop s om) + (Hp : valid_state_message_prop s om) : initial_state_prop s /\ option_initial_message_prop om \/ exists (is : state) (tr : list transition_item), - finite_protocol_trace_init_to is s tr + finite_valid_trace_init_to is s tr /\ finite_trace_last_output tr = om. Proof. - apply finite_protocol_trace_init_to_emit_protocol_rev in Hp as [is [tr Htr]]. - apply finite_protocol_trace_init_to_emit_output in Htr as Houtput. + apply finite_valid_trace_init_to_emit_valid_state_message_rev in Hp as [is [tr Htr]]. + apply finite_valid_trace_init_to_emit_output in Htr as Houtput. unfold empty_initial_message_or_final_output in Houtput. destruct_list_last tr tr' item Heqtr. - left. split; [|assumption]. inversion Htr; [assumption|]. destruct tl; simpl in *; congruence. - - right. apply finite_protocol_trace_init_to_emit_forget_emit in Htr. + - right. apply finite_valid_trace_init_to_emit_forget_emit in Htr. eexists _,_. split; [exact Htr|]. rewrite finite_trace_last_output_is_last. assumption. Qed. - (** Giving a trace for [protocol_state_prop] can be stated more - simply than [protocol_is_trace], because we don't need a + (** Giving a trace for [valid_state_prop] can be stated more + simply than [valid_state_message_has_trace], because we don't need a disjunction because we are not making claims about [output] messages. *) - Lemma protocol_state_has_trace + Lemma valid_state_has_trace (s : state) - (Hp : protocol_state_prop s): + (Hp : valid_state_prop s): exists (is : state) (tr : list transition_item), - finite_protocol_trace_init_to is s tr. + finite_valid_trace_init_to is s tr. Proof using. destruct Hp as [_om Hp]. - apply protocol_is_trace in Hp. + apply valid_state_message_has_trace in Hp. destruct Hp as [[Hinit _]|Htrace]. + exists s, []. split;[|assumption]. constructor. - apply initial_is_protocol. + apply initial_state_is_valid. assumption. + destruct Htrace as [is [tr [Htr _]]]. exists is, tr. assumption. Qed. - (** For any protocol transition there exists a protocol trace ending in it. *) + (** For any input valid transition there exists a valid trace ending in it. *) Lemma exists_right_finite_trace_from l s1 iom s2 oom - (Ht : protocol_transition l (s1, iom) (s2, oom)) - : exists s0 ts, finite_protocol_trace_init_to s0 s2 (ts ++ [{| l := l; destination := s2; input := iom; output := oom |}]) + (Ht : input_valid_transition l (s1, iom) (s2, oom)) + : exists s0 ts, finite_valid_trace_init_to s0 s2 (ts ++ [{| l := l; destination := s2; input := iom; output := oom |}]) /\ finite_trace_last s0 ts = s1. Proof. - apply protocol_transition_origin in Ht as Hs1. - apply protocol_state_has_trace in Hs1. + apply input_valid_transition_origin in Ht as Hs1. + apply valid_state_has_trace in Hs1. destruct Hs1 as [s0 [ts Hts]]. exists s0, ts. destruct Hts as [Hts Hinit]. - repeat split; [|assumption|revert Hts; apply finite_protocol_trace_from_to_last]. + repeat split; [|assumption|revert Hts; apply finite_valid_trace_from_to_last]. revert Ht. apply extend_right_finite_trace_from_to. assumption. @@ -2034,63 +2024,63 @@ in <> by outputting <> *) Lemma can_emit_has_trace m : can_emit m -> exists is tr item, - finite_protocol_trace is (tr ++ [item]) /\ + finite_valid_trace is (tr ++ [item]) /\ output item = Some m. Proof. intros [(s, im) [l [s' Hm]]]. apply exists_right_finite_trace_from in Hm as [is [tr [Htr _]]]. - apply finite_protocol_trace_init_to_forget_last in Htr. + apply finite_valid_trace_init_to_forget_last in Htr. eexists is, tr, _; split; [exact Htr|reflexivity]. Qed. - (** Any trace with the 'finite_protocol_trace_from' property can be completed + (** Any trace with the 'finite_valid_trace_from' property can be completed (to the left) to start in an initial state*) - Lemma finite_protocol_trace_from_complete_left + Lemma finite_valid_trace_from_complete_left (s : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from s tr) + (Htr : finite_valid_trace_from s tr) : exists (is : state) (trs : list transition_item), - finite_protocol_trace is (trs ++ tr) /\ + finite_valid_trace is (trs ++ tr) /\ finite_trace_last is trs = s. Proof. - apply finite_ptrace_first_pstate in Htr as Hs. - apply protocol_state_has_trace in Hs. + apply finite_valid_trace_first_pstate in Htr as Hs. + apply valid_state_has_trace in Hs. destruct Hs as [is [trs [Htrs His]]]. exists is, trs. - apply finite_protocol_trace_from_to_last in Htrs as Hlast. + apply finite_valid_trace_from_to_last in Htrs as Hlast. rewrite <- Hlast in Htr. - apply finite_protocol_trace_from_to_forget_last in Htrs. + apply finite_valid_trace_from_to_forget_last in Htrs. repeat (split || assumption || - apply finite_protocol_trace_from_app_iff). + apply finite_valid_trace_from_app_iff). Qed. - (** Any trace with the 'finite_protocol_trace_from_to' property can be completed + (** Any trace with the 'finite_valid_trace_from_to' property can be completed (to the left) to start in an initial state*) - Lemma finite_protocol_trace_from_to_complete_left + Lemma finite_valid_trace_from_to_complete_left (s f : state) (tr : list transition_item) - (Htr : finite_protocol_trace_from_to s f tr) + (Htr : finite_valid_trace_from_to s f tr) : exists (is : state) (trs : list transition_item), - finite_protocol_trace_init_to is f (trs ++ tr) /\ + finite_valid_trace_init_to is f (trs ++ tr) /\ finite_trace_last is trs = s. Proof. - assert (protocol_state_prop s) as Hs - by (apply finite_protocol_trace_from_to_forget_last, - finite_ptrace_first_pstate in Htr; assumption). - apply protocol_state_has_trace in Hs. + assert (valid_state_prop s) as Hs + by (apply finite_valid_trace_from_to_forget_last, + finite_valid_trace_first_pstate in Htr; assumption). + apply valid_state_has_trace in Hs. destruct Hs as [is [trs [Htrs His]]]. exists is, trs. split. - split;[|assumption]. - apply finite_protocol_trace_from_to_app with s;assumption. - - apply finite_protocol_trace_from_to_last in Htrs;assumption. + apply finite_valid_trace_from_to_app with s;assumption. + - apply finite_valid_trace_from_to_last in Htrs;assumption. Qed. (** Another benefit of defining traces is that we can succintly describe indirect transitions between arbitrary pairs of states. We say that state <> is in state <>'s futures if -there exists a finite (possibly empty) protocol trace that begins +there exists a finite (possibly empty) valid trace that begins with <> and ends in <>. This relation is often used in stating safety and liveness properties.*) @@ -2099,12 +2089,12 @@ This relation is often used in stating safety and liveness properties.*) (first second : state) : Prop := exists (tr : list transition_item), - finite_protocol_trace_from_to first second tr. + finite_valid_trace_from_to first second tr. Lemma in_futures_preserving (R : state -> state -> Prop) (Hpre : PreOrder R) - (Ht : protocol_transition_preserving R) + (Ht : input_valid_transition_preserving R) (s1 s2 : state) (Hin : in_futures s1 s2) : R s1 s2. @@ -2122,7 +2112,7 @@ This relation is often used in stating safety and liveness properties.*) Lemma in_futures_strict_preserving (R : state -> state -> Prop) (Hpre : StrictOrder R) - (Ht : protocol_transition_preserving R) + (Ht : input_valid_transition_preserving R) (s1 s2 : state) (Hin : in_futures s1 s2) (Hneq : s1 <> s2) @@ -2131,21 +2121,21 @@ This relation is often used in stating safety and liveness properties.*) apply (StrictOrder_PreOrder eq_equiv) in Hpre. - specialize (in_futures_preserving (relation_disjunction R eq) Hpre) as Hpreserve. spec Hpreserve. - + intro; intros. left. apply (Ht s3 s4 l0 om1 om2 Hprotocol). + + intro; intros. left. apply (Ht s3 s4 l0 om1 om2 Hvalid_transition). + spec Hpreserve s1 s2 Hin. destruct Hpreserve; try assumption. elim Hneq. assumption. - intros x1 x2 Heq. subst. intros y1 y2 Heq. subst. split; intro; assumption. Qed. - Lemma in_futures_protocol_fst + Lemma in_futures_valid_fst (first second : state) (Hfuture: in_futures first second) - : protocol_state_prop first. + : valid_state_prop first. Proof. destruct Hfuture as [tr Htr]. - apply finite_protocol_trace_from_to_forget_last in Htr. - apply finite_ptrace_first_pstate in Htr. + apply finite_valid_trace_from_to_forget_last in Htr. + apply finite_valid_trace_first_pstate in Htr. assumption. Qed. @@ -2153,7 +2143,7 @@ This relation is often used in stating safety and liveness properties.*) Lemma in_futures_refl (first: state) - (Hps : protocol_state_prop first) + (Hps : valid_state_prop first) : in_futures first first. Proof. @@ -2171,40 +2161,40 @@ This relation is often used in stating safety and liveness properties.*) destruct H12 as [tr12 Htr12]. destruct H23 as [tr23 Htr23]. exists (tr12 ++ tr23). - apply finite_protocol_trace_from_to_app with second;assumption. + apply finite_valid_trace_from_to_app with second;assumption. Qed. - Lemma protocol_transition_in_futures {l s im s' om} - (Ht : protocol_transition l (s, im) (s', om)) + Lemma input_valid_transition_in_futures {l s im s' om} + (Ht : input_valid_transition l (s, im) (s', om)) : in_futures s s'. Proof. - apply finite_ptrace_singleton in Ht. - apply finite_protocol_trace_from_add_last with (f := s') in Ht; [|reflexivity]. + apply finite_valid_trace_singleton in Ht. + apply finite_valid_trace_from_add_last with (f := s') in Ht; [|reflexivity]. eexists. exact Ht. Qed. Lemma in_futures_witness (first second : state) (Hfutures : in_futures first second) - : exists (tr : protocol_trace) (n1 n2 : nat), + : exists (tr : valid_trace) (n1 n2 : nat), n1 <= n2 /\ trace_nth (proj1_sig tr) n1 = Some first /\ trace_nth (proj1_sig tr) n2 = Some second. Proof. - specialize (in_futures_protocol_fst first second Hfutures); intro Hps. - apply protocol_state_has_trace in Hps. + specialize (in_futures_valid_fst first second Hfutures); intro Hps. + apply valid_state_has_trace in Hps. destruct Hps as [prefix_start [prefix_tr [Hprefix_tr Hinit]]]. destruct Hfutures as [suffix_tr Hsuffix_tr]. - specialize (finite_protocol_trace_from_to_app _ _ _ _ _ Hprefix_tr Hsuffix_tr) as Happ. - apply finite_protocol_trace_from_to_forget_last in Happ. - assert (Htr : protocol_trace_prop (Finite prefix_start (prefix_tr ++ suffix_tr))) + specialize (finite_valid_trace_from_to_app _ _ _ _ _ Hprefix_tr Hsuffix_tr) as Happ. + apply finite_valid_trace_from_to_forget_last in Happ. + assert (Htr : valid_trace_prop (Finite prefix_start (prefix_tr ++ suffix_tr))) by (split;assumption). exists (exist _ _ Htr). simpl. exists (length prefix_tr), (length prefix_tr + length suffix_tr). split;[lia|]. - apply finite_protocol_trace_from_to_last in Hprefix_tr. - apply finite_protocol_trace_from_to_last in Hsuffix_tr. + apply finite_valid_trace_from_to_last in Hprefix_tr. + apply finite_valid_trace_from_to_last in Hsuffix_tr. split. - rewrite finite_trace_nth_app1;[|lia]. rewrite finite_trace_nth_last. @@ -2224,31 +2214,31 @@ This relation is often used in stating safety and liveness properties.*) | Infinite s l => stream_segment l n1 n2 end. - Lemma ptrace_segment + Lemma valid_trace_segment (tr : Trace) - (Htr : protocol_trace_prop tr) + (Htr : valid_trace_prop tr) (n1 n2 : nat) (Hle : n1 <= n2) (first : state) (Hfirst : trace_nth tr n1 = Some first) - : finite_protocol_trace_from first (trace_segment tr n1 n2). + : finite_valid_trace_from first (trace_segment tr n1 n2). Proof. destruct tr as [s tr | s tr]; simpl in *; destruct Htr as [Htr Hinit]. - - apply finite_protocol_trace_from_segment with s; try assumption. + - apply finite_valid_trace_from_segment with s; try assumption. - inversion Hfirst; subst; clear Hfirst. - apply (infinite_protocol_trace_from_segment s tr Htr n1 n2 Hle). + apply (infinite_valid_trace_from_segment s tr Htr n1 n2 Hle). Qed. Inductive Trace_messages : Type := | Finite_messages : list (option message) -> Trace_messages | Infinite_messages : Stream (option message) -> Trace_messages. - Definition protocol_output_messages_trace (tr : protocol_trace) : Trace_messages := + Definition protocol_output_messages_trace (tr : valid_trace) : Trace_messages := match proj1_sig tr with | Finite _ ls => Finite_messages (List.map output ls) | Infinite _ st => Infinite_messages (map output st) end. - Definition protocol_input_messages_trace (tr : protocol_trace) : Trace_messages := + Definition protocol_input_messages_trace (tr : valid_trace) : Trace_messages := match proj1_sig tr with | Finite _ ls => Finite_messages (List.map input ls) | Infinite _ st => Infinite_messages (map input st) end. @@ -2273,21 +2263,21 @@ This relation is often used in stating safety and liveness properties.*) | Infinite s st => Finite s (stream_prefix st n) end. - Lemma trace_prefix_protocol - (tr : protocol_trace) + Lemma trace_prefix_valid + (tr : valid_trace) (last : transition_item) (prefix : list transition_item) (Hprefix : trace_prefix (proj1_sig tr) last prefix) - : protocol_trace_prop (Finite (trace_first (proj1_sig tr)) (prefix ++ [last])). + : valid_trace_prop (Finite (trace_first (proj1_sig tr)) (prefix ++ [last])). Proof. destruct tr as [tr Htr]. simpl in *. generalize dependent tr. generalize dependent last. - apply (rev_ind (fun prefix => forall (last : transition_item) (tr : Trace), protocol_trace_prop tr -> trace_prefix tr last prefix -> finite_protocol_trace (trace_first tr) (prefix ++ [last]))). + apply (rev_ind (fun prefix => forall (last : transition_item) (tr : Trace), valid_trace_prop tr -> trace_prefix tr last prefix -> finite_valid_trace (trace_first tr) (prefix ++ [last]))). - intros last tr Htr Hprefix; destruct tr as [ | ]; unfold trace_prefix in Hprefix; simpl in Hprefix ; destruct Hprefix as [suffix Heq]; subst; destruct Htr as [Htr Hinit] ; unfold trace_first; simpl; constructor; try assumption ; inversion Htr; subst; clear Htr - ; apply finite_ptrace_singleton; assumption. + ; apply finite_valid_trace_singleton; assumption. - intros last_p p Hind last tr Htr Hprefix. specialize (Hind last_p tr Htr). destruct tr as [ | ]; unfold trace_prefix in Hprefix; simpl in Hprefix @@ -2309,7 +2299,7 @@ This relation is often used in stating safety and liveness properties.*) rewrite <- (app_assoc p _ _) in Htr. simpl in Htr. rewrite <- app_assoc in Htr. specialize - (finite_ptrace_consecutive_valid_transition _ _ _ _ _ _ Htr eq_refl). + (finite_valid_trace_consecutive_valid_transition _ _ _ _ _ _ Htr eq_refl). simpl. rewrite finite_trace_last_is_last. trivial. + assert @@ -2326,7 +2316,7 @@ This relation is often used in stating safety and liveness properties.*) rewrite stream_app_assoc in Htr. rewrite <- (app_assoc p _ _) in Htr. simpl in Htr. specialize - (infinite_protocol_trace_consecutive_valid_transition + (infinite_valid_trace_consecutive_valid_transition s (stream_app (p ++ [last_p; {| l := l0; input := input0; destination := destination0; output := output0 |}]) suffix) suffix @@ -2341,22 +2331,22 @@ This relation is often used in stating safety and liveness properties.*) Qed. - Definition build_trace_prefix_protocol - {tr : protocol_trace} + Definition build_trace_prefix_valid + {tr : valid_trace} {last : transition_item} {prefix : list transition_item} (Hprefix : trace_prefix (proj1_sig tr) last prefix) - : protocol_trace + : valid_trace := exist _ (Finite (trace_first (proj1_sig tr)) (prefix ++ [last])) - (trace_prefix_protocol tr last prefix Hprefix). + (trace_prefix_valid tr last prefix Hprefix). - Lemma trace_prefix_fn_protocol + Lemma trace_prefix_fn_valid (tr : Trace) - (Htr : protocol_trace_prop tr) + (Htr : valid_trace_prop tr) (n : nat) - : protocol_trace_prop (trace_prefix_fn tr n). + : valid_trace_prop (trace_prefix_fn tr n). Proof. - specialize (trace_prefix_protocol (exist _ tr Htr)); simpl; intro Hpref. + specialize (trace_prefix_valid (exist _ tr Htr)); simpl; intro Hpref. remember (trace_prefix_fn tr n) as pref_tr. destruct pref_tr as [s l | s l]. - destruct l as [| item l]. @@ -2365,14 +2355,14 @@ This relation is often used in stating safety and liveness properties.*) ; inversion Heqpref_tr; subst ; (split;[|assumption]) ; constructor - ; apply initial_is_protocol;assumption. + ; apply initial_state_is_valid;assumption. + assert (Hnnil : item ::l <> []) by (intro Hnil; inversion Hnil). specialize (exists_last Hnnil); intros [prefix [last Heq]]. rewrite Heq in *; clear Hnnil Heq l item. replace s with (trace_first (proj1_sig (exist _ tr Htr))) ; try (destruct tr; inversion Heqpref_tr; subst; reflexivity). - apply trace_prefix_protocol. + apply trace_prefix_valid. remember (prefix ++ [last]) as prefix_last. revert Heqprefix_last. destruct tr as [s' l' | s' l'] ; inversion Heqpref_tr @@ -2393,41 +2383,41 @@ This relation is often used in stating safety and liveness properties.*) - destruct tr as [s' l' | s' l']; inversion Heqpref_tr. Qed. - Lemma protocol_trace_nth + Lemma valid_trace_nth (tr : Trace) - (Htr : protocol_trace_prop tr) + (Htr : valid_trace_prop tr) (n : nat) (s : state) (Hnth : trace_nth tr n = Some s) - : protocol_state_prop s. + : valid_state_prop s. Proof. destruct tr as [s0 l | s0 l]; destruct Htr as [Htr Hinit]. - - specialize (finite_protocol_trace_from_suffix s0 l Htr n s Hnth). + - specialize (finite_valid_trace_from_suffix s0 l Htr n s Hnth). intro Hsuf. - apply finite_ptrace_first_pstate in Hsuf. + apply finite_valid_trace_first_pstate in Hsuf. assumption. - assert (Hle : n <= n) by lia. - specialize (infinite_protocol_trace_from_segment s0 l Htr n n Hle) + specialize (infinite_valid_trace_from_segment s0 l Htr n n Hle) ; simpl; intros Hseg. inversion Hnth. - apply finite_ptrace_first_pstate in Hseg. + apply finite_valid_trace_first_pstate in Hseg. assumption. Qed. - Lemma in_futures_protocol_snd + Lemma in_futures_valid_snd (first second : state) (Hfutures: in_futures first second) - : protocol_state_prop second. + : valid_state_prop second. Proof. specialize (in_futures_witness first second Hfutures) ; intros [tr [n1 [n2 [Hle [Hn1 Hn2]]]]]. destruct tr as [tr Htr]; simpl in Hn2. - apply protocol_trace_nth with tr n2; assumption. + apply valid_trace_nth with tr n2; assumption. Qed. Lemma in_futures_witness_reverse (first second : state) - (tr : protocol_trace) + (tr : valid_trace) (n1 n2 : nat) (Hle : n1 <= n2) (Hs1 : trace_nth (proj1_sig tr) n1 = Some first) @@ -2439,10 +2429,10 @@ This relation is often used in stating safety and liveness properties.*) inversion Hle; subst; clear Hle. - rewrite Hs1 in Hs2. inversion Hs2; subst; clear Hs2. exists []. - constructor. apply protocol_trace_nth with tr n2; assumption. + constructor. apply valid_trace_nth with tr n2; assumption. - exists (trace_segment tr n1 (S m)). - apply finite_protocol_trace_from_add_last. - + apply ptrace_segment; try assumption. lia. + apply finite_valid_trace_from_add_last. + + apply valid_trace_segment; try assumption. lia. + { destruct tr as [s tr | s tr]; simpl. - simpl in Hs1, Hs2. unfold list_segment. @@ -2469,9 +2459,9 @@ This relation is often used in stating safety and liveness properties.*) (** Stating livness properties will require quantifying over complete executions of the protocol. To make this possible, we will now define -_complete_ [protocol_trace]s. +_complete_ [valid_trace]s. -A [protocol_trace] is _terminating_ if there's no other [protocol_trace] +A [valid_trace] is _terminating_ if there's no other [valid_trace] that contains it as a prefix. *) @@ -2479,17 +2469,17 @@ that contains it as a prefix. := match tr with | Finite s ls => - (exists (tr : protocol_trace) + (exists (tr : valid_trace) (last : transition_item), trace_prefix (proj1_sig tr) last ls) -> False | Infinite s ls => False end. -(** A [protocol_trace] is _complete_, if it is either _terminating_ or infinite. +(** A [valid_trace] is _complete_, if it is either _terminating_ or infinite. *) Definition complete_trace_prop (tr : Trace) : Prop - := protocol_trace_prop tr + := valid_trace_prop tr /\ match tr with | Finite _ _ => terminating_trace_prop tr @@ -2503,8 +2493,8 @@ that contains it as a prefix. (* Defining equivocation on these trace definitions *) (* Section 7 : - A message m received by a protocol state s with a transition label l in a - protocol execution trace is called "an equivocation" if it wasn't produced + A message m received by a valid state s with a transition label l in a + valid execution trace is called "an equivocation" if it wasn't produced in that trace *) @@ -2520,19 +2510,19 @@ that contains it as a prefix. (* end hide *) End VLSM. -(** Make all arguments of [protocol_state_prop_ind] explicit +(** Make all arguments of [valid_state_prop_ind] explicit so it will work with the <> tactic. (closing the section added <<{message}>> as an implicit argument) *) -Arguments protocol_prop_ind : clear implicits. -Arguments protocol_state_prop_ind : clear implicits. -Arguments finite_protocol_trace_from_to_ind : clear implicits. +Arguments valid_state_message_prop_ind : clear implicits. +Arguments valid_state_prop_ind : clear implicits. +Arguments finite_valid_trace_from_to_ind : clear implicits. -Arguments finite_protocol_trace_rev_ind : clear implicits. -Arguments finite_protocol_trace_from_rev_ind : clear implicits. -Arguments finite_protocol_trace_from_to_rev_ind : clear implicits. -Arguments finite_protocol_trace_init_to_rev_ind : clear implicits. -Arguments finite_protocol_trace_init_to_rev_strong_ind : clear implicits. +Arguments finite_valid_trace_rev_ind : clear implicits. +Arguments finite_valid_trace_from_rev_ind : clear implicits. +Arguments finite_valid_trace_from_to_rev_ind : clear implicits. +Arguments finite_valid_trace_init_to_rev_ind : clear implicits. +Arguments finite_valid_trace_init_to_rev_strong_ind : clear implicits. Arguments extend_right_finite_trace_from [message] (X) [s1] [ts] (Ht12) [l3] [iom3] [s3] [oom3] (Hv23). Arguments extend_right_finite_trace_from_to [message] (X) [s1] [s2] [ts] (Ht12) [l3] [iom3] [s3] [oom3] (Hv23). @@ -2542,63 +2532,63 @@ Class TraceWithLast @state _ (@type _ X) -> list transition_item -> Prop) (trace_prop : forall {message} (X: VLSM message), state -> state -> list transition_item -> Prop) := - {ptrace_add_last: forall [msg] [X: VLSM msg] [s f tr], + {valid_trace_add_last: forall [msg] [X: VLSM msg] [s f tr], base_prop X s tr -> finite_trace_last s tr = f -> trace_prop X s f tr; - ptrace_get_last: forall [msg] [X: VLSM msg] [s f tr], + valid_trace_get_last: forall [msg] [X: VLSM msg] [s f tr], trace_prop X s f tr -> finite_trace_last s tr = f; - ptrace_last_pstate: forall [msg] [X: VLSM msg] [s f tr], - trace_prop X s f tr -> protocol_state_prop X f; - ptrace_forget_last: forall [msg] [X: VLSM msg] [s f tr], + valid_trace_last_pstate: forall [msg] [X: VLSM msg] [s f tr], + trace_prop X s f tr -> valid_state_prop X f; + valid_trace_forget_last: forall [msg] [X: VLSM msg] [s f tr], trace_prop X s f tr -> base_prop X s tr }. -Hint Mode TraceWithLast - ! : typeclass_instances. -Hint Mode TraceWithLast ! - : typeclass_instances. +Global Hint Mode TraceWithLast - ! : typeclass_instances. +Global Hint Mode TraceWithLast ! - : typeclass_instances. -Definition ptrace_add_default_last +Definition valid_trace_add_default_last `{TraceWithLast base_prop trace_prop} [msg] [X:VLSM msg] [s tr] (Htr: base_prop msg X s tr): trace_prop msg X s (finite_trace_last s tr) tr. Proof. - apply ptrace_add_last. assumption. reflexivity. + apply valid_trace_add_last. assumption. reflexivity. Defined. -Instance trace_with_last_ptrace_from: - TraceWithLast (@finite_protocol_trace_from) (@finite_protocol_trace_from_to) - := {ptrace_add_last := @finite_protocol_trace_from_add_last; - ptrace_get_last := @finite_protocol_trace_from_to_last; - ptrace_last_pstate := @finite_protocol_trace_from_to_last_pstate; - ptrace_forget_last := @finite_protocol_trace_from_to_forget_last; +Instance trace_with_last_valid_trace_from: + TraceWithLast (@finite_valid_trace_from) (@finite_valid_trace_from_to) + := {valid_trace_add_last := @finite_valid_trace_from_add_last; + valid_trace_get_last := @finite_valid_trace_from_to_last; + valid_trace_last_pstate := @finite_valid_trace_from_to_last_pstate; + valid_trace_forget_last := @finite_valid_trace_from_to_forget_last; }. -Instance trace_with_last_ptrace_init: - TraceWithLast (@finite_protocol_trace) (@finite_protocol_trace_init_to) - := {ptrace_add_last := @finite_protocol_trace_init_add_last; - ptrace_get_last := @finite_protocol_trace_init_to_last; - ptrace_last_pstate _ _ _ _ _ H := ptrace_last_pstate (proj1 H); - ptrace_forget_last := @finite_protocol_trace_init_to_forget_last; +Instance trace_with_last_valid_trace_init: + TraceWithLast (@finite_valid_trace) (@finite_valid_trace_init_to) + := {valid_trace_add_last := @finite_valid_trace_init_add_last; + valid_trace_get_last := @finite_valid_trace_init_to_last; + valid_trace_last_pstate _ _ _ _ _ H := valid_trace_last_pstate (proj1 H); + valid_trace_forget_last := @finite_valid_trace_init_to_forget_last; }. Class TraceWithStart {message} {X : VLSM message} (start : @state message (type X)) (trace_prop : list (transition_item (type X)) -> Prop) := - {ptrace_first_pstate: - forall [tr], trace_prop tr -> protocol_state_prop X start + {valid_trace_first_pstate: + forall [tr], trace_prop tr -> valid_state_prop X start }. -Hint Mode TraceWithStart - - - ! : typeclass_instances. - -Instance trace_with_start_ptrace_from message (X: VLSM message) s: - TraceWithStart s (finite_protocol_trace_from X s) - := {ptrace_first_pstate := finite_ptrace_first_pstate X s}. -Instance trace_with_start_ptrace message (X: VLSM message) s: - TraceWithStart s (finite_protocol_trace X s) - := {ptrace_first_pstate tr H := ptrace_first_pstate (proj1 H)}. -Instance trace_with_start_ptrace_from_to message (X: VLSM message) s f: - TraceWithStart s (finite_protocol_trace_from_to X s f) - := {ptrace_first_pstate tr H := ptrace_first_pstate (ptrace_forget_last H)}. -Instance trace_with_start_ptrace_init_to message (X: VLSM message) s f: - TraceWithStart s (finite_protocol_trace_init_to X s f) - := {ptrace_first_pstate tr H := ptrace_first_pstate (ptrace_forget_last H)}. +Global Hint Mode TraceWithStart - - - ! : typeclass_instances. + +Instance trace_with_start_valid_trace_from message (X: VLSM message) s: + TraceWithStart s (finite_valid_trace_from X s) + := {valid_trace_first_pstate := finite_valid_trace_first_pstate X s}. +Instance trace_with_start_valid_trace message (X: VLSM message) s: + TraceWithStart s (finite_valid_trace X s) + := {valid_trace_first_pstate tr H := valid_trace_first_pstate (proj1 H)}. +Instance trace_with_start_valid_trace_from_to message (X: VLSM message) s f: + TraceWithStart s (finite_valid_trace_from_to X s f) + := {valid_trace_first_pstate tr H := valid_trace_first_pstate (valid_trace_forget_last H)}. +Instance trace_with_start_valid_trace_init_to message (X: VLSM message) s f: + TraceWithStart s (finite_valid_trace_init_to X s f) + := {valid_trace_first_pstate tr H := valid_trace_first_pstate (valid_trace_forget_last H)}. (** *** Pre-loaded VLSMs @@ -2641,7 +2631,7 @@ Byzantine fault tolerance analysis. A message which can be emitted during a protocol run of the [pre_loaded_with_all_messages_vlsm] is called a [byzantine_message], because as shown by Lemmas [byzantine_pre_loaded_with_all_messages] and [pre_loaded_with_all_messages_alt_eq], - byzantine traces for a [VLSM] are precisely the protocol traces + byzantine traces for a [VLSM] are precisely the valid traces of the [pre_loaded_with_all_messages_vlsm], hence a byzantine message is any message which a byzantine trace [can_emit]. *) @@ -2655,141 +2645,141 @@ Byzantine fault tolerance analysis. := sig byzantine_message_prop. (* begin hide *) - Lemma pre_loaded_with_all_messages_message_protocol_prop + Lemma pre_loaded_with_all_messages_message_valid_initial_state_message (om : option message) - : protocol_prop pre_loaded_with_all_messages_vlsm (proj1_sig (vs0 X)) om. + : valid_state_message_prop pre_loaded_with_all_messages_vlsm (proj1_sig (vs0 X)) om. Proof. - apply protocol_initial;[apply proj2_sig|]. + apply valid_initial_state_message;[apply proj2_sig|]. destruct om;exact I. Qed. - Lemma pre_loaded_with_all_messages_protocol_prop + Lemma pre_loaded_with_all_messages_valid_state_message_preservation (s : state) (om : option message) - (Hps : protocol_prop X s om) - : protocol_prop pre_loaded_with_all_messages_vlsm s om. + (Hps : valid_state_message_prop X s om) + : valid_state_message_prop pre_loaded_with_all_messages_vlsm s om. Proof. induction Hps. - - apply (protocol_initial pre_loaded_with_all_messages_vlsm). + - apply (valid_initial_state_message pre_loaded_with_all_messages_vlsm). assumption. destruct om;exact I. - - apply (protocol_generated pre_loaded_with_all_messages_vlsm) with s _om _s om l0; assumption. + - apply (valid_generated_state_message pre_loaded_with_all_messages_vlsm) with s _om _s om l0; assumption. Qed. - Lemma pre_loaded_with_all_messages_protocol_state_prop + Lemma pre_loaded_with_all_messages_valid_state_prop (s : state) - (Hps : protocol_state_prop X s) - : protocol_state_prop pre_loaded_with_all_messages_vlsm s. + (Hps : valid_state_prop X s) + : valid_state_prop pre_loaded_with_all_messages_vlsm s. Proof. - unfold protocol_state_prop in *. + unfold valid_state_prop in *. destruct Hps as [om Hprs]. exists om. - apply pre_loaded_with_all_messages_protocol_prop. + apply pre_loaded_with_all_messages_valid_state_message_preservation. intuition. Qed. (* end hide *) - Lemma any_message_is_protocol_in_preloaded (om: option message): - option_protocol_message_prop pre_loaded_with_all_messages_vlsm om. + Lemma any_message_is_valid_in_preloaded (om: option message): + option_valid_message_prop pre_loaded_with_all_messages_vlsm om. Proof. eexists. - apply pre_loaded_with_all_messages_message_protocol_prop. + apply pre_loaded_with_all_messages_message_valid_initial_state_message. Qed. - Inductive preloaded_protocol_state_prop : state -> Prop := - | preloaded_protocol_initial_state + Inductive preloaded_valid_state_prop : state -> Prop := + | preloaded_valid_initial_state (s:state) (Hs: initial_state_prop (VLSMSign:=pre_loaded_with_all_messages_vlsm_sig) s): - preloaded_protocol_state_prop s + preloaded_valid_state_prop s | preloaded_protocol_generated (l : label) (s : state) - (Hps : preloaded_protocol_state_prop s) + (Hps : preloaded_valid_state_prop s) (om : option message) (Hv : valid (VLSMClass:=pre_loaded_with_all_messages_vlsm_machine) l (s, om)) s' om' (Ht : transition (VLSMClass:=pre_loaded_with_all_messages_vlsm_machine) l (s, om) = (s', om')) - : preloaded_protocol_state_prop s'. + : preloaded_valid_state_prop s'. - Lemma preloaded_protocol_state_prop_iff s: - protocol_state_prop pre_loaded_with_all_messages_vlsm s - <-> preloaded_protocol_state_prop s. + Lemma preloaded_valid_state_prop_iff s: + valid_state_prop pre_loaded_with_all_messages_vlsm s + <-> preloaded_valid_state_prop s. Proof. split. - - intros [om Hproto]. - induction Hproto. - + apply preloaded_protocol_initial_state. + - intros [om Hvalid]. + induction Hvalid. + + apply preloaded_valid_initial_state. assumption. + apply preloaded_protocol_generated with l0 s om om';assumption. - induction 1. + exists None. - apply protocol_initial;[assumption|exact I]. - + exists om'. destruct IHpreloaded_protocol_state_prop as [_om Hs]. - specialize (any_message_is_protocol_in_preloaded om) as [_s Hom]. - apply (protocol_generated pre_loaded_with_all_messages_vlsm) with s _om _s om l0;assumption. + apply valid_initial_state_message;[assumption|exact I]. + + exists om'. destruct IHpreloaded_valid_state_prop as [_om Hs]. + specialize (any_message_is_valid_in_preloaded om) as [_s Hom]. + apply (valid_generated_state_message pre_loaded_with_all_messages_vlsm) with s _om _s om l0;assumption. Qed. - Lemma preloaded_weaken_protocol_prop s om: - protocol_prop X s om -> - protocol_prop pre_loaded_with_all_messages_vlsm s om. + Lemma preloaded_weaken_valid_state_message_prop s om: + valid_state_message_prop X s om -> + valid_state_message_prop pre_loaded_with_all_messages_vlsm s om. Proof. induction 1. - - refine (protocol_initial pre_loaded_with_all_messages_vlsm s Hs om _). + - refine (valid_initial_state_message pre_loaded_with_all_messages_vlsm s Hs om _). destruct om;exact I. - - exact (protocol_generated pre_loaded_with_all_messages_vlsm - _ _ IHprotocol_prop1 - _ _ IHprotocol_prop2 l0 Hv _ _ Ht). + - exact (valid_generated_state_message pre_loaded_with_all_messages_vlsm + _ _ IHvalid_state_message_prop1 + _ _ IHvalid_state_message_prop2 l0 Hv _ _ Ht). Qed. - Lemma preloaded_weaken_protocol_transition + Lemma preloaded_weaken_input_valid_transition l s om s' om': - protocol_transition X l (s,om) (s',om') -> - protocol_transition pre_loaded_with_all_messages_vlsm l (s,om) (s',om'). + input_valid_transition X l (s,om) (s',om') -> + input_valid_transition pre_loaded_with_all_messages_vlsm l (s,om) (s',om'). Proof. - unfold protocol_transition. - intros [[[_om Hproto_s] [_ Hpvalid]] Htrans]. + unfold input_valid_transition. + intros [[[_om valid_s] [_ Hvalid]] Htrans]. split;[clear Htrans|assumption]. split. - exists _om. - apply preloaded_weaken_protocol_prop. + apply preloaded_weaken_valid_state_message_prop. assumption. - - clear _om Hproto_s. + - clear _om valid_s. split. - + apply any_message_is_protocol_in_preloaded. + + apply any_message_is_valid_in_preloaded. + assumption. Qed. - Lemma preloaded_weaken_protocol_trace_from s tr - : finite_protocol_trace_from X s tr -> - finite_protocol_trace_from pre_loaded_with_all_messages_vlsm s tr. + Lemma preloaded_weaken_valid_trace_from s tr + : finite_valid_trace_from X s tr -> + finite_valid_trace_from pre_loaded_with_all_messages_vlsm s tr. Proof. - intros H. induction H using finite_protocol_trace_from_rev_ind. - - apply (finite_ptrace_empty pre_loaded_with_all_messages_vlsm). + intros H. induction H using finite_valid_trace_from_rev_ind. + - apply (finite_valid_trace_from_empty pre_loaded_with_all_messages_vlsm). destruct H as [om H]. exists om. - revert H. apply preloaded_weaken_protocol_prop. - - apply (finite_protocol_trace_from_app_iff pre_loaded_with_all_messages_vlsm). + revert H. apply preloaded_weaken_valid_state_message_prop. + - apply (finite_valid_trace_from_app_iff pre_loaded_with_all_messages_vlsm). split; [assumption|]. - apply (finite_ptrace_singleton pre_loaded_with_all_messages_vlsm). - revert Hx. apply preloaded_weaken_protocol_transition. + apply (finite_valid_trace_singleton pre_loaded_with_all_messages_vlsm). + revert Hx. apply preloaded_weaken_input_valid_transition. Qed. End pre_loaded_with_all_messages_vlsm. -Lemma non_empty_protocol_trace_from_protocol_generated_prop +Lemma non_empty_valid_trace_from_can_produce `(X : VLSM message) (s : state) (m : message) - : protocol_generated_prop X s m + : can_produce X s m <-> exists (is : state) (tr : list transition_item) (item : transition_item), - finite_protocol_trace X is tr /\ + finite_valid_trace X is tr /\ last_error tr = Some item /\ destination item = s /\ output item = Some m. Proof. split. - intros [(s', om') [l Hsm]]. destruct (id Hsm) as [[Hp _] _]. - pose proof (finite_ptrace_singleton _ Hsm) as Htr. - apply finite_protocol_trace_from_complete_left in Htr. + pose proof (finite_valid_trace_singleton _ Hsm) as Htr. + apply finite_valid_trace_from_complete_left in Htr. destruct Htr as [is [trs [Htrs _]]]. exists is. match type of Htrs with @@ -2805,7 +2795,7 @@ Proof. clear Heq. rewrite last_error_is_last in Hitem. inversion Hitem. clear Hitem. subst item'. destruct Htr as [Htr _]. - apply finite_protocol_trace_from_app_iff in Htr. + apply finite_valid_trace_from_app_iff in Htr. destruct Htr as [_ Htr]. inversion Htr. clear Htr. subst. simpl in Hm. subst. eexists _, l0. apply Ht. diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index 18523735..2602eff5 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -14,11 +14,11 @@ Section VLSM_partial_projection. (** ** VLSM partial projections A generic notion of VLSM projection. We say that VLSM X partially projects to -VLSM Y (sharing the same messages) if there exists a partial map @partial_trace_project@ +VLSM Y (sharing the same messages) if there exists a partial map <> from traces over X (pairs of state and list of transitions from that state) to traces over Y such that: -- [partial_trace_project_preserves_protocol_trace]s, if the projection is defined. +- [partial_trace_project_preserves_valid_trace]s, if the projection is defined. - The projection operation is stable to adding valid prefixes (property [partial_trace_project_extends_left]). More precisely, if the projection of a @@ -43,7 +43,7 @@ Record VLSM_partial_projection_type partial_trace_project (sX, trX) = Some (sY, trY) -> forall s'X preX, finite_trace_last s'X preX = sX -> - finite_protocol_trace_from X s'X (preX ++ trX) -> + finite_valid_trace_from X s'X (preX ++ trX) -> exists s'Y preY, partial_trace_project (s'X, preX ++ trX) = Some (s'Y, preY ++ trY) /\ finite_trace_last s'Y preY = sY @@ -62,10 +62,10 @@ Record VLSM_weak_partial_projection (partial_trace_project : vstate X * list (vtransition_item X) -> option (vstate Y * list (vtransition_item Y))) := { weak_partial_projection_type :> VLSM_partial_projection_type X Y partial_trace_project - ; weak_partial_trace_project_preserves_protocol_trace : + ; weak_partial_trace_project_preserves_valid_trace : forall sX trX sY trY, partial_trace_project (sX, trX) = Some (sY, trY) -> - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y sY trY + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y sY trY }. Record VLSM_partial_projection @@ -74,10 +74,10 @@ Record VLSM_partial_projection (partial_trace_project : vstate X * list (vtransition_item X) -> option (vstate Y * list (vtransition_item Y))) := { partial_projection_type :> VLSM_partial_projection_type X Y partial_trace_project - ; partial_trace_project_preserves_protocol_trace : + ; partial_trace_project_preserves_valid_trace : forall sX trX sY trY, partial_trace_project (sX, trX) = Some (sY, trY) -> - finite_protocol_trace X sX trX -> finite_protocol_trace Y sY trY + finite_valid_trace X sX trX -> finite_valid_trace Y sY trY }. Section weak_partial_projection_properties. @@ -89,43 +89,43 @@ Context (Hsimul : VLSM_weak_partial_projection X Y trace_project) . -Definition VLSM_weak_partial_projection_finite_protocol_trace_from +Definition VLSM_weak_partial_projection_finite_valid_trace_from : forall sX trX sY trY, trace_project (sX, trX) = Some (sY, trY) -> - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y sY trY - := weak_partial_trace_project_preserves_protocol_trace _ _ _ Hsimul. + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y sY trY + := weak_partial_trace_project_preserves_valid_trace _ _ _ Hsimul. -Lemma VLSM_weak_partial_projection_protocol_state +Lemma VLSM_weak_partial_projection_valid_state : forall sX sY trY, trace_project (sX, []) = Some (sY, trY) -> - protocol_state_prop X sX -> protocol_state_prop Y sY. + valid_state_prop X sX -> valid_state_prop Y sY. Proof. intros sX sY trY Hpr HsX. - apply protocol_state_has_trace in HsX. + apply valid_state_has_trace in HsX. destruct HsX as [isX [trX HtrX]]. - apply finite_protocol_trace_init_to_last in HtrX as HsX. - apply finite_protocol_trace_init_to_forget_last, proj1 in HtrX. + apply finite_valid_trace_init_to_last in HtrX as HsX. + apply finite_valid_trace_init_to_forget_last, proj1 in HtrX. specialize (partial_trace_project_extends_left _ _ _ Hsimul _ _ _ _ Hpr _ _ HsX) as Hpr_extends_left. spec Hpr_extends_left. { rewrite app_nil_r. assumption. } destruct Hpr_extends_left as [isY [preY [Hpr_tr HsY]]]. rewrite !app_nil_r in Hpr_tr. - specialize (VLSM_weak_partial_projection_finite_protocol_trace_from _ _ _ _ Hpr_tr HtrX) + specialize (VLSM_weak_partial_projection_finite_valid_trace_from _ _ _ _ Hpr_tr HtrX) as Hinit_to. - apply finite_protocol_trace_from_app_iff, proj1, finite_ptrace_last_pstate in Hinit_to. + apply finite_valid_trace_from_app_iff, proj1, finite_valid_trace_last_pstate in Hinit_to. subst sY. assumption. Qed. -Lemma VLSM_weak_partial_projection_protocol_transition +Lemma VLSM_weak_partial_projection_input_valid_transition : forall sX itemX sY itemY, trace_project (sX, [itemX]) = Some (sY, [itemY]) -> - protocol_transition X (l itemX) (sX, input itemX) (destination itemX, output itemX) -> - protocol_transition Y (l itemY) (sY, input itemY) (destination itemY, output itemY). + input_valid_transition X (l itemX) (sX, input itemX) (destination itemX, output itemX) -> + input_valid_transition Y (l itemY) (sY, input itemY) (destination itemY, output itemY). Proof. intros sX itemX sY itemY Hpr HtX. - apply finite_ptrace_singleton in HtX. - apply VLSM_weak_partial_projection_finite_protocol_trace_from with (sY := sY) (trY := [itemY]) in HtX + apply finite_valid_trace_singleton in HtX. + apply VLSM_weak_partial_projection_finite_valid_trace_from with (sY := sY) (trY := [itemY]) in HtX ; [|destruct itemX; assumption]. inversion HtX. subst. assumption. Qed. @@ -141,28 +141,28 @@ Context (Hsimul : VLSM_partial_projection X Y trace_project) . -Definition VLSM_partial_projection_finite_protocol_trace +Definition VLSM_partial_projection_finite_valid_trace : forall sX trX sY trY, trace_project (sX, trX) = Some (sY, trY) -> - finite_protocol_trace X sX trX -> finite_protocol_trace Y sY trY - := partial_trace_project_preserves_protocol_trace _ _ _ Hsimul. + finite_valid_trace X sX trX -> finite_valid_trace Y sY trY + := partial_trace_project_preserves_valid_trace _ _ _ Hsimul. -Lemma VLSM_partial_projection_finite_protocol_trace_from +Lemma VLSM_partial_projection_finite_valid_trace_from : forall sX trX sY trY, trace_project (sX, trX) = Some (sY, trY) -> - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y sY trY. + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y sY trY. Proof. intros sX trX sY trY Hpr_tr HtrX. - apply (finite_protocol_trace_from_complete_left X) in HtrX + apply (finite_valid_trace_from_complete_left X) in HtrX as [isX [preX [Htr'X HsX]]]. specialize (partial_trace_project_extends_left _ _ _ Hsimul _ _ _ _ Hpr_tr _ _ HsX) as Hpr_extends_left. spec Hpr_extends_left. { apply proj1 in Htr'X. assumption. } destruct Hpr_extends_left as [isY [preY [Hpr_tr' HsY]]]. - specialize (VLSM_partial_projection_finite_protocol_trace _ _ _ _ Hpr_tr' Htr'X) + specialize (VLSM_partial_projection_finite_valid_trace _ _ _ _ Hpr_tr' Htr'X) as Hinit_to. - apply proj1, finite_protocol_trace_from_app_iff, proj2 in Hinit_to. + apply proj1, finite_valid_trace_from_app_iff, proj2 in Hinit_to. subst sY. assumption. Qed. @@ -172,28 +172,28 @@ Lemma VLSM_partial_projection_initial_state vinitial_state_prop X sX -> vinitial_state_prop Y sY. Proof. intros sX sY trY Hpr HsX. - assert (HtrX : finite_protocol_trace X sX []). - { split; [|assumption]. constructor. apply initial_is_protocol. assumption. } - apply (VLSM_partial_projection_finite_protocol_trace _ _ _ _ Hpr HtrX). + assert (HtrX : finite_valid_trace X sX []). + { split; [|assumption]. constructor. apply initial_state_is_valid. assumption. } + apply (VLSM_partial_projection_finite_valid_trace _ _ _ _ Hpr HtrX). Qed. Definition VLSM_partial_projection_weaken : VLSM_weak_partial_projection X Y trace_project := {| weak_partial_projection_type := partial_projection_type _ _ _ Hsimul - ; weak_partial_trace_project_preserves_protocol_trace := VLSM_partial_projection_finite_protocol_trace_from + ; weak_partial_trace_project_preserves_valid_trace := VLSM_partial_projection_finite_valid_trace_from |}. -Definition VLSM_partial_projection_protocol_state +Definition VLSM_partial_projection_valid_state : forall sX sY trY, trace_project (sX, []) = Some (sY, trY) -> - protocol_state_prop X sX -> protocol_state_prop Y sY - := VLSM_weak_partial_projection_protocol_state VLSM_partial_projection_weaken. + valid_state_prop X sX -> valid_state_prop Y sY + := VLSM_weak_partial_projection_valid_state VLSM_partial_projection_weaken. -Definition VLSM_partial_projection_protocol_transition +Definition VLSM_partial_projection_input_valid_transition : forall sX itemX sY itemY, trace_project (sX, [itemX]) = Some (sY, [itemY]) -> - protocol_transition X (l itemX) (sX, input itemX) (destination itemX, output itemX) -> - protocol_transition Y (l itemY) (sY, input itemY) (destination itemY, output itemY) - := VLSM_weak_partial_projection_protocol_transition VLSM_partial_projection_weaken. + input_valid_transition X (l itemX) (sX, input itemX) (destination itemX, output itemX) -> + input_valid_transition Y (l itemY) (sY, input itemY) (destination itemY, output itemY) + := VLSM_weak_partial_projection_input_valid_transition VLSM_partial_projection_weaken. End partial_projection_properties. @@ -205,10 +205,10 @@ Section VLSM_projection. A VLSM projection guaranteeing the existence of projection for all states and traces. We say that VLSM X projects to VLSM Y (sharing the same messages) if -there exists maps @state_project@ taking X-states to Y-states, -and @trace_project@, taking list of transitions from X to Y, such that: +there exists maps <> taking X-states to Y-states, +and <>, taking list of transitions from X to Y, such that: -- state and [trace_project_preserves_protocol_trace]s. +- state and [trace_project_preserves_valid_trace]s. - [trace_project_app]: trace projection commutes with concatenation of traces @@ -341,7 +341,7 @@ Record VLSM_projection_type := { final_state_project : forall sX trX, - finite_protocol_trace_from X sX trX -> + finite_valid_trace_from X sX trX -> state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (trace_project trX) }. @@ -376,7 +376,7 @@ Proof. exists (state_project s'X), (trace_project preX). split. + simpl. f_equal. f_equal. apply pre_VLSM_projection_trace_project_app. + symmetry. apply (final_state_project _ _ _ _ Hsimul). - apply (finite_protocol_trace_from_app_iff X) in H1. + apply (finite_valid_trace_from_app_iff X) in H1. apply H1. Qed. @@ -399,7 +399,7 @@ after the transition must coincide. *) Definition weak_projection_transition_consistency_None : Prop := forall lX, label_project lX = None -> - forall s om s' om', protocol_transition X lX (s, om) (s', om') -> + forall s om s' om', input_valid_transition X lX (s, om) (s', om') -> state_project s' = state_project s. Definition strong_projection_transition_consistency_None : Prop := @@ -437,24 +437,24 @@ support base for [VLSM_weak_full_projection]s for which we have proper examples. *) Record VLSM_weak_projection := { weak_projection_type :> VLSM_projection_type X (type Y) label_project state_project - ; weak_trace_project_preserves_protocol_trace : + ; weak_trace_project_preserves_valid_trace : forall sX trX, - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y (state_project sX) (trace_project trX) + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y (state_project sX) (trace_project trX) }. Record VLSM_projection := { projection_type :> VLSM_projection_type X (type Y) label_project state_project - ; trace_project_preserves_protocol_trace : + ; trace_project_preserves_valid_trace : forall sX trX, - finite_protocol_trace X sX trX -> finite_protocol_trace Y (state_project sX) (trace_project trX) + finite_valid_trace X sX trX -> finite_valid_trace Y (state_project sX) (trace_project trX) }. Definition weak_projection_valid_preservation : Prop := forall lX lY (HlX : label_project lX = Some lY), forall s om - (Hv : protocol_valid X lX (s, om)) - (HsY : protocol_state_prop Y (state_project s)) - (HomY : option_protocol_message_prop Y om), + (Hv : input_valid X lX (s, om)) + (HsY : valid_state_prop Y (state_project s)) + (HomY : option_valid_message_prop Y om), vvalid Y lY ((state_project s), om). Definition strong_projection_valid_preservation : Prop := @@ -472,7 +472,7 @@ Qed. Definition weak_projection_transition_preservation_Some : Prop := forall lX lY, label_project lX = Some lY -> - forall s om s' om', protocol_transition X lX (s, om) (s', om') -> + forall s om s' om', input_valid_transition X lX (s, om) (s', om') -> vtransition Y lY (state_project s, om) = (state_project s', om'). Definition strong_projection_transition_preservation_Some : Prop := @@ -488,20 +488,20 @@ Proof. apply (Hstrong lX lY Hl). apply Ht. Qed. -Definition weak_projection_protocol_message_preservation : Prop := +Definition weak_projection_valid_message_preservation : Prop := forall lX lY (HlX : label_project lX = Some lY), forall s m - (Hv : protocol_valid X lX (s, Some m)) - (HsY : protocol_state_prop Y (state_project s)), - protocol_message_prop Y m. + (Hv : input_valid X lX (s, Some m)) + (HsY : valid_state_prop Y (state_project s)), + valid_message_prop Y m. -Definition strong_projection_protocol_message_preservation : Prop := +Definition strong_projection_valid_message_preservation : Prop := forall m : message, - protocol_message_prop X m -> protocol_message_prop Y m. + valid_message_prop X m -> valid_message_prop Y m. -Lemma strong_projection_protocol_message_preservation_weaken - : strong_projection_protocol_message_preservation -> - weak_projection_protocol_message_preservation. +Lemma strong_projection_valid_message_preservation_weaken + : strong_projection_valid_message_preservation -> + weak_projection_valid_message_preservation. Proof. intros Hstrong lX lY Hl s m [_ [Hm _]] HsY. apply Hstrong in Hm. assumption. @@ -572,22 +572,22 @@ Definition VLSM_weak_projection_trace_project_app_rev Definition VLSM_weak_projection_finite_trace_last : forall sX trX, - finite_protocol_trace_from X sX trX -> + finite_valid_trace_from X sX trX -> state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (VLSM_weak_projection_trace_project Hsimul trX) := final_state_project _ _ _ _ Hsimul. -Definition VLSM_weak_projection_finite_protocol_trace_from +Definition VLSM_weak_projection_finite_valid_trace_from : forall sX trX, - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y (state_project sX) (VLSM_weak_projection_trace_project Hsimul trX) - := weak_trace_project_preserves_protocol_trace _ _ _ _ Hsimul. + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y (state_project sX) (VLSM_weak_projection_trace_project Hsimul trX) + := weak_trace_project_preserves_valid_trace _ _ _ _ Hsimul. -Lemma VLSM_weak_projection_infinite_protocol_trace_from +Lemma VLSM_weak_projection_infinite_valid_trace_from : forall sX trX (Hinf : InfinitelyOften (VLSM_weak_projection_in Hsimul) trX), - infinite_protocol_trace_from X sX trX -> - infinite_protocol_trace_from Y (state_project sX) (VLSM_weak_projection_infinite_trace_project Hsimul trX Hinf). + infinite_valid_trace_from X sX trX -> + infinite_valid_trace_from Y (state_project sX) (VLSM_weak_projection_infinite_trace_project Hsimul trX Hinf). Proof. intros sX trX Hinf HtrX. - apply infinite_protocol_trace_from_prefix_rev. + apply infinite_valid_trace_from_prefix_rev. intros n. specialize @@ -597,20 +597,20 @@ Proof. as [m Hrew]. unfold VLSM_weak_projection_infinite_trace_project, pre_VLSM_projection_infinite_trace_project. replace (stream_prefix _ _) with (VLSM_weak_projection_trace_project Hsimul (stream_prefix trX m)). - apply VLSM_weak_projection_finite_protocol_trace_from. + apply VLSM_weak_projection_finite_valid_trace_from. - apply infinite_protocol_trace_from_prefix with (n0 := m) in HtrX. + apply infinite_valid_trace_from_prefix with (n0 := m) in HtrX. assumption. Qed. -Lemma VLSM_weak_projection_infinite_finite_protocol_trace_from +Lemma VLSM_weak_projection_infinite_finite_valid_trace_from : forall sX trX (Hfin : FinitelyManyBound (VLSM_weak_projection_in Hsimul) trX), - infinite_protocol_trace_from X sX trX -> - finite_protocol_trace_from Y (state_project sX) (VLSM_weak_projection_infinite_finite_trace_project Hsimul trX Hfin). + infinite_valid_trace_from X sX trX -> + finite_valid_trace_from Y (state_project sX) (VLSM_weak_projection_infinite_finite_trace_project Hsimul trX Hfin). Proof. intros sX trX Hfin HtrX. - apply VLSM_weak_projection_finite_protocol_trace_from. - apply infinite_protocol_trace_from_prefix with (n := `Hfin) in HtrX. + apply VLSM_weak_projection_finite_valid_trace_from. + apply infinite_valid_trace_from_prefix with (n := `Hfin) in HtrX. assumption. Qed. @@ -624,46 +624,46 @@ Proof. - apply VLSM_partial_projection_type_from_projection. apply Hsimul. - simpl. intros sX trX sY trY Heq. inversion Heq. - apply VLSM_weak_projection_finite_protocol_trace_from. + apply VLSM_weak_projection_finite_valid_trace_from. Qed. -Lemma VLSM_weak_projection_protocol_state +Lemma VLSM_weak_projection_valid_state : forall sX, - protocol_state_prop X sX -> protocol_state_prop Y (state_project sX). + valid_state_prop X sX -> valid_state_prop Y (state_project sX). Proof. specialize VLSM_weak_partial_projection_from_projection as Hpart_simul. - specialize (VLSM_weak_partial_projection_protocol_state Hpart_simul) as Hps. + specialize (VLSM_weak_partial_projection_valid_state Hpart_simul) as Hps. intro sX. eapply Hps; reflexivity. Qed. -Lemma VLSM_weak_projection_protocol_transition +Lemma VLSM_weak_projection_input_valid_transition : forall lX lY, label_project lX = Some lY -> forall s im s' om, - protocol_transition X lX (s, im) (s', om ) -> - protocol_transition Y lY (state_project s, im) (state_project s', om). + input_valid_transition X lX (s, im) (s', om ) -> + input_valid_transition Y lY (state_project s, im) (state_project s', om). Proof. specialize VLSM_weak_partial_projection_from_projection as Hpart_simul. - specialize (VLSM_weak_partial_projection_protocol_transition Hpart_simul) as Hpt. + specialize (VLSM_weak_partial_projection_input_valid_transition Hpart_simul) as Hivt. intros. apply - (Hpt s {| l := lX; input := im; destination := s'; output := om|} + (Hivt s {| l := lX; input := im; destination := s'; output := om|} (state_project s) {| l := lY; input := im; destination := state_project s'; output := om|}) ; [|assumption]. simpl. unfold pre_VLSM_projection_transition_item_project. simpl. rewrite H. reflexivity. Qed. -Lemma VLSM_weak_projection_finite_protocol_trace_from_to +Lemma VLSM_weak_projection_finite_valid_trace_from_to : forall sX s'X trX, - finite_protocol_trace_from_to X sX s'X trX -> finite_protocol_trace_from_to Y (state_project sX) (state_project s'X) (VLSM_weak_projection_trace_project Hsimul trX). + finite_valid_trace_from_to X sX s'X trX -> finite_valid_trace_from_to Y (state_project sX) (state_project s'X) (VLSM_weak_projection_trace_project Hsimul trX). Proof. specialize VLSM_weak_partial_projection_from_projection as Hpart_simul. - specialize (VLSM_weak_partial_projection_finite_protocol_trace_from Hpart_simul) as Htr. + specialize (VLSM_weak_partial_projection_finite_valid_trace_from Hpart_simul) as Htr. intros sX s'X trX HtrX. - apply ptrace_get_last in HtrX as Hs'X. - apply ptrace_forget_last in HtrX. subst. + apply valid_trace_get_last in HtrX as Hs'X. + apply valid_trace_forget_last in HtrX. subst. rewrite (final_state_project _ _ _ _ Hsimul). - - apply ptrace_add_default_last. revert HtrX. + - apply valid_trace_add_default_last. revert HtrX. apply Htr. reflexivity. - assumption. Qed. @@ -675,7 +675,7 @@ Proof. intros s1 s2 [tr Htr]. exists (VLSM_weak_projection_trace_project Hsimul tr). revert Htr. - apply VLSM_weak_projection_finite_protocol_trace_from_to. + apply VLSM_weak_projection_finite_valid_trace_from_to. Qed. End weak_projection_properties. @@ -748,14 +748,14 @@ Definition VLSM_projection_trace_project_in Definition VLSM_projection_finite_trace_last : forall sX trX, - finite_protocol_trace_from X sX trX -> + finite_valid_trace_from X sX trX -> state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (VLSM_projection_trace_project Hsimul trX) := final_state_project _ _ _ _ Hsimul. -Definition VLSM_projection_finite_protocol_trace +Definition VLSM_projection_finite_valid_trace : forall sX trX, - finite_protocol_trace X sX trX -> finite_protocol_trace Y (state_project sX) (VLSM_projection_trace_project Hsimul trX) - := trace_project_preserves_protocol_trace _ _ _ _ Hsimul. + finite_valid_trace X sX trX -> finite_valid_trace Y (state_project sX) (VLSM_projection_trace_project Hsimul trX) + := trace_project_preserves_valid_trace _ _ _ _ Hsimul. (** Any [VLSM_projection] determines a [VLSM_partial_projection], allowing us to lift to VLSM projection the generic results proved about VLSM partial projections. @@ -767,57 +767,57 @@ Proof. - apply VLSM_partial_projection_type_from_projection. apply Hsimul. - simpl. intros sX trX sY trY Heq. inversion Heq. - apply VLSM_projection_finite_protocol_trace. + apply VLSM_projection_finite_valid_trace. Qed. -Lemma VLSM_projection_finite_protocol_trace_from +Lemma VLSM_projection_finite_valid_trace_from : forall sX trX, - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y (state_project sX) (VLSM_projection_trace_project Hsimul trX). + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y (state_project sX) (VLSM_projection_trace_project Hsimul trX). Proof. specialize VLSM_partial_projection_from_projection as Hpart_simul. - specialize (VLSM_partial_projection_finite_protocol_trace_from Hpart_simul) as Hpt. + specialize (VLSM_partial_projection_finite_valid_trace_from Hpart_simul) as Hivt. intros sX trX. - apply Hpt. simpl. reflexivity. + apply Hivt. simpl. reflexivity. Qed. Definition VLSM_projection_weaken : VLSM_weak_projection X Y label_project state_project := {| weak_projection_type := projection_type _ _ _ _ Hsimul - ; weak_trace_project_preserves_protocol_trace := VLSM_projection_finite_protocol_trace_from + ; weak_trace_project_preserves_valid_trace := VLSM_projection_finite_valid_trace_from |}. -Definition VLSM_projection_protocol_state +Definition VLSM_projection_valid_state : forall sX, - protocol_state_prop X sX -> protocol_state_prop Y (state_project sX) - := VLSM_weak_projection_protocol_state VLSM_projection_weaken. + valid_state_prop X sX -> valid_state_prop Y (state_project sX) + := VLSM_weak_projection_valid_state VLSM_projection_weaken. -Definition VLSM_projection_protocol_transition +Definition VLSM_projection_input_valid_transition : forall lX lY, label_project lX = Some lY -> forall s im s' om, - protocol_transition X lX (s, im) (s', om ) -> - protocol_transition Y lY (state_project s, im) (state_project s', om) - := VLSM_weak_projection_protocol_transition VLSM_projection_weaken. + input_valid_transition X lX (s, im) (s', om ) -> + input_valid_transition Y lY (state_project s, im) (state_project s', om) + := VLSM_weak_projection_input_valid_transition VLSM_projection_weaken. -Definition VLSM_projection_finite_protocol_trace_from_to +Definition VLSM_projection_finite_valid_trace_from_to : forall sX s'X trX, - finite_protocol_trace_from_to X sX s'X trX -> finite_protocol_trace_from_to Y (state_project sX) (state_project s'X) (VLSM_projection_trace_project Hsimul trX) - := VLSM_weak_projection_finite_protocol_trace_from_to VLSM_projection_weaken. + finite_valid_trace_from_to X sX s'X trX -> finite_valid_trace_from_to Y (state_project sX) (state_project s'X) (VLSM_projection_trace_project Hsimul trX) + := VLSM_weak_projection_finite_valid_trace_from_to VLSM_projection_weaken. Definition VLSM_projection_in_futures : forall s1 s2, in_futures X s1 s2 -> in_futures Y (state_project s1) (state_project s2) := VLSM_weak_projection_in_futures VLSM_projection_weaken. -Definition VLSM_projection_infinite_protocol_trace_from +Definition VLSM_projection_infinite_valid_trace_from : forall sX trX (Hinf : InfinitelyOften (VLSM_projection_in Hsimul) trX), - infinite_protocol_trace_from X sX trX -> - infinite_protocol_trace_from Y (state_project sX) (VLSM_projection_infinite_trace_project Hsimul trX Hinf) - := VLSM_weak_projection_infinite_protocol_trace_from VLSM_projection_weaken. + infinite_valid_trace_from X sX trX -> + infinite_valid_trace_from Y (state_project sX) (VLSM_projection_infinite_trace_project Hsimul trX Hinf) + := VLSM_weak_projection_infinite_valid_trace_from VLSM_projection_weaken. -Definition VLSM_projection_infinite_finite_protocol_trace_from +Definition VLSM_projection_infinite_finite_valid_trace_from : forall sX trX (Hfin : FinitelyManyBound (VLSM_projection_in Hsimul) trX), - infinite_protocol_trace_from X sX trX -> - finite_protocol_trace_from Y (state_project sX) (VLSM_projection_infinite_finite_trace_project Hsimul trX Hfin) - := VLSM_weak_projection_infinite_finite_protocol_trace_from VLSM_projection_weaken. + infinite_valid_trace_from X sX trX -> + finite_valid_trace_from Y (state_project sX) (VLSM_projection_infinite_finite_trace_project Hsimul trX Hfin) + := VLSM_weak_projection_infinite_finite_valid_trace_from VLSM_projection_weaken. Lemma VLSM_projection_initial_state : forall sX, vinitial_state_prop X sX -> vinitial_state_prop Y (state_project sX). @@ -827,37 +827,108 @@ Proof. intro sX. eapply His; reflexivity. Qed. -Lemma VLSM_projection_finite_protocol_trace_init_to +Lemma VLSM_projection_finite_valid_trace_init_to : forall sX s'X trX, - finite_protocol_trace_init_to X sX s'X trX -> finite_protocol_trace_init_to Y (state_project sX) (state_project s'X) (VLSM_projection_trace_project Hsimul trX). + finite_valid_trace_init_to X sX s'X trX -> finite_valid_trace_init_to Y (state_project sX) (state_project s'X) (VLSM_projection_trace_project Hsimul trX). Proof. intros. destruct H as [H Hinit]. split. - - revert H. apply VLSM_projection_finite_protocol_trace_from_to. + - revert H. apply VLSM_projection_finite_valid_trace_from_to. - revert Hinit. apply VLSM_projection_initial_state. Qed. -Lemma VLSM_projection_infinite_protocol_trace +Lemma VLSM_projection_infinite_valid_trace : forall sX trX (Hinf : InfinitelyOften (VLSM_projection_in Hsimul) trX), - infinite_protocol_trace X sX trX -> - infinite_protocol_trace Y (state_project sX) (VLSM_projection_infinite_trace_project Hsimul trX Hinf). + infinite_valid_trace X sX trX -> + infinite_valid_trace Y (state_project sX) (VLSM_projection_infinite_trace_project Hsimul trX Hinf). Proof. intros sX trX Hinf [HtrX HsX]. split. - - apply VLSM_projection_infinite_protocol_trace_from. assumption. + - apply VLSM_projection_infinite_valid_trace_from. assumption. - apply VLSM_projection_initial_state. assumption. Qed. -Lemma VLSM_projection_infinite_finite_protocol_trace +Lemma VLSM_projection_infinite_finite_valid_trace : forall sX trX (Hfin : FinitelyManyBound (VLSM_projection_in Hsimul) trX), - infinite_protocol_trace X sX trX -> - finite_protocol_trace Y (state_project sX) (VLSM_projection_infinite_finite_trace_project Hsimul trX Hfin). + infinite_valid_trace X sX trX -> + finite_valid_trace Y (state_project sX) (VLSM_projection_infinite_finite_trace_project Hsimul trX Hfin). Proof. intros sX trX Hfin [HtrX HsX]. split. - - apply VLSM_projection_infinite_finite_protocol_trace_from. assumption. + - apply VLSM_projection_infinite_finite_valid_trace_from. assumption. - apply VLSM_projection_initial_state. assumption. Qed. +(** ** Projection Friendliness + +A projection is friendly if all the valid traces of the projection are +projections of the valid traces of the source VLSM. +*) + +Section projection_friendliness. + +(** We axiomatize projection friendliness as the converse of +[VLSM_projection_finite_valid_trace] *) +Definition projection_friendly_prop + := forall + (sY : vstate Y) + (trY : list (vtransition_item Y)) + (HtrY : finite_valid_trace Y sY trY), + exists (sX : vstate X) (trX : list (vtransition_item X)), + finite_valid_trace X sX trX + /\ state_project sX = sY + /\ VLSM_projection_trace_project Hsimul trX = trY. + +Lemma projection_friendly_in_futures + (Hfr : projection_friendly_prop) + (s1 s2 : vstate Y) + (Hfuture : in_futures Y s1 s2) + : exists (sX1 sX2 : vstate X), + state_project sX1 = s1 /\ state_project sX2 = s2 /\ in_futures X sX1 sX2. +Proof. + destruct Hfuture as [tr_s2 Hfuture]. + apply finite_valid_trace_from_to_complete_left in Hfuture + as [is [tr_s1 [Htr Heq_s1]]]. + apply valid_trace_get_last in Htr as Heq_s2. + apply valid_trace_forget_last in Htr. + apply Hfr in Htr as [isX [trX [Htr [His Htr_pr]]]]. + apply VLSM_projection_trace_project_app_rev in Htr_pr + as [trX_s1 [trX_s2 [HeqtrX [Htr_s1_pr Htr_s2_pr]]]]. + subst. + destruct Htr as [HtrX HisX]. + apply finite_valid_trace_from_app_iff in HtrX as HtrX12. + destruct HtrX12 as [HtrX1 HtrX2]. + apply valid_trace_add_default_last in HtrX2. + exists (finite_trace_last isX trX_s1). + exists (finite_trace_last isX (trX_s1 ++ trX_s2)). + rewrite !VLSM_projection_finite_trace_last, + VLSM_projection_trace_project_app. + - repeat split. + rewrite finite_trace_last_app. + eexists; exact HtrX2. + - assumption. + - assumption. +Qed. + +(** A consequence of the [projection_friendly_prop]erty is that the valid +traces of the projection are precisely the projections of all the valid traces +of the source VLSM. +*) +Lemma projection_friendly_trace_char + (Hfriendly : projection_friendly_prop) + : forall sY trY, finite_valid_trace Y sY trY <-> + exists (sX : vstate X) (trX : list (vtransition_item X)), + finite_valid_trace X sX trX + /\ state_project sX = sY + /\ VLSM_projection_trace_project Hsimul trX = trY. +Proof. + split; [apply Hfriendly|]. + intros [sX [trX [HtrX [<- <-]]]]. + apply VLSM_projection_finite_valid_trace. + assumption. +Qed. + +End projection_friendliness. + End projection_properties. End VLSM_projection. @@ -869,9 +940,9 @@ Section VLSM_full_projection. A VLSM projection guaranteeing the existence of projection for all labels and states, and the full correspondence between [transition_item]s. We say that VLSM X fully projects to VLSM Y (sharing the same messages) -if there exists maps @label_project@ taking X-labels to Y-labels -and @state_project@ taking X-states to Y-states, such that the -[finite_protocol_trace_prop]erty is preserved bu the trace +if there exists maps <> taking X-labels to Y-labels +and <> taking X-states to Y-states, such that the +[finite_valid_trace_prop]erty is preserved bu the trace transformation induced by the label and state projection functions, in which each X-[transition_item] is projected to an Y-[transition_item] preserving the messages and transforming labels and states accordingly. @@ -954,29 +1025,29 @@ are not required to preserve initial states. Proper examples of [VLSM_weak_full_projection] are presented in Lemmas [PreSubFree_PreFree_weak_full_projection] and [EquivPreloadedBase_Fixed_weak_full_projection], which show that a trace over -a subset of nodes can be replayed on top of a protocol state for the full +a subset of nodes can be replayed on top of a valid state for the full composition. Note that in this case, the initial state of the trace not -translated to an initial state but rather to a regular protocol state. +translated to an initial state but rather to a regular valid state. *) Record VLSM_weak_full_projection := - { weak_full_trace_project_preserves_protocol_trace : + { weak_full_trace_project_preserves_valid_trace : forall sX trX, - finite_protocol_trace_from X sX trX -> finite_protocol_trace_from Y (state_project sX) (pre_VLSM_full_projection_finite_trace_project _ _ label_project state_project trX) + finite_valid_trace_from X sX trX -> finite_valid_trace_from Y (state_project sX) (pre_VLSM_full_projection_finite_trace_project _ _ label_project state_project trX) }. Record VLSM_full_projection := - { full_trace_project_preserves_protocol_trace : + { full_trace_project_preserves_valid_trace : forall sX trX, - finite_protocol_trace X sX trX -> finite_protocol_trace Y (state_project sX) (pre_VLSM_full_projection_finite_trace_project _ _ label_project state_project trX) + finite_valid_trace X sX trX -> finite_valid_trace Y (state_project sX) (pre_VLSM_full_projection_finite_trace_project _ _ label_project state_project trX) }. Definition weak_full_projection_valid_preservation : Prop := forall (l : label) (s : state) (om : option message) - (Hv : protocol_valid X l (s, om)) - (HsY : protocol_state_prop Y (state_project s)) - (HomY : option_protocol_message_prop Y om), + (Hv : input_valid X l (s, om)) + (HsY : valid_state_prop Y (state_project s)) + (HomY : option_valid_message_prop Y om), vvalid Y (label_project l) ((state_project s), om). Lemma weak_projection_valid_preservation_from_full @@ -1009,7 +1080,7 @@ Qed. Definition weak_full_projection_transition_preservation : Prop := forall l s om s' om', - protocol_transition X l (s, om) (s', om') -> + input_valid_transition X l (s, om) (s', om') -> vtransition Y (label_project l) (state_project s, om) = (state_project s', om'). Lemma weak_projection_transition_preservation_Some_from_full @@ -1055,7 +1126,7 @@ Qed. Definition weak_full_projection_initial_state_preservation : Prop := forall s : state, - vinitial_state_prop X s -> protocol_state_prop Y (state_project s). + vinitial_state_prop X s -> valid_state_prop Y (state_project s). Definition strong_full_projection_initial_state_preservation : Prop := forall s : state, @@ -1066,15 +1137,15 @@ Lemma strong_full_projection_initial_state_preservation_weaken weak_full_projection_initial_state_preservation. Proof. intros Hstrong s Hs. apply Hstrong in Hs. - apply initial_is_protocol. assumption. + apply initial_state_is_valid. assumption. Qed. Definition weak_full_projection_initial_message_preservation : Prop := forall (l : label) (s : state) (m : message) - (Hv : protocol_valid X l (s, Some m)) - (HsY : protocol_state_prop Y (state_project s)) + (Hv : input_valid X l (s, Some m)) + (HsY : valid_state_prop Y (state_project s)) (HmX : vinitial_message_prop X m), - protocol_message_prop Y m. + valid_message_prop Y m. Definition strong_full_projection_initial_message_preservation : Prop := forall m : message, @@ -1085,7 +1156,7 @@ Lemma strong_full_projection_initial_message_preservation_weaken weak_full_projection_initial_message_preservation. Proof. intros Hstrong l s m Hv HsY Him. apply Hstrong in Him. - apply initial_message_is_protocol. assumption. + apply initial_message_is_valid. assumption. Qed. End basic_definitions. @@ -1158,12 +1229,12 @@ Definition VLSM_weak_full_projection_finite_trace_last state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (VLSM_weak_full_projection_finite_trace_project Hsimul trX) := pre_VLSM_full_projection_finite_trace_last _ _ label_project state_project. -Definition VLSM_weak_full_projection_finite_protocol_trace_from +Definition VLSM_weak_full_projection_finite_valid_trace_from : forall s tr, - finite_protocol_trace_from X s tr -> - finite_protocol_trace_from Y (state_project s) (VLSM_weak_full_projection_finite_trace_project Hsimul tr) + finite_valid_trace_from X s tr -> + finite_valid_trace_from Y (state_project s) (VLSM_weak_full_projection_finite_trace_project Hsimul tr) := - (weak_full_trace_project_preserves_protocol_trace _ _ _ _ Hsimul). + (weak_full_trace_project_preserves_valid_trace _ _ _ _ Hsimul). (** Any [VLSM_full_projection] determines a [VLSM_projection], allowing us to lift to VLSM full projections the generic results proved about VLSM projections. @@ -1173,70 +1244,70 @@ Lemma VLSM_weak_full_projection_is_projection Proof. split. - apply VLSM_full_projection_projection_type. - - apply VLSM_weak_full_projection_finite_protocol_trace_from. + - apply VLSM_weak_full_projection_finite_valid_trace_from. Qed. -Definition VLSM_weak_full_projection_protocol_state - : forall (s : vstate X) (Hs : protocol_state_prop X s), protocol_state_prop Y (state_project s) - := VLSM_weak_projection_protocol_state VLSM_weak_full_projection_is_projection. +Definition VLSM_weak_full_projection_valid_state + : forall (s : vstate X) (Hs : valid_state_prop X s), valid_state_prop Y (state_project s) + := VLSM_weak_projection_valid_state VLSM_weak_full_projection_is_projection. -Definition VLSM_weak_full_projection_finite_protocol_trace_from_to +Definition VLSM_weak_full_projection_finite_valid_trace_from_to : forall (s f : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from_to X s f tr), - finite_protocol_trace_from_to Y (state_project s) (state_project f) (VLSM_weak_full_projection_finite_trace_project Hsimul tr) - := VLSM_weak_projection_finite_protocol_trace_from_to VLSM_weak_full_projection_is_projection. + (Htr : finite_valid_trace_from_to X s f tr), + finite_valid_trace_from_to Y (state_project s) (state_project f) (VLSM_weak_full_projection_finite_trace_project Hsimul tr) + := VLSM_weak_projection_finite_valid_trace_from_to VLSM_weak_full_projection_is_projection. Definition VLSM_weak_full_projection_in_futures : forall (s1 s2 : vstate X), in_futures X s1 s2 -> in_futures Y (state_project s1) (state_project s2) := VLSM_weak_projection_in_futures VLSM_weak_full_projection_is_projection. -Lemma VLSM_weak_full_projection_protocol_transition +Lemma VLSM_weak_full_projection_input_valid_transition : forall l s im s' om, - protocol_transition X l (s,im) (s',om) -> - protocol_transition Y (label_project l) (state_project s,im) (state_project s',om). + input_valid_transition X l (s,im) (s',om) -> + input_valid_transition Y (label_project l) (state_project s,im) (state_project s',om). Proof. intros. - apply (VLSM_weak_projection_protocol_transition VLSM_weak_full_projection_is_projection) + apply (VLSM_weak_projection_input_valid_transition VLSM_weak_full_projection_is_projection) with (lY := label_project l) in H ; [assumption|reflexivity]. Qed. -Lemma VLSM_weak_full_projection_infinite_protocol_trace_from +Lemma VLSM_weak_full_projection_infinite_valid_trace_from : forall sX trX, - infinite_protocol_trace_from X sX trX -> - infinite_protocol_trace_from Y (state_project sX) (VLSM_weak_full_projection_infinite_trace_project Hsimul trX). + infinite_valid_trace_from X sX trX -> + infinite_valid_trace_from Y (state_project sX) (VLSM_weak_full_projection_infinite_trace_project Hsimul trX). Proof. intros. specialize (pre_VLSM_full_projection_infinite_trace_project_EqSt _ _ label_project state_project trX) as Heq. apply Streams.sym_EqSt in Heq. - apply (infinite_protocol_trace_from_EqSt Y _ _ _ Heq). - apply (VLSM_weak_projection_infinite_protocol_trace_from VLSM_weak_full_projection_is_projection sX trX). + apply (infinite_valid_trace_from_EqSt Y _ _ _ Heq). + apply (VLSM_weak_projection_infinite_valid_trace_from VLSM_weak_full_projection_is_projection sX trX). assumption. Qed. -Lemma VLSM_weak_full_projection_protocol_valid +Lemma VLSM_weak_full_projection_input_valid : forall l s im, - protocol_valid X l (s,im) -> - protocol_valid Y (label_project l) (state_project s,im). + input_valid X l (s,im) -> + input_valid Y (label_project l) (state_project s,im). Proof. intros. destruct (vtransition X l (s, im)) as (s', om) eqn:Ht. - assert (Hpt : protocol_transition X l (s, im) (s', om)). + assert (Hivt : input_valid_transition X l (s, im) (s', om)). { split; assumption. } - apply VLSM_weak_full_projection_protocol_transition in Hpt. - apply Hpt. + apply VLSM_weak_full_projection_input_valid_transition in Hivt. + apply Hivt. Qed. -Lemma VLSM_weak_full_projection_protocol_generated +Lemma VLSM_weak_full_projection_can_produce (s : state) (om : option message) - : option_protocol_generated_prop X s om -> option_protocol_generated_prop Y (state_project s) om. + : option_can_produce X s om -> option_can_produce Y (state_project s) om. Proof. intros [(s0, im) [l Ht]]. - apply VLSM_weak_full_projection_protocol_transition in Ht. + apply VLSM_weak_full_projection_input_valid_transition in Ht. eexists. eexists. exact Ht. Qed. @@ -1246,18 +1317,18 @@ Lemma VLSM_weak_full_projection_can_emit Proof. repeat rewrite can_emit_iff. intros [s Hsm]. exists (state_project s). revert Hsm. - apply VLSM_weak_full_projection_protocol_generated. + apply VLSM_weak_full_projection_can_produce. Qed. -Lemma VLSM_weak_full_projection_protocol_message - (Hinitial_protocol_message : strong_full_projection_initial_message_preservation X Y) +Lemma VLSM_weak_full_projection_valid_message + (Hinitial_valid_message : strong_full_projection_initial_message_preservation X Y) (m : message) - : protocol_message_prop X m -> protocol_message_prop Y m. + : valid_message_prop X m -> valid_message_prop Y m. Proof. intros Hm. - apply can_emit_protocol_iff in Hm as [Hinit | Hemit]. - - apply Hinitial_protocol_message in Hinit. apply initial_message_is_protocol. assumption. - - apply can_emit_protocol. revert Hemit. apply VLSM_weak_full_projection_can_emit. + apply emitted_messages_are_valid_iff in Hm as [Hinit | Hemit]. + - apply Hinitial_valid_message in Hinit. apply initial_message_is_valid. assumption. + - apply emitted_messages_are_valid. revert Hemit. apply VLSM_weak_full_projection_can_emit. Qed. End weak_projection_properties. @@ -1277,11 +1348,11 @@ Definition VLSM_full_projection_finite_trace_last state_project (finite_trace_last sX trX) = finite_trace_last (state_project sX) (VLSM_full_projection_finite_trace_project Hsimul trX) := pre_VLSM_full_projection_finite_trace_last _ _ label_project state_project. -Definition VLSM_full_projection_finite_protocol_trace +Definition VLSM_full_projection_finite_valid_trace : forall s tr, - finite_protocol_trace X s tr -> - finite_protocol_trace Y (state_project s) (VLSM_full_projection_finite_trace_project Hsimul tr) - := full_trace_project_preserves_protocol_trace _ _ _ _ Hsimul. + finite_valid_trace X s tr -> + finite_valid_trace Y (state_project s) (VLSM_full_projection_finite_trace_project Hsimul tr) + := full_trace_project_preserves_valid_trace _ _ _ _ Hsimul. (** Any [VLSM_full_projection] determines a [VLSM_projection], allowing us to lift to VLSM full projections the generic results proved about VLSM projections. @@ -1291,24 +1362,24 @@ Lemma VLSM_full_projection_is_projection Proof. split. - apply VLSM_full_projection_projection_type. - - apply VLSM_full_projection_finite_protocol_trace. + - apply VLSM_full_projection_finite_valid_trace. Qed. -Definition VLSM_full_projection_finite_protocol_trace_from +Definition VLSM_full_projection_finite_valid_trace_from : forall (s : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from X s tr), - finite_protocol_trace_from Y (state_project s) (VLSM_full_projection_finite_trace_project Hsimul tr) - := VLSM_projection_finite_protocol_trace_from VLSM_full_projection_is_projection. + (Htr : finite_valid_trace_from X s tr), + finite_valid_trace_from Y (state_project s) (VLSM_full_projection_finite_trace_project Hsimul tr) + := VLSM_projection_finite_valid_trace_from VLSM_full_projection_is_projection. -Definition VLSM_full_projection_finite_protocol_trace_init_to +Definition VLSM_full_projection_finite_valid_trace_init_to : forall (s f : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_init_to X s f tr), - finite_protocol_trace_init_to Y (state_project s) (state_project f) (VLSM_full_projection_finite_trace_project Hsimul tr) - := VLSM_projection_finite_protocol_trace_init_to VLSM_full_projection_is_projection. + (Htr : finite_valid_trace_init_to X s f tr), + finite_valid_trace_init_to Y (state_project s) (state_project f) (VLSM_full_projection_finite_trace_project Hsimul tr) + := VLSM_projection_finite_valid_trace_init_to VLSM_full_projection_is_projection. Definition VLSM_full_projection_initial_state : forall (is : vstate X), @@ -1318,54 +1389,54 @@ Definition VLSM_full_projection_initial_state Lemma VLSM_full_projection_weaken : VLSM_weak_full_projection X Y label_project state_project. Proof. - constructor. apply VLSM_full_projection_finite_protocol_trace_from. + constructor. apply VLSM_full_projection_finite_valid_trace_from. Qed. -Definition VLSM_full_projection_protocol_state - : forall (s : vstate X) (Hs : protocol_state_prop X s), protocol_state_prop Y (state_project s) - := VLSM_weak_full_projection_protocol_state VLSM_full_projection_weaken. +Definition VLSM_full_projection_valid_state + : forall (s : vstate X) (Hs : valid_state_prop X s), valid_state_prop Y (state_project s) + := VLSM_weak_full_projection_valid_state VLSM_full_projection_weaken. -Definition VLSM_full_projection_finite_protocol_trace_from_to +Definition VLSM_full_projection_finite_valid_trace_from_to : forall (s f : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from_to X s f tr), - finite_protocol_trace_from_to Y (state_project s) (state_project f) (VLSM_full_projection_finite_trace_project Hsimul tr) - := VLSM_weak_full_projection_finite_protocol_trace_from_to VLSM_full_projection_weaken. + (Htr : finite_valid_trace_from_to X s f tr), + finite_valid_trace_from_to Y (state_project s) (state_project f) (VLSM_full_projection_finite_trace_project Hsimul tr) + := VLSM_weak_full_projection_finite_valid_trace_from_to VLSM_full_projection_weaken. Definition VLSM_full_projection_in_futures : forall (s1 s2 : vstate X), in_futures X s1 s2 -> in_futures Y (state_project s1) (state_project s2) := VLSM_weak_full_projection_in_futures VLSM_full_projection_weaken. -Definition VLSM_full_projection_protocol_transition +Definition VLSM_full_projection_input_valid_transition : forall l s im s' om, - protocol_transition X l (s,im) (s',om) -> - protocol_transition Y (label_project l) (state_project s,im) (state_project s',om) - := VLSM_weak_full_projection_protocol_transition VLSM_full_projection_weaken. + input_valid_transition X l (s,im) (s',om) -> + input_valid_transition Y (label_project l) (state_project s,im) (state_project s',om) + := VLSM_weak_full_projection_input_valid_transition VLSM_full_projection_weaken. -Definition VLSM_full_projection_protocol_valid +Definition VLSM_full_projection_input_valid : forall l s im, - protocol_valid X l (s,im) -> - protocol_valid Y (label_project l) (state_project s,im) - := VLSM_weak_full_projection_protocol_valid VLSM_full_projection_weaken. + input_valid X l (s,im) -> + input_valid Y (label_project l) (state_project s,im) + := VLSM_weak_full_projection_input_valid VLSM_full_projection_weaken. -Definition VLSM_full_projection_protocol_generated +Definition VLSM_full_projection_can_produce : forall (s : state) (om : option message), - option_protocol_generated_prop X s om -> option_protocol_generated_prop Y (state_project s) om - := VLSM_weak_full_projection_protocol_generated VLSM_full_projection_weaken. + option_can_produce X s om -> option_can_produce Y (state_project s) om + := VLSM_weak_full_projection_can_produce VLSM_full_projection_weaken. Definition VLSM_full_projection_can_emit : forall (m : message), can_emit X m -> can_emit Y m := VLSM_weak_full_projection_can_emit VLSM_full_projection_weaken. -Definition VLSM_full_projection_protocol_message - (Hinitial_protocol_message : strong_full_projection_initial_message_preservation X Y) +Definition VLSM_full_projection_valid_message + (Hinitial_valid_message : strong_full_projection_initial_message_preservation X Y) : forall (m : message), - protocol_message_prop X m -> protocol_message_prop Y m - := VLSM_weak_full_projection_protocol_message VLSM_full_projection_weaken Hinitial_protocol_message. + valid_message_prop X m -> valid_message_prop Y m + := VLSM_weak_full_projection_valid_message VLSM_full_projection_weaken Hinitial_valid_message. Definition VLSM_full_projection_trace_project (t : vTrace X) : vTrace Y := match t with @@ -1373,47 +1444,47 @@ Definition VLSM_full_projection_trace_project (t : vTrace X) : vTrace Y := | Infinite s tr => Infinite (state_project s) (VLSM_full_projection_infinite_trace_project Hsimul tr) end. -Definition VLSM_full_projection_infinite_protocol_trace_from +Definition VLSM_full_projection_infinite_valid_trace_from s ls - : infinite_protocol_trace_from X s ls -> - infinite_protocol_trace_from Y (state_project s) (VLSM_full_projection_infinite_trace_project Hsimul ls) - := VLSM_weak_full_projection_infinite_protocol_trace_from VLSM_full_projection_weaken s ls. + : infinite_valid_trace_from X s ls -> + infinite_valid_trace_from Y (state_project s) (VLSM_full_projection_infinite_trace_project Hsimul ls) + := VLSM_weak_full_projection_infinite_valid_trace_from VLSM_full_projection_weaken s ls. -Lemma VLSM_full_projection_infinite_protocol_trace +Lemma VLSM_full_projection_infinite_valid_trace s ls - : infinite_protocol_trace X s ls -> - infinite_protocol_trace Y (state_project s) (VLSM_full_projection_infinite_trace_project Hsimul ls). + : infinite_valid_trace X s ls -> + infinite_valid_trace Y (state_project s) (VLSM_full_projection_infinite_trace_project Hsimul ls). Proof. intros [Htr His]. split. - - revert Htr. apply VLSM_full_projection_infinite_protocol_trace_from. + - revert Htr. apply VLSM_full_projection_infinite_valid_trace_from. - revert His. apply VLSM_full_projection_initial_state. Qed. -Lemma VLSM_full_projection_protocol_trace +Lemma VLSM_full_projection_valid_trace : forall t, - protocol_trace_prop X t -> - protocol_trace_prop Y (VLSM_full_projection_trace_project t). + valid_trace_prop X t -> + valid_trace_prop Y (VLSM_full_projection_trace_project t). Proof. intros [s tr | s tr]; simpl. - - apply VLSM_full_projection_finite_protocol_trace. - - apply VLSM_full_projection_infinite_protocol_trace. + - apply VLSM_full_projection_finite_valid_trace. + - apply VLSM_full_projection_infinite_valid_trace. Qed. (** - [VLSM_full_projection] almost implies inclusion of the [protocol_prop] sets. + [VLSM_full_projection] almost implies inclusion of the [valid_state_message_prop] sets. Some additional hypotheses are required because [VLSM_full_projection] only - refers to traces, and [protocol_initial] means that - [protocol_prop] includes some pairs that do not appear in any + refers to traces, and [valid_initial_state_message] means that + [valid_state_message_prop] includes some pairs that do not appear in any transition. *) -Lemma VLSM_full_projection_protocol_prop +Lemma VLSM_full_projection_valid_state_message (Hmessage : strong_full_projection_initial_message_preservation X Y) - : forall s om, protocol_prop X s om -> protocol_prop Y (state_project s) om. + : forall s om, valid_state_message_prop X s om -> valid_state_message_prop Y (state_project s) om. Proof. intros s om Hsom. - apply option_protocol_generated_prop_protocol_iff. - apply option_protocol_generated_prop_protocol_iff in Hsom as [Hgen | [His Him]]. - - left. revert Hgen. apply VLSM_full_projection_protocol_generated. + apply option_can_produce_valid_iff. + apply option_can_produce_valid_iff in Hsom as [Hgen | [His Him]]. + - left. revert Hgen. apply VLSM_full_projection_can_produce. - right. split. + revert His. apply VLSM_full_projection_initial_state. + destruct om as [m|]; [|exact I]. @@ -1429,9 +1500,9 @@ End VLSM_full_projection. We can also define VLSM _inclusion_ and _equality_ in terms of traces. When both VLSMs have the same state and label types they also share the same [Trace] type, and sets of traces can be compared without conversion. -- VLSM X is _included_ in VLSM Y if every [protocol_trace] available to X +- VLSM X is _included_ in VLSM Y if every [valid_trace] available to X is also available to Y. -- VLSM X and VLSM Y are _equal_ if their [protocol_trace]s are exactly the same. +- VLSM X and VLSM Y are _equal_ if their [valid_trace]s are exactly the same. *) Section VLSM_equality. @@ -1446,7 +1517,7 @@ Definition VLSM_eq_part (X := mk_vlsm MX) (Y := mk_vlsm MY) := forall t : Trace, - protocol_trace_prop X t <-> protocol_trace_prop Y t . + valid_trace_prop X t <-> valid_trace_prop Y t . Local Notation VLSM_eq X Y := (VLSM_eq_part (machine X) (machine Y)). Definition VLSM_incl_part @@ -1455,7 +1526,7 @@ Definition VLSM_incl_part (X := mk_vlsm MX) (Y := mk_vlsm MY) := forall t : Trace, - protocol_trace_prop X t -> protocol_trace_prop Y t. + valid_trace_prop X t -> valid_trace_prop Y t. Local Notation VLSM_incl X Y := (VLSM_incl_part (machine X) (machine Y)). Lemma VLSM_incl_refl @@ -1527,7 +1598,7 @@ Lemma VLSM_incl_finite_traces_characterization : VLSM_incl X Y <-> forall (s : vstate X) (tr : list (vtransition_item X)), - finite_protocol_trace X s tr -> finite_protocol_trace Y s tr. + finite_valid_trace X s tr -> finite_valid_trace Y s tr. Proof. split; intros Hincl. - intros. specialize (Hincl (Finite s tr)). apply Hincl. assumption. @@ -1535,16 +1606,16 @@ Proof. destruct tr as [is tr | is tr]; simpl in *. + revert Htr. apply Hincl. + destruct Htr as [HtrX HisX]. - assert (His_tr: finite_protocol_trace X is []). + assert (His_tr: finite_valid_trace X is []). { split; [|assumption]. constructor. - apply initial_is_protocol. assumption. + apply initial_state_is_valid. assumption. } apply Hincl in His_tr. destruct His_tr as [_ HisY]. split; [|assumption]. - apply infinite_protocol_trace_from_prefix_rev. + apply infinite_valid_trace_from_prefix_rev. intros. - apply infinite_protocol_trace_from_prefix with (n0 := n) in HtrX. + apply infinite_valid_trace_from_prefix with (n0 := n) in HtrX. apply (Hincl _ _ (conj HtrX HisX)). Qed. @@ -1565,7 +1636,7 @@ Proof. replace (pre_VLSM_full_projection_finite_trace_project _ _ _ _ trX) with trX; [assumption|]. apply Hid. - intro Hproject. apply VLSM_incl_finite_traces_characterization. - intros. apply (VLSM_full_projection_finite_protocol_trace Hproject) in H. + intros. apply (VLSM_full_projection_finite_valid_trace Hproject) in H. replace (VLSM_full_projection_finite_trace_project Hproject _) with tr in H; [assumption|]. apply Hid. Qed. @@ -1676,36 +1747,36 @@ Context (** VLSM inclusion specialized to finite trace. *) -Lemma VLSM_incl_finite_protocol_trace +Lemma VLSM_incl_finite_valid_trace (s : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace X s tr) - : finite_protocol_trace Y s tr. + (Htr : finite_valid_trace X s tr) + : finite_valid_trace Y s tr. Proof. - apply (VLSM_full_projection_finite_protocol_trace (VLSM_incl_is_full_projection Hincl)) + apply (VLSM_full_projection_finite_valid_trace (VLSM_incl_is_full_projection Hincl)) in Htr. rewrite (VLSM_incl_is_full_projection_finite_trace_project Hincl) in Htr. assumption. Qed. -Lemma VLSM_incl_finite_protocol_trace_init_to +Lemma VLSM_incl_finite_valid_trace_init_to (s f : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_init_to X s f tr) - : finite_protocol_trace_init_to Y s f tr. + (Htr : finite_valid_trace_init_to X s f tr) + : finite_valid_trace_init_to Y s f tr. Proof. - apply (VLSM_full_projection_finite_protocol_trace_init_to (VLSM_incl_is_full_projection Hincl)) + apply (VLSM_full_projection_finite_valid_trace_init_to (VLSM_incl_is_full_projection Hincl)) in Htr. rewrite (VLSM_incl_is_full_projection_finite_trace_project Hincl) in Htr. assumption. Qed. -Lemma VLSM_incl_protocol_state +Lemma VLSM_incl_valid_state (s : vstate X) - (Hs : protocol_state_prop X s) - : protocol_state_prop Y s. + (Hs : valid_state_prop X s) + : valid_state_prop Y s. Proof. - revert Hs. apply (VLSM_full_projection_protocol_state (VLSM_incl_is_full_projection Hincl)). + revert Hs. apply (VLSM_full_projection_valid_state (VLSM_incl_is_full_projection Hincl)). Qed. Lemma VLSM_incl_initial_state @@ -1715,25 +1786,25 @@ Proof. apply (VLSM_full_projection_initial_state (VLSM_incl_is_full_projection Hincl)). Qed. -Lemma VLSM_incl_finite_protocol_trace_from +Lemma VLSM_incl_finite_valid_trace_from (s : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from X s tr) - : finite_protocol_trace_from Y s tr. + (Htr : finite_valid_trace_from X s tr) + : finite_valid_trace_from Y s tr. Proof. - apply (VLSM_full_projection_finite_protocol_trace_from (VLSM_incl_is_full_projection Hincl)) + apply (VLSM_full_projection_finite_valid_trace_from (VLSM_incl_is_full_projection Hincl)) in Htr. rewrite (VLSM_incl_is_full_projection_finite_trace_project Hincl) in Htr. assumption. Qed. -Lemma VLSM_incl_finite_protocol_trace_from_to +Lemma VLSM_incl_finite_valid_trace_from_to (s f : vstate X) (tr : list (vtransition_item X)) - (Htr : finite_protocol_trace_from_to X s f tr) - : finite_protocol_trace_from_to Y s f tr. + (Htr : finite_valid_trace_from_to X s f tr) + : finite_valid_trace_from_to Y s f tr. Proof. - apply (VLSM_full_projection_finite_protocol_trace_from_to (VLSM_incl_is_full_projection Hincl)) + apply (VLSM_full_projection_finite_valid_trace_from_to (VLSM_incl_is_full_projection Hincl)) in Htr. rewrite (VLSM_incl_is_full_projection_finite_trace_project Hincl) in Htr. assumption. @@ -1746,46 +1817,46 @@ Proof. apply (VLSM_full_projection_in_futures (VLSM_incl_is_full_projection Hincl)). Qed. -Lemma VLSM_incl_protocol_transition +Lemma VLSM_incl_input_valid_transition : forall l s im s' om, - protocol_transition X l (s,im) (s',om) -> - protocol_transition Y l (s,im) (s',om). + input_valid_transition X l (s,im) (s',om) -> + input_valid_transition Y l (s,im) (s',om). Proof. apply - (VLSM_full_projection_protocol_transition (VLSM_incl_is_full_projection Hincl)). + (VLSM_full_projection_input_valid_transition (VLSM_incl_is_full_projection Hincl)). Qed. -Lemma VLSM_incl_protocol_valid +Lemma VLSM_incl_input_valid : forall l s im, - protocol_valid X l (s,im) -> - protocol_valid Y l (s,im). + input_valid X l (s,im) -> + input_valid Y l (s,im). Proof. apply - (VLSM_full_projection_protocol_valid (VLSM_incl_is_full_projection Hincl)). + (VLSM_full_projection_input_valid (VLSM_incl_is_full_projection Hincl)). Qed. (** - [VLSM_incl] almost implies inclusion of the [protocol_prop] sets. + [VLSM_incl] almost implies inclusion of the [valid_state_message_prop] sets. Some additional hypotheses are required because [VLSM_incl] only - refers to traces, and [protocol_initial] means that - [protocol_prop] includes some pairs that do not appear in any + refers to traces, and [valid_initial_state_message] means that + [valid_state_message_prop] includes some pairs that do not appear in any transition. *) -Lemma VLSM_incl_protocol_prop +Lemma VLSM_incl_valid_state_message (Hmessage : strong_incl_initial_message_preservation MX MY) - : forall s om, protocol_prop X s om -> protocol_prop Y s om. + : forall s om, valid_state_message_prop X s om -> valid_state_message_prop Y s om. Proof. intros s om. - apply (VLSM_full_projection_protocol_prop (VLSM_incl_is_full_projection Hincl)). + apply (VLSM_full_projection_valid_state_message (VLSM_incl_is_full_projection Hincl)). assumption. Qed. -Lemma VLSM_incl_protocol_generated +Lemma VLSM_incl_can_produce (s : state) (om : option message) - : option_protocol_generated_prop X s om -> option_protocol_generated_prop Y s om. + : option_can_produce X s om -> option_can_produce Y s om. Proof. - apply (VLSM_full_projection_protocol_generated (VLSM_incl_is_full_projection Hincl)). + apply (VLSM_full_projection_can_produce (VLSM_incl_is_full_projection Hincl)). Qed. Lemma VLSM_incl_can_emit @@ -1795,46 +1866,46 @@ Proof. apply (VLSM_full_projection_can_emit (VLSM_incl_is_full_projection Hincl)). Qed. -Definition VLSM_incl_protocol_message - (Hinitial_protocol_message : strong_incl_initial_message_preservation MX MY) +Definition VLSM_incl_valid_message + (Hinitial_valid_message : strong_incl_initial_message_preservation MX MY) : forall (m : message), - protocol_message_prop X m -> protocol_message_prop Y m. + valid_message_prop X m -> valid_message_prop Y m. Proof. intros m [s Hm]. - exists s. revert Hm. apply VLSM_incl_protocol_prop. + exists s. revert Hm. apply VLSM_incl_valid_state_message. assumption. Qed. -Lemma VLSM_incl_infinite_protocol_trace_from +Lemma VLSM_incl_infinite_valid_trace_from s ls - : infinite_protocol_trace_from X s ls -> - infinite_protocol_trace_from Y s ls. + : infinite_valid_trace_from X s ls -> + infinite_valid_trace_from Y s ls. Proof. intros Hls. - apply (VLSM_full_projection_infinite_protocol_trace_from (VLSM_incl_is_full_projection Hincl)) in Hls. + apply (VLSM_full_projection_infinite_valid_trace_from (VLSM_incl_is_full_projection Hincl)) in Hls. revert Hls. - apply infinite_protocol_trace_from_EqSt. + apply infinite_valid_trace_from_EqSt. apply Streams.ntheq_eqst. unfold VLSM_full_projection_infinite_trace_project, pre_VLSM_full_projection_infinite_trace_project. intro n. rewrite Streams.Str_nth_map. destruct (Streams.Str_nth _ _). reflexivity. Qed. -Lemma VLSM_incl_infinite_protocol_trace +Lemma VLSM_incl_infinite_valid_trace s ls - : infinite_protocol_trace X s ls -> infinite_protocol_trace Y s ls. + : infinite_valid_trace X s ls -> infinite_valid_trace Y s ls. Proof. intros [Htr His]. split. - - revert Htr. apply VLSM_incl_infinite_protocol_trace_from. + - revert Htr. apply VLSM_incl_infinite_valid_trace_from. - revert His. apply VLSM_incl_initial_state. Qed. -Lemma VLSM_incl_protocol_trace - : forall t, protocol_trace_prop X t -> protocol_trace_prop Y t. +Lemma VLSM_incl_valid_trace + : forall t, valid_trace_prop X t -> valid_trace_prop Y t. Proof. intros [s tr | s tr]; simpl. - - apply VLSM_incl_finite_protocol_trace. - - apply VLSM_incl_infinite_protocol_trace. + - apply VLSM_incl_finite_valid_trace. + - apply VLSM_incl_infinite_valid_trace. Qed. End VLSM_incl_properties. @@ -1846,7 +1917,7 @@ Lemma vlsm_incl_pre_loaded_with_all_messages_vlsm Proof. apply VLSM_incl_finite_traces_characterization. intros. split; [|apply H]. - apply preloaded_weaken_protocol_trace_from. destruct X as (T, (S, M)). apply H. + apply preloaded_weaken_valid_trace_from. destruct X as (T, (S, M)). apply H. Qed. Section VLSM_eq_properties. @@ -1873,34 +1944,34 @@ Proof. apply Hincl. Qed. -Lemma VLSM_eq_finite_protocol_trace +Lemma VLSM_eq_finite_valid_trace (s : vstate X) (tr : list (vtransition_item X)) - : finite_protocol_trace X s tr <-> finite_protocol_trace Y s tr. + : finite_valid_trace X s tr <-> finite_valid_trace Y s tr. Proof. split. - - apply (VLSM_incl_finite_protocol_trace VLSM_eq_proj1). - - apply (VLSM_incl_finite_protocol_trace VLSM_eq_proj2). + - apply (VLSM_incl_finite_valid_trace VLSM_eq_proj1). + - apply (VLSM_incl_finite_valid_trace VLSM_eq_proj2). Qed. -Lemma VLSM_eq_finite_protocol_trace_init_to +Lemma VLSM_eq_finite_valid_trace_init_to (s f : vstate X) (tr : list (vtransition_item X)) - : finite_protocol_trace_init_to X s f tr <-> - finite_protocol_trace_init_to Y s f tr. + : finite_valid_trace_init_to X s f tr <-> + finite_valid_trace_init_to Y s f tr. Proof. split. - - apply (VLSM_incl_finite_protocol_trace_init_to VLSM_eq_proj1). - - apply (VLSM_incl_finite_protocol_trace_init_to VLSM_eq_proj2). + - apply (VLSM_incl_finite_valid_trace_init_to VLSM_eq_proj1). + - apply (VLSM_incl_finite_valid_trace_init_to VLSM_eq_proj2). Qed. -Lemma VLSM_eq_protocol_state +Lemma VLSM_eq_valid_state (s : vstate X) - : protocol_state_prop X s <-> protocol_state_prop Y s. + : valid_state_prop X s <-> valid_state_prop Y s. Proof. split. - - apply (VLSM_incl_protocol_state VLSM_eq_proj1). - - apply (VLSM_incl_protocol_state VLSM_eq_proj2). + - apply (VLSM_incl_valid_state VLSM_eq_proj1). + - apply (VLSM_incl_valid_state VLSM_eq_proj2). Qed. Lemma VLSM_eq_initial_state @@ -1912,25 +1983,25 @@ Proof. - apply (VLSM_incl_initial_state VLSM_eq_proj2). Qed. -Lemma VLSM_eq_finite_protocol_trace_from +Lemma VLSM_eq_finite_valid_trace_from (s : vstate X) (tr : list (vtransition_item X)) - : finite_protocol_trace_from X s tr <-> - finite_protocol_trace_from Y s tr. + : finite_valid_trace_from X s tr <-> + finite_valid_trace_from Y s tr. Proof. split. - - apply (VLSM_incl_finite_protocol_trace_from VLSM_eq_proj1). - - apply (VLSM_incl_finite_protocol_trace_from VLSM_eq_proj2). + - apply (VLSM_incl_finite_valid_trace_from VLSM_eq_proj1). + - apply (VLSM_incl_finite_valid_trace_from VLSM_eq_proj2). Qed. -Lemma VLSM_eq_finite_protocol_trace_from_to +Lemma VLSM_eq_finite_valid_trace_from_to (s f : vstate X) (tr : list (vtransition_item X)) - : finite_protocol_trace_from_to X s f tr <-> finite_protocol_trace_from_to Y s f tr. + : finite_valid_trace_from_to X s f tr <-> finite_valid_trace_from_to Y s f tr. Proof. split. - - apply (VLSM_incl_finite_protocol_trace_from_to VLSM_eq_proj1). - - apply (VLSM_incl_finite_protocol_trace_from_to VLSM_eq_proj2). + - apply (VLSM_incl_finite_valid_trace_from_to VLSM_eq_proj1). + - apply (VLSM_incl_finite_valid_trace_from_to VLSM_eq_proj2). Qed. Lemma VLSM_eq_in_futures @@ -1942,33 +2013,33 @@ Proof. - apply (VLSM_incl_in_futures VLSM_eq_proj2). Qed. -Lemma VLSM_eq_protocol_transition +Lemma VLSM_eq_input_valid_transition : forall l s im s' om, - protocol_transition X l (s,im) (s',om) <-> - protocol_transition Y l (s,im) (s',om). + input_valid_transition X l (s,im) (s',om) <-> + input_valid_transition Y l (s,im) (s',om). Proof. split. - - apply (VLSM_incl_protocol_transition VLSM_eq_proj1). - - apply (VLSM_incl_protocol_transition VLSM_eq_proj2). + - apply (VLSM_incl_input_valid_transition VLSM_eq_proj1). + - apply (VLSM_incl_input_valid_transition VLSM_eq_proj2). Qed. -Lemma VLSM_eq_protocol_valid +Lemma VLSM_eq_input_valid : forall l s im, - protocol_valid X l (s,im) <-> protocol_valid Y l (s,im). + input_valid X l (s,im) <-> input_valid Y l (s,im). Proof. split. - - apply (VLSM_incl_protocol_valid VLSM_eq_proj1). - - apply (VLSM_incl_protocol_valid VLSM_eq_proj2). + - apply (VLSM_incl_input_valid VLSM_eq_proj1). + - apply (VLSM_incl_input_valid VLSM_eq_proj2). Qed. -Lemma VLSM_eq_protocol_generated +Lemma VLSM_eq_can_produce (s : state) (om : option message) - : option_protocol_generated_prop X s om <-> option_protocol_generated_prop Y s om. + : option_can_produce X s om <-> option_can_produce Y s om. Proof. split. - - apply (VLSM_incl_protocol_generated VLSM_eq_proj1). - - apply (VLSM_incl_protocol_generated VLSM_eq_proj2). + - apply (VLSM_incl_can_produce VLSM_eq_proj1). + - apply (VLSM_incl_can_produce VLSM_eq_proj2). Qed. Lemma VLSM_eq_can_emit @@ -1980,31 +2051,31 @@ Proof. - apply (VLSM_incl_can_emit VLSM_eq_proj2). Qed. -Lemma VLSM_eq_infinite_protocol_trace_from +Lemma VLSM_eq_infinite_valid_trace_from s ls - : infinite_protocol_trace_from X s ls <-> - infinite_protocol_trace_from Y s ls. + : infinite_valid_trace_from X s ls <-> + infinite_valid_trace_from Y s ls. Proof. split. - - apply (VLSM_incl_infinite_protocol_trace_from VLSM_eq_proj1). - - apply (VLSM_incl_infinite_protocol_trace_from VLSM_eq_proj2). + - apply (VLSM_incl_infinite_valid_trace_from VLSM_eq_proj1). + - apply (VLSM_incl_infinite_valid_trace_from VLSM_eq_proj2). Qed. -Lemma VLSM_eq_infinite_protocol_trace +Lemma VLSM_eq_infinite_valid_trace s ls - : infinite_protocol_trace X s ls <-> infinite_protocol_trace Y s ls. + : infinite_valid_trace X s ls <-> infinite_valid_trace Y s ls. Proof. split. - - apply (VLSM_incl_infinite_protocol_trace VLSM_eq_proj1). - - apply (VLSM_incl_infinite_protocol_trace VLSM_eq_proj2). + - apply (VLSM_incl_infinite_valid_trace VLSM_eq_proj1). + - apply (VLSM_incl_infinite_valid_trace VLSM_eq_proj2). Qed. -Lemma VLSM_eq_protocol_trace - : forall t, protocol_trace_prop X t <-> protocol_trace_prop Y t. +Lemma VLSM_eq_valid_trace + : forall t, valid_trace_prop X t <-> valid_trace_prop Y t. Proof. split. - - apply (VLSM_incl_protocol_trace VLSM_eq_proj1). - - apply (VLSM_incl_protocol_trace VLSM_eq_proj2). + - apply (VLSM_incl_valid_trace VLSM_eq_proj1). + - apply (VLSM_incl_valid_trace VLSM_eq_proj2). Qed. End VLSM_eq_properties. @@ -2013,10 +2084,10 @@ End VLSM_eq_properties. For VLSM <> to project to a VLSM <>, the following set of conditions is sufficient: - <>'s [initial_state]s project to <>'s [initial state]s - Every message <> (including the empty one) which can be input to a - projectable [protocol_valid] transition in <>, is a [protocol_message] + projectable [input_valid] transition in <>, is a [valid_message] in <> -- <>'s [protocol_valid] is included in <>'s [valid]. -- For all projectable [protocol_valid] inputs (in <>), <>'s [transition] +- <>'s [input_valid] is included in <>'s [valid]. +- For all projectable [input_valid] inputs (in <>), <>'s [transition] acts like <>'s [transition]. - All non-projectable transitions preserve the projected state *) @@ -2039,7 +2110,7 @@ Lemma basic_VLSM_projection_type Proof. constructor. intros is tr Htr. - induction Htr using finite_protocol_trace_from_rev_ind + induction Htr using finite_valid_trace_from_rev_ind ; [reflexivity|]. rewrite (pre_VLSM_projection_trace_project_app _ _ label_project state_project). rewrite finite_trace_last_is_last. @@ -2074,27 +2145,27 @@ Section weak_projection. Context (Hstate : weak_full_projection_initial_state_preservation X Y state_project) - (Hmessage : weak_projection_protocol_message_preservation X Y label_project state_project) + (Hmessage : weak_projection_valid_message_preservation X Y label_project state_project) . -Local Lemma basic_VLSM_projection_finite_protocol_trace_init_to +Local Lemma basic_VLSM_projection_finite_valid_trace_init_to is s tr - (Htr : finite_protocol_trace_init_to X is s tr) - : finite_protocol_trace_from_to Y (state_project is) (state_project s) (pre_VLSM_projection_trace_project _ _ label_project state_project tr). + (Htr : finite_valid_trace_init_to X is s tr) + : finite_valid_trace_from_to Y (state_project is) (state_project s) (pre_VLSM_projection_trace_project _ _ label_project state_project tr). Proof. - induction Htr using finite_protocol_trace_init_to_rev_strong_ind. + induction Htr using finite_valid_trace_init_to_rev_strong_ind. - constructor. apply Hstate. assumption. - unfold pre_VLSM_projection_trace_project. rewrite map_option_app. - apply finite_protocol_trace_from_to_app with (state_project s) + apply finite_valid_trace_from_to_app with (state_project s) ; [assumption|]. simpl. unfold pre_VLSM_projection_transition_item_project. simpl. - apply ptrace_last_pstate in IHHtr1. + apply valid_trace_last_pstate in IHHtr1. destruct (label_project l) as [lY|] eqn:Hl. - + apply finite_ptrace_from_to_singleton. - assert (Hiom : option_protocol_message_prop Y iom). - { destruct iom as [im|]; [|apply option_protocol_message_None]. + + apply finite_valid_trace_from_to_singleton. + assert (Hiom : option_valid_message_prop Y iom). + { destruct iom as [im|]; [|apply option_valid_message_None]. apply (Hmessage _ _ Hl _ _ (proj1 Ht)). assumption. } @@ -2105,26 +2176,26 @@ Proof. rewrite Ht. constructor. assumption. Qed. -Local Lemma basic_VLSM_projection_finite_protocol_trace_from +Local Lemma basic_VLSM_projection_finite_valid_trace_from (s : state) (ls : list transition_item) - (Hpxt : finite_protocol_trace_from X s ls) - : finite_protocol_trace_from Y (state_project s) (pre_VLSM_projection_trace_project _ _ label_project state_project ls). -Proof. - apply ptrace_add_default_last in Hpxt. - apply ptrace_first_pstate in Hpxt as Hs. - apply protocol_state_has_trace in Hs as [is_s [tr_s Hs]]. - specialize (finite_protocol_trace_from_to_app X _ _ _ _ _ (proj1 Hs) Hpxt) as Happ. - specialize (basic_VLSM_projection_finite_protocol_trace_init_to _ _ _ (conj Happ (proj2 Hs))) + (Hpxt : finite_valid_trace_from X s ls) + : finite_valid_trace_from Y (state_project s) (pre_VLSM_projection_trace_project _ _ label_project state_project ls). +Proof. + apply valid_trace_add_default_last in Hpxt. + apply valid_trace_first_pstate in Hpxt as Hs. + apply valid_state_has_trace in Hs as [is_s [tr_s Hs]]. + specialize (finite_valid_trace_from_to_app X _ _ _ _ _ (proj1 Hs) Hpxt) as Happ. + specialize (basic_VLSM_projection_finite_valid_trace_init_to _ _ _ (conj Happ (proj2 Hs))) as Happ_pr. rewrite (pre_VLSM_projection_trace_project_app _ _ label_project state_project) in Happ_pr. - apply finite_protocol_trace_from_to_app_split, proj2 in Happ_pr. - apply ptrace_get_last in Hs as Heqs. - apply ptrace_forget_last, proj1 in Hs. + apply finite_valid_trace_from_to_app_split, proj2 in Happ_pr. + apply valid_trace_get_last in Hs as Heqs. + apply valid_trace_forget_last, proj1 in Hs. rewrite <- (final_state_project X (type Y) label_project state_project Htype) in Happ_pr by assumption. - apply ptrace_forget_last in Happ_pr. + apply valid_trace_forget_last in Happ_pr. subst. assumption. Qed. @@ -2135,7 +2206,7 @@ Lemma basic_VLSM_weak_projection Proof. constructor. - assumption. - - apply basic_VLSM_projection_finite_protocol_trace_from. + - apply basic_VLSM_projection_finite_valid_trace_from. Qed. End weak_projection. @@ -2147,13 +2218,13 @@ Lemma basic_VLSM_weak_projection_strengthen Proof. constructor; [apply Hweak|]. intros sX trX [HtrX HsX]. split. - - revert HtrX. apply (VLSM_weak_projection_finite_protocol_trace_from Hweak). + - revert HtrX. apply (VLSM_weak_projection_finite_valid_trace_from Hweak). - revert HsX. apply Hstate. Qed. Lemma basic_VLSM_projection (Hstate : strong_full_projection_initial_state_preservation X Y state_project) - (Hmessage : weak_projection_protocol_message_preservation X Y label_project state_project) + (Hmessage : weak_projection_valid_message_preservation X Y label_project state_project) : VLSM_projection X Y label_project state_project. Proof. apply basic_VLSM_weak_projection_strengthen; [|assumption]. @@ -2173,7 +2244,7 @@ Lemma basic_VLSM_strong_projection (Htransition_Some : strong_projection_transition_preservation_Some X Y label_project state_project) (Htransition_None : strong_projection_transition_consistency_None _ _ label_project state_project) (Hstate : strong_full_projection_initial_state_preservation X Y state_project) - (Hmessage : strong_projection_protocol_message_preservation X Y) + (Hmessage : strong_projection_valid_message_preservation X Y) : VLSM_projection X Y label_project state_project. Proof. apply basic_VLSM_projection. @@ -2181,7 +2252,7 @@ Proof. - apply strong_projection_transition_preservation_Some_weaken. assumption. - apply strong_projection_transition_consistency_None_weaken. assumption. - assumption. - - apply strong_projection_protocol_message_preservation_weaken. assumption. + - apply strong_projection_valid_message_preservation_weaken. assumption. Qed. Lemma basic_VLSM_projection_type_preloaded @@ -2194,7 +2265,7 @@ Lemma basic_VLSM_projection_type_preloaded Proof. constructor. intros is tr Htr. - induction Htr using finite_protocol_trace_from_rev_ind + induction Htr using finite_valid_trace_from_rev_ind ; [reflexivity|]. rewrite (@pre_VLSM_projection_trace_project_app _ (type (pre_loaded_with_all_messages_vlsm X)) (type Y) label_project state_project). rewrite finite_trace_last_is_last. @@ -2223,27 +2294,27 @@ Proof. constructor; [assumption|]. intros sX trX HtrX. split; [|apply Hstate; apply HtrX]. - induction HtrX using finite_protocol_trace_rev_ind. - - constructor. apply initial_is_protocol. + induction HtrX using finite_valid_trace_rev_ind. + - constructor. apply initial_state_is_valid. apply Hstate; assumption. - rewrite (@pre_VLSM_projection_trace_project_app _ (type (pre_loaded_with_all_messages_vlsm X)) (type Y) label_project state_project). - apply (finite_protocol_trace_from_app_iff (pre_loaded_with_all_messages_vlsm Y)). + apply (finite_valid_trace_from_app_iff (pre_loaded_with_all_messages_vlsm Y)). split; [assumption|]. simpl. unfold pre_VLSM_projection_transition_item_project. simpl. - apply finite_ptrace_last_pstate in IHHtrX. + apply finite_valid_trace_last_pstate in IHHtrX. destruct Hx as [[_ [_ Hv]] Ht]. rewrite <- (final_state_project _ _ _ _ Htype) in IHHtrX |- * by apply HtrX. destruct (label_project l) as [lY|] eqn:Hl. - + apply (finite_ptrace_singleton (pre_loaded_with_all_messages_vlsm Y)). - assert (Hiom : option_protocol_message_prop (pre_loaded_with_all_messages_vlsm Y) iom). - { destruct iom as [im|]; [|apply option_protocol_message_None]. - apply (any_message_is_protocol_in_preloaded Y). + + apply (finite_valid_trace_singleton (pre_loaded_with_all_messages_vlsm Y)). + assert (Hiom : option_valid_message_prop (pre_loaded_with_all_messages_vlsm Y) iom). + { destruct iom as [im|]; [|apply option_valid_message_None]. + apply (any_message_is_valid_in_preloaded Y). } apply (Hvalid _ _ Hl) in Hv. apply (Htransition_Some _ _ Hl) in Ht. repeat split; assumption. - + apply (finite_ptrace_empty (pre_loaded_with_all_messages_vlsm Y)). assumption. + + apply (finite_valid_trace_from_empty (pre_loaded_with_all_messages_vlsm Y)). assumption. Qed. Lemma basic_VLSM_projection_type_preloaded_with @@ -2257,7 +2328,7 @@ Lemma basic_VLSM_projection_type_preloaded_with Proof. constructor. intros is tr Htr. - induction Htr using finite_protocol_trace_from_rev_ind + induction Htr using finite_valid_trace_from_rev_ind ; [reflexivity|]. rewrite (@pre_VLSM_projection_trace_project_app _ (type (pre_loaded_vlsm X P)) (type Y) label_project state_project). rewrite finite_trace_last_is_last. @@ -2281,35 +2352,35 @@ Lemma basic_VLSM_projection_preloaded_with (Htransition_Some : strong_projection_transition_preservation_Some X Y label_project state_project) (Htransition_None : strong_projection_transition_consistency_None _ _ label_project state_project) (Hstate : strong_full_projection_initial_state_preservation X Y state_project) - (Hmessage : weak_projection_protocol_message_preservation (pre_loaded_vlsm X P) (pre_loaded_vlsm Y Q) label_project state_project) + (Hmessage : weak_projection_valid_message_preservation (pre_loaded_vlsm X P) (pre_loaded_vlsm Y Q) label_project state_project) : VLSM_projection (pre_loaded_vlsm X P) (pre_loaded_vlsm Y Q) label_project state_project. Proof. specialize (basic_VLSM_projection_type_preloaded_with X Y P Q label_project state_project Htransition_None) as Htype. constructor; [assumption|]. intros sX trX HtrX. split; [|apply Hstate; apply HtrX]. - induction HtrX using finite_protocol_trace_rev_ind. - - constructor. apply initial_is_protocol. + induction HtrX using finite_valid_trace_rev_ind. + - constructor. apply initial_state_is_valid. apply Hstate; assumption. - rewrite (@pre_VLSM_projection_trace_project_app _ (type (pre_loaded_vlsm X P)) (type Y) label_project state_project). - apply (finite_protocol_trace_from_app_iff (pre_loaded_vlsm Y Q)). + apply (finite_valid_trace_from_app_iff (pre_loaded_vlsm Y Q)). split; [assumption|]. simpl. unfold pre_VLSM_projection_transition_item_project. simpl. - apply finite_ptrace_last_pstate in IHHtrX. + apply finite_valid_trace_last_pstate in IHHtrX. apply proj1 in Hx as Hpv. destruct Hx as [[_ [_ Hv]] Ht]. rewrite <- (final_state_project _ _ _ _ Htype) in IHHtrX |- * by apply HtrX. destruct (label_project l) as [lY|] eqn:Hl. - + apply (finite_ptrace_singleton (pre_loaded_vlsm Y Q)). - assert (Hiom : option_protocol_message_prop (pre_loaded_vlsm Y Q) iom). - { destruct iom as [im|]; [|apply option_protocol_message_None]. + + apply (finite_valid_trace_singleton (pre_loaded_vlsm Y Q)). + assert (Hiom : option_valid_message_prop (pre_loaded_vlsm Y Q) iom). + { destruct iom as [im|]; [|apply option_valid_message_None]. apply (Hmessage _ _ Hl) in Hpv; assumption. } apply (Hvalid _ _ Hl) in Hv. apply (Htransition_Some _ _ Hl) in Ht. repeat split; assumption. - + apply (finite_ptrace_empty (pre_loaded_vlsm Y Q)). assumption. + + apply (finite_valid_trace_from_empty (pre_loaded_vlsm Y Q)). assumption. Qed. @@ -2318,10 +2389,10 @@ which are easy to verify in a practical setting. One such result is the followin For VLSM <> to fully-project to a VLSM <>, the following set of conditions is sufficient: - <>'s [initial_state]s project to <>'s [initial state]s -- Every message <> (including the empty one) which can be input to a -[protocol_valid] transition in <>, is a [protocol_message] in <> -- <>'s [protocol_valid] is included in <>'s [valid]. -- For all [protocol_valid] inputs (in <>), <>'s [transition] acts +- Every message <> (including the empty one) which can be input to + an [input_valid] transition in <>, is a [valid_message] in <> +- <>'s [input_valid] is included in <>'s [valid]. +- For all [input_valid] inputs (in <>), <>'s [transition] acts like <>'s [transition]. *) @@ -2346,51 +2417,51 @@ Context (Hmessage : weak_full_projection_initial_message_preservation X Y state_project) . -Lemma weak_projection_protocol_message_preservation_from_full - : weak_projection_protocol_message_preservation X Y (Some ∘ label_project) state_project. +Lemma weak_projection_valid_message_preservation_from_full + : weak_projection_valid_message_preservation X Y (Some ∘ label_project) state_project. Proof. intros lX lY Hl s m Hv HsY. apply proj2 in Hv as Hom. apply proj1 in Hom. - apply can_emit_protocol_iff in Hom. + apply emitted_messages_are_valid_iff in Hom. destruct Hom as [Him | Hemit]. - apply (Hmessage _ _ _ Hv); assumption. - apply can_emit_has_trace in Hemit as [is [tr [item [Htr Hm]]]]. destruct item. simpl in *. subst. - apply ptrace_add_default_last in Htr. + apply valid_trace_add_default_last in Htr. rewrite finite_trace_last_is_last in Htr. simpl in Htr. remember (tr ++ _) as tr'. - cut (option_protocol_message_prop Y (Some m)); [exact id|]. + cut (option_valid_message_prop Y (Some m)); [exact id|]. exists (state_project destination). clear Hv Hl lX lY. revert tr l input Heqtr'. generalize (Some m) as om. - induction Htr using finite_protocol_trace_init_to_rev_strong_ind + induction Htr using finite_valid_trace_init_to_rev_strong_ind ; intros; [destruct tr; simpl in *; congruence|]. apply app_inj_tail in Heqtr' as [Heqtr Heqitem]. subst tr0. inversion Heqitem. subst l0 input oom. clear Heqitem. - assert (Hs : protocol_state_prop Y (state_project s0)). + assert (Hs : valid_state_prop Y (state_project s0)). { destruct_list_last tr s_tr' s_item Heqtr. - subst tr. destruct Htr1 as [Hs His]. inversion Hs. subst. apply Hstate. assumption. - subst. - apply ptrace_get_last in Htr1 as Hs0. + apply valid_trace_get_last in Htr1 as Hs0. rewrite finite_trace_last_is_last in Hs0. destruct s_item. simpl in Hs0. subst destination. specialize (IHHtr1 _ _ _ _ eq_refl). eexists; exact IHHtr1. } destruct Hs as [_om Hs]. - assert (Hom : option_protocol_message_prop Y iom). - { destruct iom as [im|]; [|apply option_protocol_message_None]. + assert (Hom : option_valid_message_prop Y iom). + { destruct iom as [im|]; [|apply option_valid_message_None]. unfold empty_initial_message_or_final_output in Heqiom. destruct_list_last iom_tr iom_tr' iom_item Heqiom_tr. - apply (Hmessage _ _ _ (proj1 Ht)); [|assumption]. eexists; exact Hs. - subst. - apply ptrace_get_last in Htr2 as Hs0. + apply valid_trace_get_last in Htr2 as Hs0. rewrite finite_trace_last_is_last in Hs0. destruct iom_item. simpl in *. subst. specialize (IHHtr2 _ _ _ _ eq_refl). @@ -2398,7 +2469,7 @@ Proof. } destruct Hom as [_s Hom]. apply - (protocol_generated Y _ _ Hs _ _ Hom (label_project l)). + (valid_generated_state_message Y _ _ Hs _ _ Hom (label_project l)). + apply Hvalid; [apply Ht|exists _om|exists _s]; assumption. + apply Htransition. assumption. Qed. @@ -2410,9 +2481,9 @@ Proof. spec Hproj; [apply weak_projection_transition_preservation_Some_from_full; assumption|]. spec Hproj; [apply weak_projection_transition_consistency_None_from_full|]. spec Hproj; [assumption|]. - spec Hproj; [apply weak_projection_protocol_message_preservation_from_full|]. + spec Hproj; [apply weak_projection_valid_message_preservation_from_full|]. constructor. intro; intros. - apply (VLSM_weak_projection_finite_protocol_trace_from Hproj) in H. + apply (VLSM_weak_projection_finite_valid_trace_from Hproj) in H. assumption. Qed. @@ -2425,7 +2496,7 @@ Lemma basic_VLSM_weak_full_projection_strengthen Proof. constructor. intros sX trX [HtrX HsX]. split. - - revert HtrX. apply (VLSM_weak_full_projection_finite_protocol_trace_from Hweak). + - revert HtrX. apply (VLSM_weak_full_projection_finite_valid_trace_from Hweak). - revert HsX. apply Hstate. Qed. @@ -2473,19 +2544,19 @@ Proof. constructor. intros sX trX HtrX. split; [|apply Hstate; apply HtrX]. - induction HtrX using finite_protocol_trace_rev_ind. - - constructor. apply initial_is_protocol. + induction HtrX using finite_valid_trace_rev_ind. + - constructor. apply initial_state_is_valid. apply Hstate; assumption. - - setoid_rewrite map_app. apply finite_protocol_trace_from_app_iff. + - setoid_rewrite map_app. apply finite_valid_trace_from_app_iff. split; [assumption|]. - simpl. apply (finite_ptrace_singleton (pre_loaded_with_all_messages_vlsm Y)). + simpl. apply (finite_valid_trace_singleton (pre_loaded_with_all_messages_vlsm Y)). destruct Hx as [[_ [_ Hv]] Ht]. apply Hvalid in Hv. apply Htransition in Ht. rewrite (pre_VLSM_full_projection_finite_trace_last _ _ label_project state_project) in Hv, Ht. repeat split; [..|assumption|assumption]. - + apply finite_ptrace_last_pstate in IHHtrX. assumption. - + apply any_message_is_protocol_in_preloaded. + + apply finite_valid_trace_last_pstate in IHHtrX. assumption. + + apply any_message_is_valid_in_preloaded. Qed. Lemma basic_VLSM_full_projection_preloaded_with @@ -2503,30 +2574,30 @@ Lemma basic_VLSM_full_projection_preloaded_with Proof. constructor. intros sX trX HtrX. - apply ptrace_add_default_last in HtrX. + apply valid_trace_add_default_last in HtrX. split; [|apply Hstate; apply HtrX]. - induction HtrX using finite_protocol_trace_init_to_rev_strong_ind. - - constructor. apply initial_is_protocol. + induction HtrX using finite_valid_trace_init_to_rev_strong_ind. + - constructor. apply initial_state_is_valid. apply Hstate; assumption. - - setoid_rewrite map_app. apply finite_protocol_trace_from_app_iff. + - setoid_rewrite map_app. apply finite_valid_trace_from_app_iff. split; [assumption|]. - simpl. apply (finite_ptrace_singleton (pre_loaded_vlsm Y Q)). + simpl. apply (finite_valid_trace_singleton (pre_loaded_vlsm Y Q)). destruct Ht as [[_ [_ Hv]] Ht]. apply Hvalid in Hv. apply Htransition in Ht. - apply ptrace_get_last in HtrX1. subst s. + apply valid_trace_get_last in HtrX1. subst s. rewrite (pre_VLSM_full_projection_finite_trace_last _ _ label_project state_project) in Hv, Ht. simpl. repeat split; [..|assumption|assumption]. - + apply finite_ptrace_last_pstate in IHHtrX1. assumption. - + destruct iom as [m|]; [|apply option_protocol_message_None]. + + apply finite_valid_trace_last_pstate in IHHtrX1. assumption. + + destruct iom as [m|]; [|apply option_valid_message_None]. unfold empty_initial_message_or_final_output in Heqiom. destruct_list_last iom_tr iom_tr' iom_lst Heqiom_tr - ; [apply option_initial_message_is_protocol; destruct Heqiom as [Him | Hp]|]. + ; [apply option_initial_message_is_valid; destruct Heqiom as [Him | Hp]|]. * left. revert Him. apply Hmessage. * right. apply PimpliesQ. assumption. * apply - (protocol_trace_output_is_protocol (pre_loaded_vlsm Y Q) _ _ IHHtrX2 m). + (valid_trace_output_is_valid (pre_loaded_vlsm Y Q) _ _ IHHtrX2 m). setoid_rewrite map_app. apply Exists_app. right. left. assumption. Qed. @@ -2548,7 +2619,7 @@ Context Lemma basic_VLSM_incl (Hinitial_state : strong_incl_initial_state_preservation MX MY) - (Hinitial_protocol_message : weak_incl_initial_message_preservation MX MY) + (Hinitial_valid_message : weak_incl_initial_message_preservation MX MY) (Hvalid : weak_incl_valid_preservation MX MY) (Htransition : weak_incl_transition_preservation MX MY) : VLSM_incl X Y. @@ -2559,7 +2630,7 @@ Qed. Lemma basic_VLSM_strong_incl (Hinitial_state : strong_incl_initial_state_preservation MX MY) - (Hinitial_protocol_message : strong_incl_initial_message_preservation MX MY) + (Hinitial_valid_message : strong_incl_initial_message_preservation MX MY) (Hvalid : strong_incl_valid_preservation MX MY) (Htransition : strong_incl_transition_preservation MX MY) : VLSM_incl X Y. @@ -2701,13 +2772,13 @@ Proof. - apply pre_loaded_with_all_messages_vlsm_idem_r. Qed. -Lemma vlsm_is_pre_loaded_with_False_protocol_prop s om - : protocol_prop X s om <-> protocol_prop (pre_loaded_vlsm X (fun m => False)) s om. +Lemma vlsm_is_pre_loaded_with_False_valid_state_message s om + : valid_state_message_prop X s om <-> valid_state_message_prop (pre_loaded_vlsm X (fun m => False)) s om. Proof. pose proof vlsm_is_pre_loaded_with_False as Heq. apply VLSM_eq_incl_iff in Heq. destruct X as (T, (S, M)); simpl in *. - split; (apply VLSM_incl_protocol_prop; [|cbv;tauto]); apply Heq. + split; (apply VLSM_incl_valid_state_message; [|cbv;tauto]); apply Heq. Qed. Lemma pre_loaded_with_all_messages_can_emit @@ -2721,23 +2792,23 @@ Qed. End VLSM_incl_preloaded_properties. -Lemma preloaded_weaken_finite_protocol_trace_from +Lemma preloaded_weaken_finite_valid_trace_from {message : Type} (X : VLSM message) (from : state) (tr : list transition_item) : - finite_protocol_trace_from X from tr -> - finite_protocol_trace_from (pre_loaded_with_all_messages_vlsm X) from tr. + finite_valid_trace_from X from tr -> + finite_valid_trace_from (pre_loaded_with_all_messages_vlsm X) from tr. Proof. destruct X as (T, (S, M)). - apply VLSM_incl_finite_protocol_trace_from. + apply VLSM_incl_finite_valid_trace_from. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. Qed. -Lemma preloaded_weaken_finite_protocol_trace_from_to +Lemma preloaded_weaken_finite_valid_trace_from_to {message : Type} (X : VLSM message) (from to : state) (tr : list transition_item) : - finite_protocol_trace_from_to X from to tr -> - finite_protocol_trace_from_to (pre_loaded_with_all_messages_vlsm X) from to tr. + finite_valid_trace_from_to X from to tr -> + finite_valid_trace_from_to (pre_loaded_with_all_messages_vlsm X) from to tr. Proof. destruct X as (T, (S, M)). - apply VLSM_incl_finite_protocol_trace_from_to. + apply VLSM_incl_finite_valid_trace_from_to. apply vlsm_incl_pre_loaded_with_all_messages_vlsm. Qed. @@ -2778,7 +2849,7 @@ Next Obligation. Qed. Definition projection_induced_initial_message_prop : message -> Prop := - protocol_message_prop X. + valid_message_prop X. Definition projection_induced_vlsm_sig : VLSMSign TY := {| initial_message_prop := projection_induced_initial_message_prop @@ -2799,7 +2870,7 @@ Definition projection_induced_valid : Prop := let (sY, om) := somY in exists lX sX, label_project lX = Some lY /\ state_project sX = sY /\ - protocol_valid X lX (sX, om). + input_valid X lX (sX, om). Definition projection_induced_vlsm_class : VLSMClass projection_induced_vlsm_sig := {| transition := projection_induced_transition @@ -2809,33 +2880,86 @@ Definition projection_induced_vlsm_class : VLSMClass projection_induced_vlsm_sig Definition projection_induced_vlsm : VLSM message := mk_vlsm projection_induced_vlsm_class. -(** The results of a transition and the results of the transition -performed on the lifting of the projection of its inputs should have equal +(** <> is a left-inverse of <> *) +Definition induced_projection_label_lift_prop : Prop := + forall lY, label_project (label_lift lY) = Some lY. + +(** <> is a left-inverse of <> *) +Definition induced_projection_state_lift_prop : Prop := + forall sY, state_project (state_lift sY) = sY. + +(** Transitions through states and labels with the same projections using the +same message should lead to the same output message and states with the same projections. *) +Definition induced_projection_transition_consistency_Some : Prop := + forall lX1 lX2 lY, label_project lX1 = Some lY -> label_project lX2 = Some lY -> + forall sX1 sX2, state_project sX1 = state_project sX2 -> + forall iom sX1' oom1, vtransition X lX1 (sX1, iom) = (sX1', oom1) -> + forall sX2' oom2, vtransition X lX2 (sX2, iom) = (sX2', oom2) -> + state_project sX1' = state_project sX2' /\ oom1 = oom2. + +(** A weaker version of [induced_projection_transition_consistency_Some] *) Definition weak_projection_transition_consistency_Some : Prop := forall lX lY, label_project lX = Some lY -> - forall s1 om s1' om1', protocol_transition X lX (s1, om) (s1', om1') -> + forall s1 om s1' om1', input_valid_transition X lX (s1, om) (s1', om1') -> forall s2' om2', vtransition X (label_lift lY) (state_lift (state_project s1), om) = (s2', om2') -> state_project s1' = state_project s2' /\ om1' = om2'. -(* TODO(traiansf): the definition above follows very specifically conditions -that we encounter in practice. A more principled approach would be to split -this property into -(1) vtransition X is compatible with the kernels of label_project - and state_project -(2) label_project ∘ label_lift ∘ label_project = id and -(3) state_project ∘ state_lift ∘ state_project = id +(* TODO(traiansf): remove the definition above and assume the properties +deriving it in the next lemma instead. *) + +Lemma basic_weak_projection_transition_consistency_Some + : induced_projection_label_lift_prop -> induced_projection_state_lift_prop -> + induced_projection_transition_consistency_Some -> + weak_projection_transition_consistency_Some. +Proof. + intros Hlabel Hstate Htransition lX lY HlX_pr sX1 iom sX1' oom1 Ht1 + sX2' oom2 Ht2. + apply proj2 in Ht1. + eapply Htransition; [eassumption| apply Hlabel| |exact Ht1|exact Ht2]. + symmetry. + apply Hstate. +Qed. + +(** Under transition-consistency assumptions, valid messages of the +[projection_induced_vlsm] coincide with those of the source [VLSM]. *) +Lemma projection_induced_valid_message_char + (Htransition_Some : weak_projection_transition_consistency_Some) + : forall om, option_valid_message_prop projection_induced_vlsm om <-> + option_valid_message_prop X om. +Proof. + split; cycle 1. + - intros Hm. + destruct om as [m |]. + + apply initial_message_is_valid. assumption. + + apply option_valid_message_None. + - intros [s Hsom]. + induction Hsom. + + destruct om as [m |]. + * assumption. + * apply option_valid_message_None. + + destruct Hv as [lX [sX [HlX_pr [HsX_pr [HsX [HomX Hv]]]]]]. + cbn in Ht. + destruct (vtransition _ _ _) as (_s'X, __om') eqn: H_tX. + inversion Ht; subst; clear Ht. + destruct (vtransition X lX (sX, om)) as (s'X, _om') eqn: HtX. + assert (HivtX : input_valid_transition X lX (sX, om) (s'X, _om')) + by (repeat split; assumption). + replace om' with _om'. + * eapply input_valid_transition_out. eassumption. + * eapply Htransition_Some; eassumption. +Qed. Context (Htransition_None : weak_projection_transition_consistency_None _ _ label_project state_project) - (Htransition_Some : weak_projection_transition_consistency_Some) (Htype : VLSM_projection_type X TY label_project state_project := basic_VLSM_projection_type X TY label_project state_project Htransition_None) . Lemma projection_induced_vlsm_is_projection + (Htransition_Some : weak_projection_transition_consistency_Some) : VLSM_projection X projection_induced_vlsm label_project state_project. Proof. apply basic_VLSM_projection; intro; intros. @@ -2850,12 +2974,167 @@ Proof. - apply (Htransition_None _ H _ _ _ _ H0). - exists s. split; [reflexivity|assumption]. - destruct Hv as [_ [Hm _]]. - apply initial_message_is_protocol. + apply initial_message_is_valid. assumption. Qed. +Section projection_induced_friendliness. + +Context + (Hlabel_lift : induced_projection_label_lift_prop) + (Hstate_lift : induced_projection_state_lift_prop) + (Htransition_consistency : induced_projection_transition_consistency_Some) + (Htransition_Some : weak_projection_transition_consistency_Some + := basic_weak_projection_transition_consistency_Some Hlabel_lift Hstate_lift Htransition_consistency) + (Hproj := projection_induced_vlsm_is_projection Htransition_Some) + . + +Lemma induced_projection_transition_item_lift + (item : @transition_item _ TY) + : @pre_VLSM_projection_transition_item_project _ (type X) _ + label_project state_project + (pre_VLSM_full_projection_trace_item_project _ _ label_lift state_lift item) + = Some item. +Proof. + destruct item. + unfold pre_VLSM_full_projection_trace_item_project, + pre_VLSM_projection_transition_item_project. + simpl. + rewrite Hlabel_lift. + f_equal. f_equal. + apply Hstate_lift. +Qed. + +Lemma induced_projection_trace_lift + (tr : list (@transition_item _ TY)) + : @pre_VLSM_projection_trace_project _ (type X) _ + label_project state_project + (pre_VLSM_full_projection_finite_trace_project _ _ label_lift state_lift tr) + = tr. +Proof. + induction tr; [reflexivity|]. + simpl. + rewrite induced_projection_transition_item_lift. + f_equal. + assumption. +Qed. + +(** If there is a way to "lift" valid traces of the [projection_induced_vlsm] +to the original [VLSM], then the induced [VLSM_projection] is friendly. +*) +Lemma basic_projection_induces_friendliness + : VLSM_full_projection projection_induced_vlsm X label_lift state_lift -> + projection_friendly_prop Hproj. +Proof. + intros Hfull_proj isY trY HtrY. + exists (state_lift isY). + exists (VLSM_full_projection_finite_trace_project Hfull_proj trY). + intuition. + + apply (VLSM_full_projection_finite_valid_trace Hfull_proj). + assumption. + + apply induced_projection_trace_lift. +Qed. + +End projection_induced_friendliness. + End projection_induced_vlsm. +(** Under [weak_projection_transition_consistency_Some] assumptions, +[VLSM_incl]usion between source [VLSM]s implies [VLSM_incl]usion between +their projections induced by the same maps. +*) +Lemma projection_induced_vlsm_incl + {message : Type} + {TX : VLSMType message} + {SigX1 SigX2: VLSMSign TX} + (MX1 : VLSMClass SigX1) (MX2 : VLSMClass SigX2) + (X1 := mk_vlsm MX1) (X2 := mk_vlsm MX2) + (TY : VLSMType message) + (label_project : @label _ TX -> option (@label _ TY)) + (state_project : @state _ TX -> @state _ TY) + (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (label_lift : @label _ TY -> @label _ TX) + (state_lift : @state _ TY -> @state _ TX) + (XY1 : VLSM message := projection_induced_vlsm X1 TY label_project state_project label_lift state_lift) + (XY2 : VLSM message := projection_induced_vlsm X2 TY label_project state_project label_lift state_lift) + (Htransition_Some1 : weak_projection_transition_consistency_Some X1 TY label_project state_project label_lift state_lift) + (Htransition_Some2 : weak_projection_transition_consistency_Some X2 TY label_project state_project label_lift state_lift) + : VLSM_incl X1 X2 -> VLSM_incl XY1 XY2. +Proof. + intros Hincl. + apply VLSM_incl_finite_traces_characterization. + assert (His : forall s, vinitial_state_prop XY1 s -> vinitial_state_prop XY2 s). + { + intros is [s1 [Hs1_pr Hs1]]. + exists s1. + split; [assumption|]. + apply VLSM_incl_initial_state; assumption. + } + intros is tr Htr. + split; [|apply His, Htr]. + induction Htr using finite_valid_trace_rev_ind + ; [apply (finite_valid_trace_from_empty XY2), initial_state_is_valid, His; assumption|]. + apply (finite_valid_trace_from_app_iff XY2). + split; [apply IHHtr|]. + apply (finite_valid_trace_singleton XY2). + destruct Hx as [[_ [_ [lX [sX [HlX_pr [HsX_pr HpvX1]]]]]] Ht]. + cbn in Ht. + destruct (vtransition _ _ _) as (_s'X, _oom) eqn:H_tX1. + inversion Ht. subst. clear Ht. + destruct (vtransition X1 lX (sX, iom)) as (s'X, _oom) eqn:HtX1. + assert (HivtX1 : input_valid_transition X1 lX (sX, iom) (s'X, _oom)) + by (split; assumption). + simpl in HsX_pr, H_tX1. rewrite <- HsX_pr in H_tX1. + apply (Htransition_Some1 _ _ HlX_pr _ _ _ _ HivtX1) in H_tX1 + as [Heq_s'X_pr Heq_oom]. + subst. + apply (VLSM_incl_input_valid_transition Hincl) in HivtX1. + repeat split. + - eapply finite_valid_trace_last_pstate; eassumption. + - apply projection_induced_valid_message_char; [assumption| apply HivtX1]. + - exists lX, sX. intuition. apply HivtX1. + - cbn in *. rewrite <- HsX_pr. + destruct (vtransition X2 _ _) as (_s'X2, _oom) eqn:H_tX2. + apply (Htransition_Some2 _ _ HlX_pr _ _ _ _ HivtX1) in H_tX2 + as [Heq_s'X_pr' Heq_oom]. + congruence. +Qed. + +(** Under [weak_projection_transition_consistency_Some] assumptions, +[VLSM_eq]uality between source [VLSM]s implies [VLSM_eq]uality between +their projections induced by the same maps. +*) +Lemma projection_induced_vlsm_eq + {message : Type} + {TX : VLSMType message} + {SigX1 SigX2: VLSMSign TX} + (MX1 : VLSMClass SigX1) (MX2 : VLSMClass SigX2) + (X1 := mk_vlsm MX1) (X2 := mk_vlsm MX2) + (TY : VLSMType message) + (label_project : @label _ TX -> option (@label _ TY)) + (state_project : @state _ TX -> @state _ TY) + (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (label_lift : @label _ TY -> @label _ TX) + (state_lift : @state _ TY -> @state _ TX) + (XY1 : VLSM message := projection_induced_vlsm X1 TY label_project state_project label_lift state_lift) + (XY2 : VLSM message := projection_induced_vlsm X2 TY label_project state_project label_lift state_lift) + (Htransition_Some1 : weak_projection_transition_consistency_Some X1 TY label_project state_project label_lift state_lift) + (Htransition_Some2 : weak_projection_transition_consistency_Some X2 TY label_project state_project label_lift state_lift) + : VLSM_eq X1 X2 -> VLSM_eq XY1 XY2. +Proof. + intro Heq. + apply VLSM_eq_incl_iff. + split. + - apply + (projection_induced_vlsm_incl MX1 MX2 TY + label_project state_project label_lift state_lift + Htransition_Some1 Htransition_Some2 (VLSM_eq_proj1 Heq)). + - apply + (projection_induced_vlsm_incl MX2 MX1 TY + label_project state_project label_lift state_lift + Htransition_Some2 Htransition_Some1 (VLSM_eq_proj2 Heq)). +Qed. + Section same_VLSM_full_projection. Context diff --git a/theories/VLSM/Core/Validating.v b/theories/VLSM/Core/Validating.v deleted file mode 100644 index d0d781ff..00000000 --- a/theories/VLSM/Core/Validating.v +++ /dev/null @@ -1,314 +0,0 @@ -From stdpp Require Import prelude. -From Coq Require Import FinFun. -From VLSM Require Import Lib.Preamble Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces. - -(** * VLSM Validating Projections - -In the sequel we fix a composite VLSM <> obtained from an indexed family -of VLSMs <> and a <>, and an index <>, corresponding to -component <>. -*) - -Section validating_projection. - -Context - {message : Type} - {index : Type} - {IndEqDec : EqDecision index} - (IM : index -> VLSM message) - (constraint : composite_label IM -> composite_state IM * option message -> Prop) - (X := composite_vlsm IM constraint) - (i : index) - (Xi := composite_vlsm_constrained_projection IM constraint i) - . - -(** -We say that the component <> of X is validating received messages if -non-[projection_valid]itiy implied non-component-[valid]ity (or non-reachability). -*) - -Definition validating_projection_received_messages_prop - := - forall (li : vlabel (IM i)) (si : vstate (IM i)) (mi : message), - ~ vvalid Xi li (si, Some mi) - -> ~ protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (si, Some mi). - -(** -We can slightly generalize the definition above to also include empty messages -and state it in a positive manner as the [validating_projection_prop]erty below, -requiring that [valid]ity in the component (for reachable states) implies -[projection_valid]ity. -*) - -Definition validating_projection_prop := - forall (li : vlabel (IM i)) (siomi : vstate (IM i) * option message), - protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) li siomi -> - vvalid Xi li siomi. - -(** An alternative formulation of the above property with a seemingly -stronger hypoyesis, states that component (IM i) is validating for constraint -if for any @si@ protocol state in the projection @Xi@, @li valid (IM i) (si, om)@ -implies that there exists a state @s@ whose ith component is @si@, -and @s@ and @om@ are protocol in @X@, and @(i,li) valid (s,om)@ in @X@, that is, -we have that @li valid (si, om)@ in the projection @Xi@. -*) -Definition validating_projection_prop_alt := - forall (li : vlabel (IM i)) (siom : vstate (IM i) * option message), - let (si, om) := siom in - vvalid (IM i) li siom -> - protocol_state_prop Xi si -> - vvalid Xi li siom. - -(** Under validating assumptions, all reachable states for component @IM i@ are -protocol states in the projection @Xi@. -*) -Lemma validating_alt_free_states_are_projection_states - : validating_projection_prop_alt -> - forall s, - protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s -> - protocol_state_prop Xi s. -Proof. - intros Hvalidating s Hs. - induction Hs using protocol_state_prop_ind; - [apply initial_is_protocol;assumption|]. - change s' with (fst (s',om')). - destruct Ht as [[_ [_ Hvalid]] Htrans]. - apply (projection_valid_implies_destination_projection_protocol_state IM constraint i) with l s om om' - ; [|assumption]. - apply (Hvalidating l (s,om));assumption. -Qed. - -(** Below we show that the two definitions above are actually equivalent. -*) -Lemma validating_projection_prop_alt_iff - : validating_projection_prop_alt <-> validating_projection_prop. -Proof. - split. - - intros Hvalidating l (si, om) Hvalid. - destruct Hvalid as [Hsi [_ Hvalid]]. - apply validating_alt_free_states_are_projection_states in Hsi - ; [|assumption]. - exact (Hvalidating l (si,om) Hvalid Hsi). - - intros Hvalidating l (si, om) Hvalid HXisi. - specialize (Hvalidating l (si, om)). - apply Hvalidating. - repeat split; [| apply any_message_is_protocol_in_preloaded | assumption]. - revert HXisi. - apply VLSM_incl_protocol_state. - apply proj_pre_loaded_with_all_messages_incl. -Qed. - -Lemma validating_free_states_are_projection_states - : validating_projection_prop -> - forall s, - protocol_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s -> - protocol_state_prop Xi s. -Proof. - rewrite <- validating_projection_prop_alt_iff. - apply validating_alt_free_states_are_projection_states. -Qed. - -(** -It is easy to see that the [validating_projection_prop]erty includes the -[validating_projection_received_messages_prop]erty. -*) -Lemma validating_projection_messages_received - : validating_projection_prop -> validating_projection_received_messages_prop. -Proof. - unfold validating_projection_prop. unfold validating_projection_received_messages_prop. intros. - intro Hvalid. elim H0. clear H0. - specialize (H li (si, Some mi) Hvalid). assumption. -Qed. - -(** -We say that component <> is [validating_transitions] if any [valid] -transition (from a reachable state) in component <> can be "lifted" to -a [protocol_transition] in <>. - -*) - -Definition validating_transitions := - forall - (si : vstate (IM i)) - (omi : option message) - (li : vlabel (IM i)) - , - protocol_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (si, omi) - -> - (exists - (s s' : vstate X) - (om' : option message), - si = s i - /\ - protocol_transition X (existT i li) (s, omi) (s', om') - ). - -(** -Next two results show that the (simpler) [validating_projection_prop]erty -is equivalent with the [validating_transitions] property. -*) - -Lemma validating_projection_messages_transitions - : validating_projection_prop -> validating_transitions. -Proof. - intros Hvalidating si omi li Hpvi. - specialize (Hvalidating li (si, omi) Hpvi). clear Hpvi. - destruct Hvalidating as [s [Hsi [Hps [Hopm [Hvalid Hctr]]]]]. - destruct (vtransition X (existT i li) (s, omi)) as (s', om') eqn:Heqt. - exists s. exists s'. exists om'. - subst si. - repeat split; assumption. -Qed. - -Lemma validating_transitions_messages - : validating_transitions -> validating_projection_prop. -Proof. - intros Hvalidating li (si,omi) Hpvi. - specialize (Hvalidating si omi li Hpvi). clear Hpvi. - destruct Hvalidating as [s [s' [om' [Hsi [Hvalid Htransition]]]]]. - symmetry in Hsi. - exists s. split; assumption. -Qed. - -(** ** Validating projections and Byzantine behavior - -In the sequel we assume that <> has the [validating_projection_prop]erty for -component <>. Let <> be the projection of <> to component <> -and <> be the [pre_loaded_with_all_messages_vlsm] associated to component <>. -*) - -Section pre_loaded_with_all_messages_validating_proj. - Context - (Hvalidating : validating_projection_prop) - (PreLoaded := pre_loaded_with_all_messages_vlsm (IM i)) - . - -(** -We can show that <> is included in <> by applying the meta-lemma -[VLSM_incl_finite_traces_characterization], and by induction on the length -of a trace. The [validating_projection_prop]erty is used to translate -[protocol_valid]ity for the preloaded machine into the [projection_valid]ity. -*) - - Lemma pre_loaded_with_all_messages_validating_proj_incl - : VLSM_incl PreLoaded Xi. - Proof. - (* reduce inclusion to inclusion of finite traces. *) - apply VLSM_incl_finite_traces_characterization. - intros. - split; [|apply H]. - induction H using finite_protocol_trace_rev_ind. - - apply (finite_ptrace_empty Xi). apply initial_is_protocol. assumption. - - apply (extend_right_finite_trace_from Xi);[assumption|]. - destruct Hx as [Hvx Htx]. - split; [|assumption]. - apply projection_valid_protocol. - apply Hvalidating. - assumption. - Qed. - -(** -Given that any projection is included in the [pre_loaded_with_all_messages_vlsm] -of its component (Lemma [proj_pre_loaded_with_all_messages_incl]), we conclude -that <> and <> are trace-equal. This means that all the -byzantine behavior of a validating component is exhibited by its corresponding -projection. -*) - Lemma pre_loaded_with_all_messages_validating_proj_eq - : VLSM_eq PreLoaded Xi. - Proof. - split. - - apply pre_loaded_with_all_messages_validating_proj_incl. - - apply proj_pre_loaded_with_all_messages_incl. - Qed. - -End pre_loaded_with_all_messages_validating_proj. - -End validating_projection. - -(** ** VLSM self-validation *) - -Section validating_vlsm. - -Context - {message : Type} - (X : VLSM message) - . - -(** -Let us fix a (regular) VLSM <>. <> is (self-)validating if every [valid] -reachable [state] and <> are [protocol_state] and [protocol_message] -for that VLSM, respectively. -*) - -Definition validating_vlsm_prop - := - forall (l : label) (s : state) (om : option message), - protocol_valid (pre_loaded_with_all_messages_vlsm X) l (s, om) -> - protocol_valid X l (s, om). - -(** -In the sequel we will show that a VLSM with the [validating_vlsm_prop]erty -is trace-equal to its associated [pre_loaded_with_all_messages_vlsm], basically -meaning (due to Lemma [byzantine_pre_loaded_with_all_messages]) that all traces -with the [byzantine_trace_prop]erty associated to a validating VLSMs are also -[protocol_trace]s for that VLSM, meaning that the VLSM cannot exhibit -byzantine behavior. -*) - -Context - (Hvalidating : validating_vlsm_prop) - (PreLoaded := pre_loaded_with_all_messages_vlsm X) - . - -(** -Let <> be the [pre_loaded_with_all_messages_vlsm] associated to X. -From Lemma [vlsm_incl_pre_loaded_with_all_messages_vlsm] we know that <> is -included in <>. - -To prove the converse we use the [validating_vlsm_prop]erty to -verify the conditions of meta-lemma [VLSM_incl_finite_traces_characterization]. -*) - - Lemma pre_loaded_with_all_messages_validating_vlsm_incl - : VLSM_incl PreLoaded X. - Proof. - unfold validating_vlsm_prop in Hvalidating. - destruct X as [T [S M]]. simpl in *. - (* redcuction to inclusion of finite traces. *) - apply VLSM_incl_finite_traces_characterization. - intros. - split; [|apply H]. - destruct H as [Htr Hs]. - (* reverse induction on the length of a trace. *) - induction tr using rev_ind. - - constructor. apply initial_is_protocol. assumption. - - apply finite_protocol_trace_from_app_iff in Htr. - destruct Htr as [Htr Hx]. - specialize (IHtr Htr). - apply (finite_protocol_trace_from_app_iff (mk_vlsm M)). - split; [assumption|]. - apply (first_transition_valid (mk_vlsm M)). - apply first_transition_valid in Hx. - destruct Hx as [Hvx Htx]. - split; [|assumption]. - (* using the [validating_vlsm_prop]erty. *) - revert Hvx. - apply Hvalidating. - Qed. - -(** -We conclude that <> and <> are trace-equal. -*) - - Lemma pre_loaded_with_all_messages_validating_vlsm_eq - : VLSM_eq PreLoaded X. - Proof. - split. - - apply pre_loaded_with_all_messages_validating_vlsm_incl. - - pose (vlsm_incl_pre_loaded_with_all_messages_vlsm X) as Hincl. - destruct X as (T, (S, M)). - apply Hincl. - Qed. - -End validating_vlsm. diff --git a/theories/VLSM/Core/Validator.v b/theories/VLSM/Core/Validator.v new file mode 100644 index 00000000..c4b3e729 --- /dev/null +++ b/theories/VLSM/Core/Validator.v @@ -0,0 +1,315 @@ +From stdpp Require Import prelude. +From Coq Require Import FinFun. +From VLSM Require Import Lib.Preamble Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces. + +(** * VLSM Projection Validators + +In the sequel we fix a composite VLSM <> obtained from an indexed family +of VLSMs <> and a <>, and an index <>, corresponding to +component <>. +*) + +Section projection_validator. + +Context + {message : Type} + {index : Type} + {IndEqDec : EqDecision index} + (IM : index -> VLSM message) + (constraint : composite_label IM -> composite_state IM * option message -> Prop) + (X := composite_vlsm IM constraint) + (i : index) + (Xi := composite_vlsm_constrained_projection IM constraint i) + . + +(** +We say that the component <> of X is a validator for received messages if +non-[projection_valid]itiy implied non-component-[valid]ity (or non-reachability). +*) + +Definition projection_validator_received_messages_prop + := + forall (li : vlabel (IM i)) (si : vstate (IM i)) (mi : message), + ~ vvalid Xi li (si, Some mi) + -> ~ input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (si, Some mi). + +(** +We can slightly generalize the definition above to also include empty messages +and state it in a positive manner as the [projection_validator_prop]erty below, +requiring that [valid]ity in the component (for reachable states) implies +[projection_valid]ity. +*) + +Definition projection_validator_prop := + forall (li : vlabel (IM i)) (siomi : vstate (IM i) * option message), + input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li siomi -> + vvalid Xi li siomi. + +(** An alternative formulation of the above property with a seemingly +stronger hypothesis, states that component (IM i) is a validator for the composition constraint +if for any <> a valid state in the projection <>, <

  • > +implies that there exists a state <> whose ith component is <>, +and <> and <> are valid in <>, and <<(i,li) valid (s,om)>> in <>, that is, +we have that <
  • > in the projection <>. +*) +Definition projection_validator_prop_alt := + forall (li : vlabel (IM i)) (siom : vstate (IM i) * option message), + let (si, om) := siom in + vvalid (IM i) li siom -> + valid_state_prop Xi si -> + vvalid Xi li siom. + +(** Under validator assumptions, all reachable states for component <> are +valid states in the projection <>. +*) +Lemma validator_alt_free_states_are_projection_states + : projection_validator_prop_alt -> + forall s, + valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s -> + valid_state_prop Xi s. +Proof. + intros Hvalidator s Hs. + induction Hs using valid_state_prop_ind; + [apply initial_state_is_valid;assumption|]. + change s' with (fst (s',om')). + destruct Ht as [[_ [_ Hvalid]] Htrans]. + apply (projection_valid_implies_destination_projection_valid_state IM constraint i) with l s om om' + ; [|assumption]. + apply (Hvalidator l (s,om));assumption. +Qed. + +(** Below we show that the two definitions above are actually equivalent. +*) +Lemma projection_validator_prop_alt_iff + : projection_validator_prop_alt <-> projection_validator_prop. +Proof. + split. + - intros Hvalidator l (si, om) Hvalid. + destruct Hvalid as [Hsi [_ Hvalid]]. + apply validator_alt_free_states_are_projection_states in Hsi + ; [|assumption]. + exact (Hvalidator l (si,om) Hvalid Hsi). + - intros Hvalidator l (si, om) Hvalid HXisi. + specialize (Hvalidator l (si, om)). + apply Hvalidator. + repeat split; [| apply any_message_is_valid_in_preloaded | assumption]. + revert HXisi. + apply VLSM_incl_valid_state. + apply proj_pre_loaded_with_all_messages_incl. +Qed. + +Lemma validator_free_states_are_projection_states + : projection_validator_prop -> + forall s, + valid_state_prop (pre_loaded_with_all_messages_vlsm (IM i)) s -> + valid_state_prop Xi s. +Proof. + rewrite <- projection_validator_prop_alt_iff. + apply validator_alt_free_states_are_projection_states. +Qed. + +(** +It is easy to see that the [projection_validator_prop]erty includes the +[projection_validator_received_messages_prop]erty. +*) +Lemma projection_validator_messages_received + : projection_validator_prop -> projection_validator_received_messages_prop. +Proof. + unfold projection_validator_prop. unfold projection_validator_received_messages_prop. intros. + intro Hvalid. elim H0. clear H0. + specialize (H li (si, Some mi) Hvalid). assumption. +Qed. + +(** +We say that component <> is a [transition_validator] if any [valid] +transition (from a reachable state) in component <> can be "lifted" to +an [input_valid_transition] in <>. + +*) + +Definition transition_validator := + forall + (si : vstate (IM i)) + (omi : option message) + (li : vlabel (IM i)) + , + input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (si, omi) + -> + (exists + (s s' : vstate X) + (om' : option message), + si = s i + /\ + input_valid_transition X (existT i li) (s, omi) (s', om') + ). + +(** +Next two results show that the (simpler) [projection_validator_prop]erty +is equivalent with the [transition_validator] property. +*) + +Lemma projection_validator_messages_transitions + : projection_validator_prop -> transition_validator. +Proof. + intros Hvalidator si omi li Hpvi. + specialize (Hvalidator li (si, omi) Hpvi). clear Hpvi. + destruct Hvalidator as [s [Hsi [Hps [Hopm [Hvalid Hctr]]]]]. + destruct (vtransition X (existT i li) (s, omi)) as (s', om') eqn:Heqt. + exists s. exists s'. exists om'. + subst si. + repeat split; assumption. +Qed. + +Lemma transition_validator_messages + : transition_validator -> projection_validator_prop. +Proof. + intros Hvalidator li (si,omi) Hpvi. + specialize (Hvalidator si omi li Hpvi). clear Hpvi. + destruct Hvalidator as [s [s' [om' [Hsi [Hvalid Htransition]]]]]. + symmetry in Hsi. + exists s. split; assumption. +Qed. + +(** ** Projection validators and Byzantine behavior + +In the sequel we assume that <> has the [projection_validator_prop]erty for +component <>. Let <> be the projection of <> to component <> +and <> be the [pre_loaded_with_all_messages_vlsm] associated to component <>. +*) + +Section pre_loaded_with_all_messages_validator_proj. + Context + (Hvalidator : projection_validator_prop) + (PreLoaded := pre_loaded_with_all_messages_vlsm (IM i)) + . + +(** +We can show that <> is included in <> by applying the meta-lemma +[VLSM_incl_finite_traces_characterization], and by induction on the length +of a trace. The [projection_validator_prop]erty is used to translate +[input_valid]ity for the preloaded machine into the [projection_valid]ity. +*) + + Lemma pre_loaded_with_all_messages_validator_proj_incl + : VLSM_incl PreLoaded Xi. + Proof. + (* reduce inclusion to inclusion of finite traces. *) + apply VLSM_incl_finite_traces_characterization. + intros. + split; [|apply H]. + induction H using finite_valid_trace_rev_ind. + - apply (finite_valid_trace_from_empty Xi). apply initial_state_is_valid. assumption. + - apply (extend_right_finite_trace_from Xi);[assumption|]. + destruct Hx as [Hvx Htx]. + split; [|assumption]. + apply projection_valid_input_valid. + apply Hvalidator. + assumption. + Qed. + +(** +Given that any projection is included in the [pre_loaded_with_all_messages_vlsm] +of its component (Lemma [proj_pre_loaded_with_all_messages_incl]), we conclude +that <> and <> are trace-equal. This means that all the +byzantine behavior of a component which is a validator +is exhibited by its corresponding projection. +*) + Lemma pre_loaded_with_all_messages_validator_proj_eq + : VLSM_eq PreLoaded Xi. + Proof. + split. + - apply pre_loaded_with_all_messages_validator_proj_incl. + - apply proj_pre_loaded_with_all_messages_incl. + Qed. + +End pre_loaded_with_all_messages_validator_proj. + +End projection_validator. + +(** ** VLSM self-validation *) + +Section self_validator_vlsm. + +Context + {message : Type} + (X : VLSM message) + . + +(** +Let us fix a (regular) VLSM <>. <> is a self-validator if for any +arguments satisfying [valid] where the state is reachable in the +[pre_loaded_with_all_messages_vlsm], the arguments are also +a [valid_state] and [valid_message] for the original VLSM. +*) + +Definition self_validator_vlsm_prop + := + forall (l : label) (s : state) (om : option message), + input_valid (pre_loaded_with_all_messages_vlsm X) l (s, om) -> + input_valid X l (s, om). + +(** +In the sequel we will show that a VLSM with the [self_validator_vlsm_prop]erty +is trace-equal to its associated [pre_loaded_with_all_messages_vlsm], basically +meaning (due to Lemma [byzantine_pre_loaded_with_all_messages]) that all traces +with the [byzantine_trace_prop]erty associated to self-validator VLSMs are also +[valid_trace]s for that VLSM, meaning that the VLSM cannot exhibit +byzantine behavior. +*) + +Context + (Hvalidator : self_validator_vlsm_prop) + (PreLoaded := pre_loaded_with_all_messages_vlsm X) + . + +(** +Let <> be the [pre_loaded_with_all_messages_vlsm] associated to X. +From Lemma [vlsm_incl_pre_loaded_with_all_messages_vlsm] we know that <> is +included in <>. + +To prove the converse we use the [self_validator_vlsm_prop]erty to +verify the conditions of meta-lemma [VLSM_incl_finite_traces_characterization]. +*) + + Lemma pre_loaded_with_all_messages_self_validator_vlsm_incl + : VLSM_incl PreLoaded X. + Proof. + unfold self_validator_vlsm_prop in Hvalidator. + destruct X as [T [S M]]. simpl in *. + (* redcuction to inclusion of finite traces. *) + apply VLSM_incl_finite_traces_characterization. + intros. + split; [|apply H]. + destruct H as [Htr Hs]. + (* reverse induction on the length of a trace. *) + induction tr using rev_ind. + - constructor. apply initial_state_is_valid. assumption. + - apply finite_valid_trace_from_app_iff in Htr. + destruct Htr as [Htr Hx]. + specialize (IHtr Htr). + apply (finite_valid_trace_from_app_iff (mk_vlsm M)). + split; [assumption|]. + apply (first_transition_valid (mk_vlsm M)). + apply first_transition_valid in Hx. + destruct Hx as [Hvx Htx]. + split; [|assumption]. + (* using the [self_validator_vlsm_prop]erty. *) + revert Hvx. + apply Hvalidator. + Qed. + +(** +We conclude that <> and <> are trace-equal. +*) + + Lemma pre_loaded_with_all_messages_self_validator_vlsm_eq + : VLSM_eq PreLoaded X. + Proof. + split. + - apply pre_loaded_with_all_messages_self_validator_vlsm_incl. + - pose (vlsm_incl_pre_loaded_with_all_messages_vlsm X) as Hincl. + destruct X as (T, (S, M)). + apply Hincl. + Qed. + +End self_validator_vlsm. diff --git a/theories/VLSM/Lib/FinSetExtras.v b/theories/VLSM/Lib/FinSetExtras.v index b63f1d53..4807c571 100644 --- a/theories/VLSM/Lib/FinSetExtras.v +++ b/theories/VLSM/Lib/FinSetExtras.v @@ -8,7 +8,7 @@ Context `{FinSet A C}. Section general. - + Lemma union_size_ge_size1 (X Y : C) : size (X ∪ Y) >= size X. @@ -17,7 +17,7 @@ Section general. apply subseteq_union. set_solver. Qed. - + Lemma union_size_ge_size2 (X Y : C) : size (X ∪ Y) >= size Y. @@ -26,7 +26,7 @@ Section general. apply subseteq_union. set_solver. Qed. - + Lemma union_size_ge_average (X Y : C) : 2 * size (X ∪ Y) >= size X + size Y. @@ -35,7 +35,7 @@ Section general. specialize (union_size_ge_size2 X Y) as Hy. lia. Qed. - + Lemma difference_size_le_self (X Y : C) : size (X ∖ Y) <= size X. @@ -45,7 +45,7 @@ Section general. intros x Hx. apply elem_of_difference in Hx. intuition. Qed. - + Lemma union_size_le_sum (X Y : C) : size (X ∪ Y) <= size X + size Y. @@ -55,7 +55,7 @@ Section general. specialize (difference_size_le_self Y X). lia. Qed. - + Lemma intersection_size1 (X Y : C) : size (X ∩ Y) <= size X. @@ -63,7 +63,7 @@ Section general. apply subseteq_size with (X0 := X ∩ Y) (Y0 := X). set_solver. Qed. - + Lemma intersection_size2 (X Y : C) : size (X ∩ Y) <= size Y. @@ -71,7 +71,7 @@ Section general. apply subseteq_size with (X0 := X ∩ Y) (Y0 := Y). set_solver. Qed. - + Lemma difference_size_subset (X Y : C) (Hsub : Y ⊆ X) : @@ -85,10 +85,10 @@ Section general. - destruct (@decide (a ∈ Y)). apply elem_of_dec_slow. + apply elem_of_union. left. intuition. - + apply elem_of_union. right. set_solver. + + apply elem_of_union. right. set_solver. } assert (Htemp2 : size Y + size (X ∖ Y) = size X). { - specialize (size_union Y (X ∖ Y)) as Hun. + specialize (size_union Y (X ∖ Y)) as Hun. spec Hun. { apply elem_of_disjoint. intros a Ha Ha2. @@ -100,14 +100,14 @@ Section general. } lia. Qed. - + Lemma difference_with_intersection (X Y : C) : X ∖ Y ≡ X ∖ (X ∩ Y). Proof. set_solver. Qed. - + Lemma difference_size (X Y : C) : (Z.of_nat (size (X ∖ Y)) = size X - size (X ∩ Y))%Z. @@ -116,7 +116,7 @@ Section general. specialize (difference_size_subset X (X ∩ Y)) as Hdif. set_solver. Qed. - + Lemma difference_size_ge_disjoint_case (X Y : C) : size (X ∖ Y) >= size X - size Y. @@ -125,13 +125,13 @@ Section general. specialize (intersection_size2 X Y). lia. Qed. - + Lemma list_to_set_size (l : list A) : size (list_to_set l (C := C)) <= length l. Proof. induction l. - - simpl. + - simpl. rewrite size_empty. lia. - simpl. specialize (union_size_le_sum ({[a]}) (list_to_set l)) as Hun_size. @@ -141,11 +141,11 @@ Section general. End general. Section filter. - Context (P P2 : A → Prop) + Context (P P2 : A → Prop) `{!∀ x, Decision (P x)} `{!∀ x, Decision (P2 x)} (X Y : C). - + Lemma filter_subset (Hsub : X ⊆ Y) : filter P X ⊆ filter P Y. @@ -155,7 +155,7 @@ Section filter. apply elem_of_filter. set_solver. Qed. - + Lemma filter_subprop (Hsub : forall a, (P a -> P2 a)) : filter P X ⊆ filter P2 X. @@ -165,15 +165,15 @@ Section filter. apply elem_of_filter. intuition. Qed. - + End filter. End fin_set. Section map. - Context + Context `{FinSet A C} `{FinSet B D}. - + Lemma set_map_subset (f : A -> B) (X Y : C) @@ -185,13 +185,13 @@ Section map. apply elem_of_map. firstorder. Qed. - + Lemma set_map_size_upper_bound (f : A -> B) (X : C) : size (set_map (D := D) f X) <= size X. Proof. - unfold set_map. + unfold set_map. remember (f <$> elements X) as fX. specialize (list_to_set_size (A := B) fX) as Hsize. assert (length fX = size X). { diff --git a/theories/VLSM/Lib/ListExtras.v b/theories/VLSM/Lib/ListExtras.v index 1dbaa1b7..9890e373 100644 --- a/theories/VLSM/Lib/ListExtras.v +++ b/theories/VLSM/Lib/ListExtras.v @@ -779,7 +779,7 @@ Proof. - inversion Hin. - rewrite list_annotate_unroll,elem_of_cons in Hin. destruct Hin as [Heq | Hin]. - + subst xP. left. + + subst xP. left. + right. specialize (IHl (Forall_tl Hs)). apply IHl. assumption. Qed. @@ -1497,7 +1497,7 @@ Proof. - exists lb1. reflexivity. - rewrite elem_of_cons in Ha. destruct Ha. right. - apply elem_of_app. + apply elem_of_app. right; left; reflexivity. - specialize (IHla1 a0 lb1 b la2 lb2 H1 Ha). destruct IHla1 as [la0b Hla0b]. @@ -1762,7 +1762,7 @@ Proof. rewrite Hn0; left. - specialize (list_max_elem_of_exists l) as Hmax. spec Hmax. lia. rewrite <- eq_max. - assumption. + assumption. Qed. (* Returns all values which occur with maximum frequency in the given list. @@ -2315,7 +2315,7 @@ Lemma filter_subseteq_fn {A} P Q forall (s : list A), filter P s ⊆ filter Q s. Proof. induction s; simpl. - - rewrite filter_nil; intros x Hx; inversion Hx. + - rewrite filter_nil; intros x Hx; inversion Hx. - intro x; intros. rewrite filter_cons in H2. destruct (decide (P a)). diff --git a/theories/VLSM/Lib/ListSetExtras.v b/theories/VLSM/Lib/ListSetExtras.v index 4032a117..d4286124 100644 --- a/theories/VLSM/Lib/ListSetExtras.v +++ b/theories/VLSM/Lib/ListSetExtras.v @@ -71,7 +71,7 @@ Lemma set_eq_cons {A} : forall (a : A) (s1 s2 : set A), set_eq s1 s2 -> set_eq (a :: s1) (a :: s2). Proof. - intros. + intros. split; intros x Hx; apply elem_of_cons in Hx; destruct Hx; subst ; (left; reflexivity) || (right; apply H; assumption). Qed. @@ -108,9 +108,9 @@ Proof. pose proof set_add_intro a a (set_union [] s1) (or_introl (eq_refl a)) as Hadd. pose proof H1 a Hadd as Ha. inversion Ha. - - simpl in H. + - simpl in H. pose proof (@list_subseteq_nil A []). - rewrite <- H in H0 at 1. + rewrite <- H in H0 at 1. pose proof set_add_intro a a (set_union s1 s2) (or_introl (eq_refl a)) as Hadd. pose proof H0 a Hadd as Ha. inversion Ha. @@ -185,7 +185,7 @@ Qed. Lemma set_union_iterated_subseteq `{EqDecision A} (ss ss': list (set A)) - (Hincl : ss ⊆ ss') : + (Hincl : ss ⊆ ss') : (fold_right set_union [] ss) ⊆ (fold_right set_union [] ss'). Proof. @@ -196,7 +196,7 @@ Proof. destruct H as [x [Hx Ha]]. exists x. unfold incl in Hincl. - split; [|assumption]. + split; [|assumption]. apply Hincl. assumption. Qed. @@ -362,7 +362,7 @@ Lemma set_remove_not_elem_of `{EqDecision A} : forall x (s : list A), Proof. induction s; intros. - reflexivity. - - simpl. + - simpl. destruct (decide (x = a)). + subst; contradict H; left. + rewrite IHs; [reflexivity|]. @@ -466,7 +466,7 @@ Qed. Lemma subseteq_remove_union `{EqDecision A} : forall x (s1 s2 : list A), NoDup s1 -> - NoDup s2 -> + NoDup s2 -> (set_remove x (set_union s1 s2)) ⊆ (set_union s1 (set_remove x s2)). Proof. @@ -655,7 +655,7 @@ Proof. intros [Ha _]. apply elem_of_list_In in Ha. intro Hab. - apply elem_of_list_In in Hab. + apply elem_of_list_In in Hab. rewrite in_map_iff in Hab. destruct Hab as [_ [[= Hax _] _]]. subst. @@ -737,7 +737,7 @@ Lemma len_set_diff_decrease `{EqDecision A} (new:A) (l a b: list A) (H_new_is_relevant: new ∈ l): length (set_diff_filter l a) < length (set_diff_filter l b). Proof. - induction l; [inversion H_new_is_relevant|]; + induction l; [inversion H_new_is_relevant|]; apply elem_of_cons in H_new_is_relevant; destruct H_new_is_relevant. - subst a0. unfold set_diff_filter. @@ -915,7 +915,7 @@ Proof. assumption. Qed. -Lemma subseteq_Forall (A:Type) (P : A -> Prop) (l1 l2 : list A) : +Lemma subseteq_Forall (A:Type) (P : A -> Prop) (l1 l2 : list A) : l2 ⊆ l1 -> Forall P l1 -> Forall P l2. Proof. intros Hsub Hfor. diff --git a/theories/VLSM/Lib/Measurable.v b/theories/VLSM/Lib/Measurable.v index dc3ead93..d3a5c987 100644 --- a/theories/VLSM/Lib/Measurable.v +++ b/theories/VLSM/Lib/Measurable.v @@ -7,7 +7,7 @@ From VLSM Require Import Lib.Preamble Lib.ListExtras Lib.StdppListSet Lib.ListSe Definition pos_R := {r : R | (r > 0)%R}. Class Measurable V := { weight : V -> pos_R}. -Hint Mode Measurable ! : typeclass_instances. +Global Hint Mode Measurable ! : typeclass_instances. Definition sum_weights `{Measurable V} (l : list V) : R := fold_right (fun v r => (proj1_sig (weight v) + r)%R) 0%R l. diff --git a/theories/VLSM/Lib/Preamble.v b/theories/VLSM/Lib/Preamble.v index 319f5109..5c39945c 100644 --- a/theories/VLSM/Lib/Preamble.v +++ b/theories/VLSM/Lib/Preamble.v @@ -145,6 +145,29 @@ Proof. apply proof_irrel. Qed. +Lemma dec_sig_sigT_eq_rev + `{EqDecision A} (P : A -> Prop) {P_dec : forall x, Decision (P x)} + (F : A -> Type) + (a : A) + (b1 b2 : F a) + (e1 e2 : P a) + (pa1 := dexist a e1) + (pa2 := dexist a e2) + : @existT _ (fun pa : dsig P => F (proj1_sig pa)) pa1 b1 + = @existT _ (fun pa : dsig P => F (proj1_sig pa)) pa2 b2 -> + b1 = b2. +Proof. + subst pa1 pa2. + unfold dexist. + replace (bool_decide_pack (P a) e1) with (bool_decide_pack (P a) e2) + ; [|apply proof_irrel]. + apply inj_pair2_eq_dec. + intros x y. + destruct (decide (` x = ` y)). + - left. revert e. apply dsig_eq. + - right. intro contra. elim n. revert contra. apply dsig_eq. +Qed. + Lemma ex_out (A : Type) (P : Prop) (Q : A -> Prop): (exists x, P /\ Q x) <-> (P /\ exists x, Q x). Proof. firstorder. Qed. @@ -290,15 +313,15 @@ Qed. Class DecidablePred {A} (r : A -> Prop) := pred_dec : forall (a : A), r a \/ ~ r a. -Hint Mode DecidablePred ! ! : typeclass_instances. +Global Hint Mode DecidablePred ! ! : typeclass_instances. Class PredicateFunction {A} (r : A -> Prop) (r_fn : A -> bool) : Prop := { equiv : forall a, r a <-> r_fn a = true; predicate_function_dec :> DecidablePred r; }. -Hint Mode PredicateFunction ! ! - : typeclass_instances. -Hint Mode PredicateFunction ! - ! : typeclass_instances. +Global Hint Mode PredicateFunction ! ! - : typeclass_instances. +Global Hint Mode PredicateFunction ! - ! : typeclass_instances. Definition predicate_not {A} (p : A -> Prop) : A -> Prop := fun a => ~ p a. @@ -315,8 +338,8 @@ Qed. Class PredicateFunction2 {A B} (r : A -> B -> Prop) (r_fn : A -> B -> bool) : Prop := predicate_function2 : forall a b, r a b <-> r_fn a b = true. -Hint Mode PredicateFunction2 ! ! ! - : typeclass_instances. -Hint Mode PredicateFunction2 ! ! - ! : typeclass_instances. +Global Hint Mode PredicateFunction2 ! ! ! - : typeclass_instances. +Global Hint Mode PredicateFunction2 ! ! - ! : typeclass_instances. Lemma predicate_function2_neg : forall A B (r : A -> B -> Prop) (r_fn : A -> B -> bool), PredicateFunction2 r r_fn -> @@ -347,7 +370,7 @@ Qed. (* Reflexivity of comparison operators *) Class CompareReflexive {A} (compare : A -> A -> comparison) : Prop := compare_eq : forall x y, compare x y = Eq <-> x = y. -Hint Mode CompareReflexive ! - : typeclass_instances. +Global Hint Mode CompareReflexive ! - : typeclass_instances. (* About reflexive comparison operators *) Lemma compare_eq_refl {A} `{CompareReflexive A} : @@ -389,7 +412,7 @@ Class CompareTransitive {A} (compare : A -> A -> comparison) : Prop := compare_transitive : forall x y z comp, compare x y = comp -> compare y z = comp -> compare x z = comp. -Hint Mode CompareTransitive ! - : typeclass_instances. +Global Hint Mode CompareTransitive ! - : typeclass_instances. (* Strict-orderedness of comparison operators *) Class CompareStrictOrder {A} (compare : A -> A -> comparison) : Prop := @@ -397,7 +420,7 @@ Class CompareStrictOrder {A} (compare : A -> A -> comparison) : Prop := StrictOrder_Reflexive :> CompareReflexive compare; StrictOrder_Transitive :> CompareTransitive compare; }. -Hint Mode CompareStrictOrder ! - : typeclass_instances. +Global Hint Mode CompareStrictOrder ! - : typeclass_instances. (* Strictly-ordered comparisons give decidable equality *) Lemma compare_eq_dec {A} `{CompareStrictOrder A} : @@ -418,7 +441,7 @@ Definition eq_bool {X} `{CompareStrictOrder X} (x y : X) : bool := (* Asymmetry of comparison operators *) Class CompareAsymmetric {A} (compare : A -> A -> comparison) : Prop := compare_asymmetric : forall x y, compare x y = Lt <-> compare y x = Gt. -Hint Mode CompareAsymmetric ! - : typeclass_instances. +Global Hint Mode CompareAsymmetric ! - : typeclass_instances. (* Strictly-ordered comparisons give asymmetry *) Lemma compare_asymmetric_intro {A} `{CompareStrictOrder A} : @@ -497,7 +520,7 @@ Class StrictlyComparable (X : Type) : Type := compare : X -> X -> comparison; compare_strictorder :> CompareStrictOrder compare; }. -Hint Mode StrictlyComparable ! : typeclass_instances. +Global Hint Mode StrictlyComparable ! : typeclass_instances. Instance strictly_comparable_eq_dec `{StrictlyComparable M} : EqDecision M. diff --git a/theories/VLSM/Lib/StdppExtras.v b/theories/VLSM/Lib/StdppExtras.v index 223c541a..a708bd5a 100644 --- a/theories/VLSM/Lib/StdppExtras.v +++ b/theories/VLSM/Lib/StdppExtras.v @@ -220,7 +220,7 @@ Proof. Qed. Lemma ext_elem_of_filter {A} P Q - `{∀ (x:A), Decision (P x)} `{∀ (x:A), Decision (Q x)} + `{∀ (x:A), Decision (P x)} `{∀ (x:A), Decision (Q x)} (l : list A) : filter P l = filter Q l -> forall a, a ∈ l -> (P a <-> Q a). Proof. @@ -242,8 +242,8 @@ Proof. destruct H4; assumption. Qed. -Lemma filter_complement {X} P Q - `{∀ (x:X), Decision (P x)} `{∀ (x:X), Decision (Q x)} +Lemma filter_complement {X} P Q + `{∀ (x:X), Decision (P x)} `{∀ (x:X), Decision (Q x)} (l : list X) : filter P l = filter Q l <-> filter (fun x => ~ P x) l = filter (fun x => ~ Q x) l. diff --git a/theories/VLSM/Lib/StdppListSet.v b/theories/VLSM/Lib/StdppListSet.v index fe2dbd87..9c38d593 100644 --- a/theories/VLSM/Lib/StdppListSet.v +++ b/theories/VLSM/Lib/StdppListSet.v @@ -41,7 +41,7 @@ Lemma set_add_intro1 : Proof. simple induction x; simpl. - intro Hel; inversion Hel. - - intros a0 l IH. + - intros a0 l IH. rewrite elem_of_cons. intros [Ha0a| Hal]. * subst; destruct (decide (b = a0)); simpl; left; assumption. @@ -114,7 +114,7 @@ induction l as [|x xs Hrec]. * tauto. * intro H. destruct H. + rewrite H. left. - + rewrite elem_of_cons. + + rewrite elem_of_cons. right; apply Hrec; assumption. Qed. diff --git a/theories/VLSM/Lib/TopSort.v b/theories/VLSM/Lib/TopSort.v index cecaf582..f9b7261a 100644 --- a/theories/VLSM/Lib/TopSort.v +++ b/theories/VLSM/Lib/TopSort.v @@ -295,7 +295,7 @@ Proof. specialize (Hts a b Hab lb1 lb2 Heqb). rewrite Heqa in Heqb. assert (Ha : ~a ∈ (b :: lb2)). - { intro Ha. apply Hts. + { intro Ha. apply Hts. rewrite elem_of_cons in Ha. destruct Ha; try assumption. subst. apply (preceeds_irreflexive preceeds P b Hpa) in Hab. @@ -417,7 +417,7 @@ Proof. apply elem_of_app. left; assumption. } - specialize (Hpc b Hb' a Hab). + specialize (Hpc b Hb' a Hab). apply elem_of_app in Hpc. destruct Hpc as [Ha | Ha]; try assumption. rewrite elem_of_cons in Ha. @@ -489,7 +489,7 @@ Proof. specialize (set_remove_length min l Hin). rewrite <- Heql'. rewrite <- H0. intro Hlen. specialize (IHn Hlen). - split; intros x Hx; rewrite elem_of_cons in Hx; + split; intros x Hx; rewrite elem_of_cons in Hx; destruct Hx as [Heq|Hinx]; try (subst x). + right. apply IHn. left. + destruct (decide (x = min)); try subst x. @@ -570,7 +570,7 @@ Lemma top_sort_sorted : topologically_sorted preceeds (top_sort l). Proof. intro a; intros. intro Ha2. - assert (Ha : a ∈ l). { + assert (Ha : a ∈ l). { apply top_sort_set_eq. rewrite Heq. simpl. apply elem_of_app. right. right. assumption. } @@ -642,7 +642,7 @@ Proof. * subst l'. rewrite elem_of_cons in Ha. destruct Ha as [Heqa | Ha']. - -- subst a0. + -- subst a0. rewrite elem_of_cons in i. contradict i. left; reflexivity. @@ -678,7 +678,7 @@ Qed. Definition get_maximal_element := ListExtras.last_error (top_sort l). -Lemma maximal_element_in +Lemma maximal_element_in (a : A) (Hmax : get_maximal_element = Some a) : a ∈ l. @@ -708,7 +708,7 @@ Proof. intuition. Qed. -Lemma get_maximal_element_correct +Lemma get_maximal_element_correct (a max : A) (Hina : a ∈ l) (Hmax : get_maximal_element = Some max) : @@ -718,9 +718,9 @@ Proof. unfold topologically_sorted in Htop. intros contra. specialize (Htop max a contra). - + assert (Hinmax: max ∈ l) by (apply maximal_element_in; intuition). - assert (Hinatop : a ∈ (top_sort l)) by (apply Hseteq; intuition). + assert (Hinatop : a ∈ (top_sort l)) by (apply Hseteq; intuition). apply elem_of_list_split in Hinatop. destruct Hinatop as [prefA [sufA HeqA]]. unfold get_maximal_element in Hmax. @@ -728,7 +728,7 @@ Proof. - rewrite HeqA in Hmax. specialize (last_error_is_last prefA a) as Hlast. assert (a = max) by intuition congruence. - subst a. + subst a. specialize StrictOrder_Irreflexive as Hirr. unfold Irreflexive in Hirr. unfold complement in Hirr. unfold Reflexive in Hirr. diff --git a/theories/VLSM/dune b/theories/VLSM/dune index bd4b388a..210a39ab 100644 --- a/theories/VLSM/dune +++ b/theories/VLSM/dune @@ -1,7 +1,6 @@ (coq.theory (name VLSM) (package coq-vlsm) - (synopsis "Coq core definitions for VLSMs") - (flags -w -deprecated-hint-without-locality)) + (synopsis "Core definitions in Coq for VLSMs")) (include_subdirs qualified)