From b33cf1b86d8e1bf3864f9953b1fb5ab418168df2 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Wed, 5 Jan 2022 15:19:37 +0200 Subject: [PATCH 01/18] induced component projection --- theories/VLSM/Core/ProjectionTraces.v | 188 +++++++++++++++++++++----- 1 file changed, 151 insertions(+), 37 deletions(-) diff --git a/theories/VLSM/Core/ProjectionTraces.v b/theories/VLSM/Core/ProjectionTraces.v index d810adc2..cfcd72c9 100644 --- a/theories/VLSM/Core/ProjectionTraces.v +++ b/theories/VLSM/Core/ProjectionTraces.v @@ -264,13 +264,158 @@ Definition composite_project_label | _ => None end. -Definition from_projection - (a : composite_transition_item IM) - : Prop - := j = projT1 (l a). +(** The [composite_vlsm_constraint_projection] is [VLSM_eq]ual (trace-equivalent) +to the [projection_induced_vlsm] by the [composite_project_label] and the +projection of the state to the component. +*) +Lemma composite_vlsm_constrained_projection_is_induced + : VLSM_eq Xj + (projection_induced_vlsm X (type (IM j)) + composite_project_label (fun s => s j) + (lift_to_composite_label IM j) (lift_to_composite_state IM j)). +Proof. + apply VLSM_eq_incl_iff. + split. + - apply basic_VLSM_strong_incl. + + intros s Hs. + exists (lift_to_composite_state IM j s). + split; [apply state_update_eq|]. + apply (lift_to_composite_state_initial IM). + assumption. + + intros m [[im Him] <-]. assumption. + + intros l s iom [sX [<- Hv]]. + exists (existT j l), sX. + intuition. + unfold composite_project_label. + case_decide; [|contradiction]. + cbn in H |- *. + replace H with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. + + intros l s iom s' oom. + cbn. + unfold lift_to_composite_state at 1. + rewrite state_update_eq. + cbn. + intros Ht. + setoid_rewrite Ht. + rewrite state_update_eq. + reflexivity. + - cbn. apply basic_VLSM_strong_incl. + + intros s [sX [<- HsX]]. + apply (HsX j). + + intros m Him. + exists (exist _ m Him). + reflexivity. + + intros l s iom [[i li] [sX [HlX [<- Hv]]]]. + exists sX. + split; [reflexivity|]. + unfold composite_project_label in HlX. + simpl in *. + case_decide; [|congruence]. + subst i. + apply Some_inj in HlX. + cbv in HlX. + subst li. + assumption. + + intros l s iom s' oom. + cbn. + unfold lift_to_composite_state at 1. + rewrite state_update_eq. + destruct (vtransition _ _ _) as (si', om'). + rewrite state_update_eq. + intuition. +Qed. + +Lemma component_label_projection_lift + : induced_projection_label_lift_prop X (type (IM j)) composite_project_label + (lift_to_composite_label IM j). +Proof. + intros lj. + unfold composite_project_label. + case_decide as Hj; [|contradiction]. + cbn in Hj |- *. + replace Hj with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. +Qed. + +Lemma component_state_projection_lift + : induced_projection_state_lift_prop X (type (IM j)) (fun s => s j) + (lift_to_composite_state IM j). +Proof. + intros sj. + apply state_update_eq. +Qed. -Instance dec_from_projection (a : transition_item) : Decision (from_projection a) := - decide (from_projection a). +Lemma component_transition_projection_None + : weak_projection_transition_consistency_None X (type (IM j)) + composite_project_label (λ s : vstate X, s j). +Proof. + intros (i, li) HlX sX iom s'X oom [_ Ht]. + cbn in Ht. + destruct (vtransition _ _ _) as (si', om'). + inversion Ht. subst. + apply state_update_neq. + unfold composite_project_label in HlX. + simpl in HlX. + case_decide; congruence. +Qed. + +Lemma component_transition_projection_Some + : induced_projection_transition_consistency_Some X (type (IM j)) + composite_project_label (λ s : vstate X, s j). +Proof. + intros [j1 lj1] [j2 lj2] lj. + unfold composite_project_label. + cbn. + case_decide as Hj1; [|congruence]. + subst j1. + intros Hlj1. cbv in Hlj1. + apply Some_inj in Hlj1. + subst lj1. + case_decide as Hj2; [|congruence]. + subst j2. + intros Hlj2. cbv in Hlj2. + apply Some_inj in Hlj2. + subst lj2. + intros sX1 sX2 <- iom. + destruct (vtransition _ _ _) as (si', om'). + intros sX1' oom1 Ht1. inversion Ht1. subst. clear Ht1. + intros sX2' oom2 Ht2. inversion Ht2. subst. clear Ht2. + split; [|reflexivity]. + rewrite !state_update_eq. + reflexivity. +Qed. + +(** The [projection_induced_vlsm] by the [composite_project_label] and the +projection of the state to the component is indeed a [VLSM_projection]. +*) +Lemma induced_component_projection + : VLSM_projection X + (projection_induced_vlsm X (type (IM j)) + composite_project_label (fun s => s j) + (lift_to_composite_label IM j) (lift_to_composite_state IM j)) + composite_project_label (fun s => s j). +Proof. + apply projection_induced_vlsm_is_projection. + - apply component_transition_projection_None. + - apply basic_weak_projection_transition_consistency_Some. + + apply component_label_projection_lift. + + apply component_state_projection_lift. + + apply component_transition_projection_Some. +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. + constructor. + - apply induced_component_projection. + - intros isX trX HtrX. + apply (VLSM_eq_finite_valid_trace composite_vlsm_constrained_projection_is_induced). + apply (VLSM_projection_finite_valid_trace induced_component_projection). + assumption. +Qed. Lemma initial_state_projection (s : vstate X) @@ -296,37 +441,6 @@ Proof. 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. - - destruct lX as (i, li). - unfold composite_project_label in HlX. - simpl in HlX. case_decide; [|congruence]. - subst. inversion HlX. subst. clear HlX. - eexists; intuition. - - destruct lX as (i, li). - unfold composite_project_label in H. - simpl in H. destruct (decide _); [|congruence]. - subst. inversion H. subst. clear H. - apply proj2 in H0. - simpl in H0. - cbn. - destruct (vtransition _ _ _) as (si', _om'). - inversion H0. rewrite state_update_eq. reflexivity. - - destruct lX as (i, li). - unfold composite_project_label in H. - simpl in H. destruct (decide _); [congruence|]. - clear H. apply proj2 in H0. cbn in H0. - destruct (vtransition _ _ _) as (si', _om'). - inversion H0. rewrite state_update_neq by congruence. - reflexivity. - - apply initial_state_projection. assumption. - - apply valid_message_projection. apply Hv. -Qed. - (* The projection of a finite valid trace remains a valid trace *) Lemma finite_valid_trace_projection (s : vstate X) From 02d983cf87f3194b995b1f15b0eab20af1552cbd Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Wed, 5 Jan 2022 15:20:21 +0200 Subject: [PATCH 02/18] Abstract validators --- theories/VLSM/Core/ByzantineTraces.v | 6 +- .../ByzantineTraces/FixedSetByzantineTraces.v | 2 +- .../ByzantineTraces/LimitedByzantineTraces.v | 2 +- theories/VLSM/Core/VLSM.v | 6 +- theories/VLSM/Core/VLSMProjections.v | 78 ++- theories/VLSM/Core/Validator.v | 562 +++++++++++------- 6 files changed, 432 insertions(+), 224 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces.v index 630f7671..f36fbc96 100644 --- a/theories/VLSM/Core/ByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces.v @@ -369,14 +369,14 @@ Lemma validator_component_byzantine_fault_tolerance (IM : index -> VLSM message) (constraint : composite_label IM -> composite_state IM * option message -> Prop) (i : index) - (Hvalidator: projection_validator_prop IM constraint i) + (Hvalidator: component_projection_validator_prop IM constraint i) : forall tr, byzantine_trace_prop (IM i) tr -> valid_trace_prop (composite_vlsm_constrained_projection IM constraint i) tr. Proof. intros tr Htr. apply (VLSM_incl_valid_trace - (pre_loaded_with_all_messages_validator_proj_incl _ _ _ Hvalidator)). + (pre_loaded_with_all_messages_validator_component_proj_incl _ _ _ Hvalidator)). revert Htr. apply byzantine_pre_loaded_with_all_messages. Qed. @@ -399,7 +399,7 @@ Section composite_validator_byzantine_traces. (X := composite_vlsm IM constraint) (PreLoadedX := pre_loaded_with_all_messages_vlsm X) (FreeX := free_composite_vlsm IM) - (Hvalidator: forall i : index, projection_validator_prop IM constraint i) + (Hvalidator: forall i : index, component_projection_validator_prop IM constraint i) . (** diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index 13d9a8d5..f4839b56 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -821,7 +821,7 @@ End assuming_initial_messages_lift. Context (Hvalidator: forall i : index, i ∉ selection -> - projection_validator_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) + component_projection_validator_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) . Lemma validator_fixed_non_byzantine_vlsm_lift_initial_message diff --git a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v index d95e448b..25e1fd8a 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -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) - (Hvalidator: forall i : index, projection_validator_prop IM limited_constraint i) + (Hvalidator: forall i : index, component_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)) diff --git a/theories/VLSM/Core/VLSM.v b/theories/VLSM/Core/VLSM.v index 8c9aa438..288ab3a2 100644 --- a/theories/VLSM/Core/VLSM.v +++ b/theories/VLSM/Core/VLSM.v @@ -671,10 +671,9 @@ given inputs and that they have a [valid_state] and a [valid_message]. (l : label) (som : state * option message) (Hv : input_valid l som) - : exists (som' : state * option message), + : forall som', transition l som = som' -> input_valid_transition l som som'. Proof. - exists (transition l som). repeat split; assumption. Qed. @@ -686,7 +685,8 @@ given inputs and that they have a [valid_state] and a [valid_message]. input_valid_transition l som som'. Proof. split. - - apply input_valid_can_transition. + - eexists. + apply input_valid_can_transition; [assumption|reflexivity]. - intros [som' Hivt]. apply input_valid_transition_valid with som'. assumption. diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index 2602eff5..d843b32f 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -130,6 +130,20 @@ Proof. inversion HtX. subst. assumption. Qed. +Lemma VLSM_weak_partial_projection_input_valid + : forall lX sX inX, input_valid X lX (sX, inX) -> + forall sX' outX, vtransition X lX (sX, inX) = (sX', outX) -> + forall sY itemY, + trace_project (sX, [{| l := lX; input := inX; destination := sX'; output := outX |}]) + = Some (sY, [itemY]) -> + input_valid Y (l itemY) (sY, input itemY). +Proof. + intros lX sX inX HvX sX' outX HtX sY itemY Hpr. + eapply proj1, VLSM_weak_partial_projection_input_valid_transition + ; [eassumption|]. + split; assumption. +Qed. + End weak_partial_projection_properties. Section partial_projection_properties. @@ -195,6 +209,15 @@ Definition VLSM_partial_projection_input_valid_transition input_valid_transition Y (l itemY) (sY, input itemY) (destination itemY, output itemY) := VLSM_weak_partial_projection_input_valid_transition VLSM_partial_projection_weaken. +Definition VLSM_partial_projection_input_valid + : forall lX sX inX, input_valid X lX (sX, inX) -> + forall sX' outX, vtransition X lX (sX, inX) = (sX', outX) -> + forall sY itemY, + trace_project (sX, [{| l := lX; input := inX; destination := sX'; output := outX |}]) + = Some (sY, [itemY]) -> + input_valid Y (l itemY) (sY, input itemY) + := VLSM_weak_partial_projection_input_valid VLSM_partial_projection_weaken. + End partial_projection_properties. End VLSM_partial_projection. @@ -653,6 +676,17 @@ Proof. simpl. rewrite H. reflexivity. Qed. +Lemma VLSM_weak_projection_input_valid + : forall lX lY, label_project lX = Some lY -> + forall sX im, + input_valid X lX (sX, im) -> input_valid Y lY (state_project sX, im). +Proof. + intros lX lY Hpr sX im HvX. + destruct (vtransition X lX (sX, im)) eqn:HtX. + eapply proj1, VLSM_weak_projection_input_valid_transition, input_valid_can_transition + ; [eassumption|assumption|exact HtX]. +Qed. + Lemma VLSM_weak_projection_finite_valid_trace_from_to : forall sX s'X 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). @@ -797,6 +831,12 @@ Definition VLSM_projection_input_valid_transition 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_input_valid + : forall lX lY, label_project lX = Some lY -> + forall s im, + input_valid X lX (s, im) -> input_valid Y lY (state_project s, im) + := VLSM_weak_projection_input_valid VLSM_projection_weaken. + Definition VLSM_projection_finite_valid_trace_from_to : forall sX s'X 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_projection_trace_project Hsimul trX) @@ -1275,6 +1315,14 @@ Proof. ; [assumption|reflexivity]. Qed. +Lemma VLSM_weak_full_projection_input_valid l s im + : input_valid X l (s,im) -> input_valid Y (label_project l) (state_project s,im). +Proof. + intros. + eapply (VLSM_weak_projection_input_valid VLSM_weak_full_projection_is_projection) + ; [reflexivity|assumption]. +Qed. + Lemma VLSM_weak_full_projection_infinite_valid_trace_from : forall sX trX, infinite_valid_trace_from X sX trX -> @@ -1289,18 +1337,6 @@ Proof. assumption. Qed. -Lemma VLSM_weak_full_projection_input_valid - : forall l 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 (Hivt : input_valid_transition X l (s, im) (s', om)). - { split; assumption. } - apply VLSM_weak_full_projection_input_valid_transition in Hivt. - apply Hivt. -Qed. - Lemma VLSM_weak_full_projection_can_produce (s : state) (om : option message) @@ -2978,6 +3014,24 @@ Proof. assumption. Qed. +(** When we have a [VLSM_projection] to the [projection_induced_vlsm], +[valid]ity is [input_valid]ity. +*) +Lemma induced_projection_valid_is_input_valid + (Hproj : VLSM_projection X projection_induced_vlsm label_project state_project) + l s om + : vvalid projection_induced_vlsm l (s, om) -> input_valid projection_induced_vlsm l (s,om). +Proof. + intro Hv. + destruct (id Hv) as [lX [sX [HlX [<- [Hps [Hopm _]]]]]]. + repeat split. + - apply (VLSM_projection_valid_state Hproj). assumption. + - destruct om as [m|]; [|apply option_valid_message_None]. + apply option_initial_message_is_valid. + assumption. + - assumption. +Qed. + Section projection_induced_friendliness. Context diff --git a/theories/VLSM/Core/Validator.v b/theories/VLSM/Core/Validator.v index c4b3e729..8e720a3e 100644 --- a/theories/VLSM/Core/Validator.v +++ b/theories/VLSM/Core/Validator.v @@ -1,147 +1,73 @@ From stdpp Require Import prelude. From Coq Require Import FinFun. -From VLSM Require Import Lib.Preamble Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces. +From VLSM.Lib Require Import Preamble ListExtras. +From VLSM.Core Require Import VLSM VLSMProjections Composition 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 <>. +In the sequel we fix a VLSMs <> and <> and a [VLSM_projection] +of <> into <>, the [pre_loaded_with_all_messages_vlsm] of <>. *) 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) - . + {message : Type} + {X Y : VLSM message} + {label_project : vlabel X -> option (vlabel Y)} + {state_project : vstate X -> vstate Y} + (PreY := pre_loaded_with_all_messages_vlsm Y) + (Hproj : VLSM_projection X PreY label_project state_project) + . (** -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). +We say that <> is a validator for received messages if +non-[projection_induced_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). +Definition projection_validator_received_messages_prop := + forall (li : vlabel Y) (si : vstate Y) (mi : message), + ~ projection_induced_valid X (type Y) label_project state_project li (si, Some mi) + -> ~ input_valid PreY 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. +[projection_induced_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. + forall li si omi, + input_valid PreY li (si,omi) -> + projection_induced_valid X (type Y) label_project state_project li (si,omi). (** 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. + : 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. + unfold projection_validator_prop, projection_validator_received_messages_prop. + intuition. Qed. (** -We say that component <> is a [transition_validator] if any [valid] -transition (from a reachable state) in component <> can be "lifted" to +We say that <> is a [transition_validator] if any [valid] +transition (from a reachable state) in <> 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') - ). + forall lY sY omi, input_valid PreY lY (sY, omi) -> + exists lX sX sX' om', + input_valid_transition X lX (sX, omi) (sX', om') /\ + label_project lX = Some lY /\ + state_project sX = sY. + (** Next two results show that the (simpler) [projection_validator_prop]erty @@ -149,91 +75,320 @@ is equivalent with the [transition_validator] property. *) Lemma projection_validator_messages_transitions - : projection_validator_prop -> transition_validator. + : 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. + intros Hvalidator li si omi Hpvi. + specialize (Hvalidator _ _ _ Hpvi) + as [l [s [Hli [Hsi [Hps [Hopm Hvalid]]]]]]. + exists l, s. + unfold input_valid_transition. + destruct (transition _ _ ) as (s', om'). + exists s', om'. + repeat split; assumption. Qed. Lemma transition_validator_messages - : transition_validator -> projection_validator_prop. + : 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. + intros Hvalidator li si omi Hpvi. + specialize (Hvalidator _ _ _ Hpvi) + as [l [s [s' [om' [[Hvalid Htransition] [Hli Hsi]]]]]]. + exists l, s. + intuition. 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 <>. +In the sequel we assume that the [projection_induced_vlsm_is_projection] and +that initial states of <> can be lifted to <>. +*) + +Section induced_projection_validators. + +Context + (Htransition_None : weak_projection_transition_consistency_None _ _ label_project state_project) + (label_lift : vlabel Y -> vlabel X) + (state_lift : vstate Y -> vstate X) + (Xi := projection_induced_vlsm X (type Y) + label_project state_project label_lift state_lift) + (Hlabel_lift : induced_projection_label_lift_prop _ _ label_project label_lift) + (Hstate_lift : induced_projection_state_lift_prop _ _ state_project state_lift) + (Hinitial_lift : strong_full_projection_initial_state_preservation Y X state_lift) + (Htransition_consistency : induced_projection_transition_consistency_Some _ _ label_project state_project) + (Htransition_Some : weak_projection_transition_consistency_Some _ _ label_project state_project label_lift state_lift + := basic_weak_projection_transition_consistency_Some _ _ _ _ _ _ Hlabel_lift Hstate_lift Htransition_consistency) + (Hproji := projection_induced_vlsm_is_projection _ _ _ _ _ _ Htransition_None Htransition_Some) + . + +(** If there is a [VLSM_projection] from <> to <> and the +[projection_induced_vlsm_is_projection], then a [transition] [valid] for the +[projection_induced_vlsm] has the same output as the transition on <>. +*) +Lemma projection_induced_valid_transition_eq + : forall l s om, vvalid Xi l (s, om) -> + vtransition Xi l (s, om) = vtransition Y l (s, om). +Proof. + intros l s im [lX [sX [Hlx [<- Hv]]]]. + replace (vtransition Y _ _) with + (state_project (vtransition X lX (sX, im)).1, (vtransition X lX (sX, im)).2). + - eapply proj2, (VLSM_projection_input_valid_transition Hproji) with (lY := l) + ; [eassumption|]. + split; [assumption|]. + apply injective_projections; reflexivity. + - symmetry. + eapply proj2, (VLSM_projection_input_valid_transition Hproj) with (lY := l) + ; [eassumption|]. + split; [assumption|]. + apply injective_projections; reflexivity. +Qed. + +Lemma induced_projection_incl_preloaded_with_all_messages + : VLSM_incl Xi PreY. +Proof. + apply basic_VLSM_incl. + - intros is [s [<- Hs]]. + apply (VLSM_projection_initial_state Hproj). + assumption. + - intro; intros. apply any_message_is_valid_in_preloaded. + - intros l s om [_ [_ [lX [sX [Hlx [<- Hv]]]]]] _ _. + simpl. + eapply proj2, proj2, (VLSM_projection_input_valid Hproj); eassumption. + - intros l s im s' om [[_ [_ HvXi]] HtXi]. + setoid_rewrite <- HtXi. + symmetry. + apply projection_induced_valid_transition_eq. + assumption. +Qed. + +(** An alternative formulation of the [projection_validator_prop]erty with a +seemingly stronger hypothesis, states that <> is a validator for <> +if for any <> a valid state in the projection <>, <
  • > +implies that <
  • > in the projection <> (i.e., +[projection_induced_valid]ity). *) +Definition projection_validator_prop_alt := + forall li si iom, + vvalid Y li (si, iom) -> + valid_state_prop Xi si -> + vvalid Xi li (si, iom). + +(** Under validator assumptions, all reachable states for component <> are +valid states in the induced projection <>. +*) +Lemma validator_alt_free_states_are_projection_states + : projection_validator_prop_alt -> + forall s, + valid_state_prop (pre_loaded_with_all_messages_vlsm Y) s -> + valid_state_prop Xi s. +Proof. + intros Hvalidator sY Hs. + induction Hs using valid_state_prop_ind. + - apply initial_state_is_valid. + exists (state_lift s). + split; [apply Hstate_lift|]. + apply Hinitial_lift. + assumption. + - destruct Ht as [[_ [_ Hvalid]] Htrans]. + specialize (Hvalidator _ _ _ Hvalid IHHs) + as [lX [sX [HlX [HsX HvX]]]]. + replace s' with (state_project (vtransition X lX (sX, om)).1). + + eapply input_valid_transition_destination, + (VLSM_projection_input_valid_transition Hproji); + [exact HlX|]. + split; [exact HvX|]. + apply injective_projections; reflexivity. + + assert (HivtX : input_valid_transition X lX (sX, om) (vtransition X lX (sX, om))) + by (split; [assumption|reflexivity]). + destruct (vtransition _ _ _) as (sX', _om'). + apply (VLSM_projection_input_valid_transition Hproj) with (lY := l) + in HivtX as [_ Hs'] + ; [|assumption]. + rewrite HsX in Hs'. + destruct Y as (TY, (SY, MY)). + cbv in Htrans, Hs'. + rewrite Htrans in Hs'. + inversion Hs'. + reflexivity. +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. + apply Hvalidator; [apply Hvalid|]. + apply validator_alt_free_states_are_projection_states; [assumption..|]. + apply Hvalid. + - intros Hvalidator l si om Hvalid HXisi. + apply Hvalidator. + repeat split; [| apply any_message_is_valid_in_preloaded | assumption]. + revert HXisi. + apply VLSM_incl_valid_state. + apply induced_projection_incl_preloaded_with_all_messages. +Qed. + +Lemma validator_free_states_are_projection_states + : projection_validator_prop -> + forall s, + valid_state_prop (pre_loaded_with_all_messages_vlsm Y) s -> + valid_state_prop Xi s. +Proof. + rewrite <- projection_validator_prop_alt_iff by assumption. + apply validator_alt_free_states_are_projection_states. +Qed. Section pre_loaded_with_all_messages_validator_proj. - Context - (Hvalidator : projection_validator_prop) - (PreLoaded := pre_loaded_with_all_messages_vlsm (IM i)) - . + Context + (Hvalidator : projection_validator_prop) + . (** -We can show that <> is included in <> by applying the meta-lemma +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. +[input_valid]ity for the PreY 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. +Lemma pre_loaded_with_all_messages_validator_proj_incl + : VLSM_incl PreY Xi. +Proof. + (* reduce inclusion to inclusion of finite traces. *) + apply VLSM_incl_finite_traces_characterization. + intros sY trY HtrY. + split; cycle 1. + - exists (state_lift sY). + split; [apply Hstate_lift|]. + apply Hinitial_lift. + apply HtrY. + - induction HtrY using finite_valid_trace_rev_ind. + + apply (finite_valid_trace_from_empty Xi). + apply initial_state_is_valid. + exists (state_lift si). + split; [apply Hstate_lift|]. + apply Hinitial_lift. + assumption. + + apply (extend_right_finite_trace_from Xi);[assumption|]. + split. + * apply induced_projection_valid_is_input_valid; [assumption|]. + apply Hvalidator. + apply Hx. + * replace (sf, _) with (vtransition Y l (finite_trace_last si tr, iom)) + by apply Hx. + apply projection_induced_valid_transition_eq. + apply Hvalidator. + apply Hx. +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 +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. +Lemma pre_loaded_with_all_messages_validator_proj_eq + : VLSM_eq PreY Xi. +Proof. + apply VLSM_eq_incl_iff. + split. + - apply pre_loaded_with_all_messages_validator_proj_incl. + - apply induced_projection_incl_preloaded_with_all_messages. +Qed. End pre_loaded_with_all_messages_validator_proj. +End induced_projection_validators. + End projection_validator. +(** ** Validator properties for the [component_projection]. + +In this section we specialize the validator-related results to the +components of a composition. +*) + +Section component_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) + (PreXi := pre_loaded_with_all_messages_vlsm (IM i)) + . + +(** +We say that the component <> of X is a validator for received messages if +if [valid]ity in the component (for reachable states) implies [projection_valid]ity. +*) +Definition component_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. + +Lemma component_projection_to_preloaded + : VLSM_projection X PreXi (composite_project_label IM i) (fun s => s i). +Proof. + constructor. + - apply component_projection. + - intros sX trX HtrX. + apply (VLSM_projection_finite_valid_trace (preloaded_component_projection IM i)). + revert HtrX. + apply VLSM_incl_finite_valid_trace. + apply constraint_preloaded_free_incl with (constraint0 := constraint). +Qed. + +(** Assuming the [component_projection_validator_prop]erty, the component +[pre_loaded_with_all_messages_vlsm] is [VLSM_eq]ual (trace-equivalent) with +its corresponding [projection_induced_vlsm]. +*) +Lemma pre_loaded_with_all_messages_validator_component_proj_eq + (Hvalidator : component_projection_validator_prop) + : VLSM_eq PreXi Xi. +Proof. + apply VLSM_eq_trans with + (machine (projection_induced_vlsm X (type (IM i)) + (composite_project_label IM i) (fun s => s i) + (lift_to_composite_label IM i) (lift_to_composite_state IM i))) + ; [|apply VLSM_eq_sym; apply composite_vlsm_constrained_projection_is_induced]. + apply pre_loaded_with_all_messages_validator_proj_eq. + - apply component_projection_to_preloaded. + - apply component_transition_projection_None. + - apply component_label_projection_lift. + - apply component_state_projection_lift. + - intro s. apply (lift_to_composite_state_initial IM). + - apply component_transition_projection_Some. + - intros li si omi Hiv. + apply Hvalidator in Hiv as [sX [<- HivX]]. + exists (existT i li), sX. + intuition. + unfold composite_project_label. + simpl. + case_decide as Hi; [|contradiction]. + replace Hi with (@eq_refl index i) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. +Qed. + +Definition pre_loaded_with_all_messages_validator_component_proj_incl + (Hvalidator : component_projection_validator_prop) + : VLSM_incl PreXi Xi := + VLSM_eq_proj1 (pre_loaded_with_all_messages_validator_component_proj_eq Hvalidator). + +End component_projection_validator. + (** ** VLSM self-validation *) Section self_validator_vlsm. Context - {message : Type} - (X : VLSM message) - . + {message : Type} + (X : VLSM message) + . (** Let us fix a (regular) VLSM <>. <> is a self-validator if for any @@ -242,11 +397,10 @@ arguments satisfying [valid] where the state is reachable in the 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). +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 @@ -258,58 +412,58 @@ byzantine behavior. *) Context - (Hvalidator : self_validator_vlsm_prop) - (PreLoaded := pre_loaded_with_all_messages_vlsm X) - . + (Hvalidator : self_validator_vlsm_prop) + (PreX := pre_loaded_with_all_messages_vlsm X) + . (** -Let <> be the [pre_loaded_with_all_messages_vlsm] associated to 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 <>. +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. +Lemma pre_loaded_with_all_messages_self_validator_vlsm_incl + : VLSM_incl PreX 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. +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. +Lemma pre_loaded_with_all_messages_self_validator_vlsm_eq + : VLSM_eq PreX 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. From 0db67fd0b3752d44dec0466e81d98208be612609 Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Thu, 20 Jan 2022 15:43:51 +0200 Subject: [PATCH 03/18] Final touches --- theories/VLSM/Core/ProjectionTraces.v | 2 +- theories/VLSM/Core/VLSMProjections.v | 2 +- theories/VLSM/Core/Validator.v | 26 +++++++++----------------- 3 files changed, 11 insertions(+), 19 deletions(-) diff --git a/theories/VLSM/Core/ProjectionTraces.v b/theories/VLSM/Core/ProjectionTraces.v index cfcd72c9..80cc1c0d 100644 --- a/theories/VLSM/Core/ProjectionTraces.v +++ b/theories/VLSM/Core/ProjectionTraces.v @@ -402,7 +402,7 @@ Proof. + apply component_label_projection_lift. + apply component_state_projection_lift. + apply component_transition_projection_Some. -Qed. +Qed. (** The projection on component <> of valid traces from <> is valid for the <>th projection. diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index d843b32f..177692ab 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -686,7 +686,7 @@ Proof. eapply proj1, VLSM_weak_projection_input_valid_transition, input_valid_can_transition ; [eassumption|assumption|exact HtX]. Qed. - + Lemma VLSM_weak_projection_finite_valid_trace_from_to : forall sX s'X 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). diff --git a/theories/VLSM/Core/Validator.v b/theories/VLSM/Core/Validator.v index 8e720a3e..ea84318d 100644 --- a/theories/VLSM/Core/Validator.v +++ b/theories/VLSM/Core/Validator.v @@ -58,16 +58,13 @@ Qed. We say that <> is a [transition_validator] if any [valid] transition (from a reachable state) in <> can be "lifted" to an [input_valid_transition] in <>. - *) - Definition transition_validator := forall lY sY omi, input_valid PreY lY (sY, omi) -> exists lX sX sX' om', input_valid_transition X lX (sX, omi) (sX', om') /\ label_project lX = Some lY /\ state_project sX = sY. - (** Next two results show that the (simpler) [projection_validator_prop]erty @@ -109,8 +106,7 @@ Context (Htransition_None : weak_projection_transition_consistency_None _ _ label_project state_project) (label_lift : vlabel Y -> vlabel X) (state_lift : vstate Y -> vstate X) - (Xi := projection_induced_vlsm X (type Y) - label_project state_project label_lift state_lift) + (Xi := projection_induced_vlsm X (type Y) label_project state_project label_lift state_lift) (Hlabel_lift : induced_projection_label_lift_prop _ _ label_project label_lift) (Hstate_lift : induced_projection_state_lift_prop _ _ state_project state_lift) (Hinitial_lift : strong_full_projection_initial_state_preservation Y X state_lift) @@ -161,10 +157,11 @@ Proof. Qed. (** An alternative formulation of the [projection_validator_prop]erty with a -seemingly stronger hypothesis, states that <> is a validator for <> -if for any <> a valid state in the projection <>, <
  • > -implies that <
  • > in the projection <> (i.e., -[projection_induced_valid]ity). +seemingly stronger hypothesis, states that <> is a validator for <> if +for any <
  • >, <>, <> such that <
  • > in <> +and <> is a valid state in the induced projection <>, +implies that <
  • > in the induced projection <> +(i.e., [projection_induced_valid]ity). *) Definition projection_validator_prop_alt := forall li si iom, @@ -177,9 +174,7 @@ valid states in the induced projection <>. *) Lemma validator_alt_free_states_are_projection_states : projection_validator_prop_alt -> - forall s, - valid_state_prop (pre_loaded_with_all_messages_vlsm Y) s -> - valid_state_prop Xi s. + forall s, valid_state_prop PreY s -> valid_state_prop Xi s. Proof. intros Hvalidator sY Hs. induction Hs using valid_state_prop_ind. @@ -211,8 +206,7 @@ Proof. reflexivity. Qed. -(** Below we show that the two definitions above are actually equivalent. -*) +(** 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. @@ -231,9 +225,7 @@ Qed. Lemma validator_free_states_are_projection_states : projection_validator_prop -> - forall s, - valid_state_prop (pre_loaded_with_all_messages_vlsm Y) s -> - valid_state_prop Xi s. + forall s, valid_state_prop PreY s -> valid_state_prop Xi s. Proof. rewrite <- projection_validator_prop_alt_iff by assumption. apply validator_alt_free_states_are_projection_states. From 47c8ccedc4b3cd8f0d89a27bf82bfe477b47be75 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:17:48 +0200 Subject: [PATCH 04/18] Library updates --- .gitignore | 1 + theories/VLSM/Lib/ListExtras.v | 27 +++++++++++++++++++++++++++ theories/VLSM/Lib/ListSetExtras.v | 17 +++++++++++++++++ 3 files changed, 45 insertions(+) diff --git a/.gitignore b/.gitignore index 437c2d05..69762a5c 100644 --- a/.gitignore +++ b/.gitignore @@ -26,3 +26,4 @@ deps.map *.dpd *.dot .coq-native/ +.vscode diff --git a/theories/VLSM/Lib/ListExtras.v b/theories/VLSM/Lib/ListExtras.v index 9890e373..b034d934 100644 --- a/theories/VLSM/Lib/ListExtras.v +++ b/theories/VLSM/Lib/ListExtras.v @@ -783,6 +783,33 @@ Proof. + right. specialize (IHl (Forall_tl Hs)). apply IHl. assumption. Qed. +Lemma elem_of_list_annotate + `{EqDecision A} + (P : A -> Prop) + {Pdec : forall a, Decision (P a)} + (l : list A) + (Hs : Forall P l) + (xP : dsig P) + : xP ∈ (list_annotate P l Hs) <-> (` xP) ∈ l. +Proof. + split; [apply elem_of_list_annotate_forget|]. + destruct_dec_sig xP x HPx HeqxP. + subst. + simpl. + intros Hx. + induction l; [inversion Hx|]. + rewrite list_annotate_unroll. + destruct (decide (x = a)) as [Heq | Hneq]. + - subst. + replace (@dexist A P _ a (Forall_hd Hs)) with (dec_exist P a HPx); [left|]. + apply dsig_eq. + reflexivity. + - right. + apply IHl. + inversion Hx; [congruence|]. + assumption. +Qed. + Lemma nth_error_list_annotate {A : Type} (P : A -> Prop) diff --git a/theories/VLSM/Lib/ListSetExtras.v b/theories/VLSM/Lib/ListSetExtras.v index d4286124..a4fb26bb 100644 --- a/theories/VLSM/Lib/ListSetExtras.v +++ b/theories/VLSM/Lib/ListSetExtras.v @@ -116,6 +116,14 @@ Proof. inversion Ha. Qed. +Lemma set_union_nodup_left `{EqDecision A} (l l' : set A) + : NoDup l -> NoDup (set_union l l'). +Proof. + intro Hl. + induction l' as [|x' l' IH]; [trivial|]. + now apply set_add_nodup. +Qed. + Lemma set_union_subseteq_left `{EqDecision A} : forall (s1 s2 : list A), s1 ⊆ (set_union s1 s2). Proof. @@ -130,6 +138,15 @@ Proof. right; assumption. Qed. +Lemma set_union_subseteq_iff `{EqDecision A} : forall (s1 s2 s : list A), + set_union s1 s2 ⊆ s <-> s1 ⊆ s /\ s2 ⊆ s. +Proof. + intros. + unfold subseteq, list_subseteq. + setoid_rewrite set_union_iff. + intuition. +Qed. + Lemma set_union_iterated_nodup `{EqDecision A} (ss : list (list A)) (H : forall s, s ∈ ss -> NoDup s) : From 4673f4ac1ef15e76de430884d61f3f13a8c83ec7 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:24:09 +0200 Subject: [PATCH 05/18] Specialization of existing results --- theories/VLSM/Core/Composition.v | 32 ++ theories/VLSM/Core/Equivocation.v | 288 ++++++++++++----- .../Core/Equivocation/FixedSetEquivocation.v | 2 +- theories/VLSM/Core/EquivocationProjections.v | 7 +- .../LimitedEquivocation/FixedEquivocation.v | 2 +- theories/VLSM/Core/ProjectionTraces.v | 32 +- theories/VLSM/Core/SubProjectionTraces.v | 291 +++++++++++++++++- theories/VLSM/Core/VLSM.v | 112 ++++++- theories/VLSM/Core/VLSMProjections.v | 50 ++- 9 files changed, 694 insertions(+), 122 deletions(-) diff --git a/theories/VLSM/Core/Composition.v b/theories/VLSM/Core/Composition.v index 1be63f00..a3e015dd 100644 --- a/theories/VLSM/Core/Composition.v +++ b/theories/VLSM/Core/Composition.v @@ -515,6 +515,19 @@ and/or the message being valid (e.g., Lemma [Fixed_incl_StrongFixed]). forall (l : composite_label) (som : composite_state * option message), input_valid (composite_vlsm constraint1) l som -> constraint2 l som. +(** +The weakest form [constraint_subsumption] also requires that the input +state and message are valid for the composition under the second constraint. +*) + Definition weak_input_valid_constraint_subsumption + (constraint1 constraint2 : composite_label -> composite_state * option message -> Prop) + := + forall (l : composite_label) (som : composite_state * option message), + input_valid (composite_vlsm constraint1) l som -> + valid_state_prop (composite_vlsm constraint2) som.1 -> + option_valid_message_prop (composite_vlsm constraint2) som.2 -> + constraint2 l som. + Context (constraint1 constraint2 : composite_label -> composite_state * option message -> Prop) (X1 := composite_vlsm constraint1) @@ -531,6 +544,18 @@ Lemma [basic_VLSM_incl] *) (* begin hide *) + Lemma weak_constraint_subsumption_incl + (Hsubsumption : weak_input_valid_constraint_subsumption constraint1 constraint2) + : VLSM_incl X1 X2. + Proof. + apply basic_VLSM_incl; intro; intros. + - assumption. + - apply initial_message_is_valid. assumption. + - split; [apply Hv|]. + apply Hsubsumption; assumption. + - apply H. + Qed. + Lemma constraint_subsumption_input_valid (Hsubsumption : input_valid_constraint_subsumption constraint1 constraint2) (l : label) @@ -587,6 +612,13 @@ Lemma [basic_VLSM_incl] - apply preloaded_constraint_subsumption_input_valid; assumption. Qed. + Lemma weak_constraint_subsumption_weakest + (Hsubsumption : input_valid_constraint_subsumption constraint1 constraint2) + : weak_input_valid_constraint_subsumption constraint1 constraint2. + Proof. + intros l som Hv _ _. apply Hsubsumption. assumption. + Qed. + Lemma preloaded_constraint_subsumption_stronger (Hpre_subsumption : preloaded_constraint_subsumption constraint1 constraint2) : input_valid_constraint_subsumption constraint1 constraint2. diff --git a/theories/VLSM/Core/Equivocation.v b/theories/VLSM/Core/Equivocation.v index b1b20fd0..70c10382 100644 --- a/theories/VLSM/Core/Equivocation.v +++ b/theories/VLSM/Core/Equivocation.v @@ -1239,6 +1239,36 @@ Proof. apply proper_not_sent. Defined. +Lemma preloaded_has_been_sent_stepwise_props + [message : Type] + [vlsm: VLSM message] + (Hhbs: HasBeenSentCapability vlsm) + (seed : message -> Prop) + (X := pre_loaded_vlsm vlsm seed): + has_been_sent_stepwise_props (vlsm := X) (has_been_sent vlsm). +Proof. + destruct (has_been_sent_stepwise_from_trace Hhbs) as [Hinit Hupdate]. + split. + - intros s Hs. apply Hinit. assumption. + - intros l s im s' om Ht msg. + apply (Hupdate l s im s' om). + revert Ht. + apply VLSM_incl_input_valid_transition. + apply basic_VLSM_incl_preloaded; cbv; intuition. +Qed. + +Lemma preloaded_HasBeenSentCapability + [message : Type] + [vlsm: VLSM message] + (Hhbs: HasBeenSentCapability vlsm) + (seed : message -> Prop): + HasBeenSentCapability (pre_loaded_vlsm vlsm seed). +Proof. + eapply HasBeenSentCapability_from_stepwise. + - apply Hhbs. + - apply preloaded_has_been_sent_stepwise_props. +Defined. + Lemma has_been_sent_step_update `{Hhbs: HasBeenSentCapability message vlsm}: forall [l s im s' om], @@ -1292,6 +1322,36 @@ Proof. apply proper_not_received. Defined. +Lemma preloaded_has_been_received_stepwise_props + [message : Type] + [vlsm: VLSM message] + (Hhbr: HasBeenReceivedCapability vlsm) + (seed : message -> Prop) + (X := pre_loaded_vlsm vlsm seed): + has_been_received_stepwise_props (vlsm := X) (has_been_received vlsm). +Proof. + destruct (has_been_received_stepwise_from_trace Hhbr) as [Hinit Hupdate]. + split. + - intros s Hs. apply Hinit. assumption. + - intros l s im s' om Ht msg. + apply (Hupdate l s im s' om). + revert Ht. + apply VLSM_incl_input_valid_transition. + apply basic_VLSM_incl_preloaded; cbv; intuition. +Qed. + +Lemma preloaded_HasBeenReceivedCapability + [message : Type] + [vlsm: VLSM message] + (Hhbr: HasBeenReceivedCapability vlsm) + (seed : message -> Prop): + HasBeenReceivedCapability (pre_loaded_vlsm vlsm seed). +Proof. + eapply HasBeenReceivedCapability_from_stepwise. + - apply Hhbr. + - apply preloaded_has_been_received_stepwise_props. +Defined. + Lemma has_been_received_step_update `{Hhbs: HasBeenReceivedCapability message vlsm}: forall [l s im s' om], @@ -1583,6 +1643,43 @@ Proof. + auto. Qed. +Lemma received_valid + [message] + (X : VLSM message) + {Hhbr: HasBeenReceivedCapability X} + (s : state) + (Hs : valid_state_prop X s) + (m : message) + (Hreceived : has_been_received X s m) : + valid_message_prop X m. +Proof. + induction Hs using valid_state_prop_ind. + - contradict Hreceived. apply (oracle_no_inits (has_been_received_stepwise_from_trace Hhbr));assumption. + - apply input_valid_transition_in in Ht as Hom'. + apply preloaded_weaken_input_valid_transition in Ht. + apply (oracle_step_update (has_been_received_stepwise_from_trace Hhbr) _ _ _ _ _ Ht) in Hreceived. + destruct Hreceived. + + simpl in H. subst om. assumption. + + auto. +Qed. + +Lemma observed_valid + [message] + (X : VLSM message) + {Hhbs: HasBeenSentCapability X} + {Hhbr: HasBeenReceivedCapability X} + (Hhbo: HasBeenObservedCapability X := HasBeenObservedCapability_from_sent_received X) + (s : state) + (Hs : valid_state_prop X s) + (m : message) + (Hobserved : has_been_observed X s m) : + valid_message_prop X m. +Proof. + destruct Hobserved as [Hsent | Hreceived]. + - eapply sent_valid; eassumption. + - eapply received_valid; eassumption. +Qed. + (** *** Equivocation in compositions We now move on to a composite context. Each component of our composition @@ -1811,6 +1908,34 @@ Section Composite. Global Instance free_composite_HasBeenReceivedCapability : HasBeenReceivedCapability Free := composite_HasBeenReceivedCapability (free_constraint IM). + Lemma preloaded_composite_has_been_received_stepwise_props + (constraint : composite_label IM -> composite_state IM * option message -> Prop) + (seed : message -> Prop) + (X := pre_loaded_vlsm (composite_vlsm IM constraint) seed) + : has_been_received_stepwise_props (vlsm := X) composite_has_been_received. + Proof. + unfold has_been_received_stepwise_props. + pose proof (composite_stepwise_props + (fun i => has_been_received_stepwise_from_trace + (has_been_received_capabilities i))) + as [Hinits Hstep]. + split;[exact Hinits|]. + (* <> doesn't work because [composite_message_selector] + pattern matches on the label l, so we instantiate and destruct + to let that simplify *) + intros l;specialize (Hstep l);destruct l. + exact Hstep. + Qed. + + Definition preloaded_composite_HasBeenReceivedCapability + (constraint : composite_label IM -> composite_state IM * option message -> Prop) + (seed : message -> Prop) + (X := pre_loaded_vlsm (composite_vlsm IM constraint) seed) + : HasBeenReceivedCapability X := + HasBeenReceivedCapability_from_stepwise (vlsm := X) + composite_has_been_received_dec + (preloaded_composite_has_been_received_stepwise_props constraint seed). + End composite_has_been_received. @@ -1942,6 +2067,33 @@ Section Composite. Definition no_initial_messages_in_IM_prop : Prop := forall i m, ~vinitial_message_prop (IM i) m. + Lemma composite_no_initial_valid_messages_emitted_by_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), valid_message_prop X m -> + exists v, sender m = Some v /\ + can_emit (pre_loaded_with_all_messages_vlsm (IM (A v))) m. + Proof. + intro m. + rewrite emitted_messages_are_valid_iff. + intros [[i [[mi Hmi] Hom]] | [(s, om) [(i, l) [s' Ht]]]] + ; [elim (no_initial_messages_in_IM i mi); assumption|]. + 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]. + specialize (can_emit_signed i m). + spec can_emit_signed; [eexists _,_,_; eassumption|]. + unfold channel_authenticated_message in can_emit_signed. + destruct (sender m) as [v|] eqn: Hsender; [|inversion can_emit_signed]. + apply Some_inj in can_emit_signed. + exists v. + split; [reflexivity|]. + subst. cbn in Ht. + eexists _, _, _. + eassumption. + Qed. + 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) @@ -1949,22 +2101,13 @@ Section Composite. (X := composite_vlsm IM constraint) : forall (m : message) (Hm : valid_message_prop X m), sender m <> None. Proof. - intros m Hm Hsender. - apply emitted_messages_are_valid_iff in Hm. - destruct Hm as [Hinitial | Hemit]. - - destruct Hinitial as [i [[mi Hmi] Hom]]. - simpl in Hom. subst. - apply (no_initial_messages_in_IM i m). - assumption. - - destruct Hemit as [(s, om) [(i, l) [s' Ht]]]. - specialize (can_emit_signed i m). - unfold channel_authenticated_message in can_emit_signed. - rewrite Hsender in can_emit_signed. simpl in can_emit_signed. - spec can_emit_signed; [|congruence]. - red. eexists _, _, _. - 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. + intros m Hm. + cut + (exists v, sender m = Some v /\ + can_emit (pre_loaded_with_all_messages_vlsm (IM (A v))) m) + ; [intros [v [Hsender _]]; congruence|]. + revert m Hm. + apply composite_no_initial_valid_messages_emitted_by_sender; assumption. Qed. Lemma has_been_sent_iff_by_sender @@ -2121,23 +2264,23 @@ Section Composite. (Hreceived : has_been_received (IM i) (s i) m) : valid_message_prop X m. Proof. - assert (Hcomp : has_been_received Free s m) by (exists i; assumption). - - apply valid_state_has_trace in Hs as H'. - destruct H' as [is [tr Hpr]]. - 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 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_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. + pose (Xhbr := composite_HasBeenReceivedCapability has_been_received_capabilities constraint). + apply (received_valid X s). assumption. exists i;assumption. + Qed. + + Lemma preloaded_messages_received_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 : valid_state_prop X s) + (i : index) + (m : message) + (Hreceived : has_been_received (IM i) (s i) m) + : valid_message_prop X m. + Proof. + pose (Xhbr := preloaded_composite_HasBeenReceivedCapability has_been_received_capabilities constraint seed). + apply (received_valid X s). assumption. exists i;assumption. Qed. Lemma composite_sent_valid @@ -2183,6 +2326,21 @@ Section Composite. ; assumption. Qed. + Lemma preloaded_composite_received_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 : valid_state_prop X s) + (m : message) + (Hreceived : composite_has_been_received has_been_received_capabilities s m) + : valid_message_prop X m. + Proof. + destruct Hreceived as [i Hreceived]. + apply preloaded_messages_received_from_component_of_valid_state_are_valid with s i + ; assumption. + Qed. + Lemma composite_observed_valid (constraint : composite_label IM -> composite_state IM * option message -> Prop) (X := composite_vlsm IM constraint) @@ -2199,6 +2357,27 @@ Section Composite. + apply messages_sent_from_component_of_valid_state_are_valid with s i; assumption. - revert Hs. apply valid_state_project_preloaded. Qed. + + Lemma preloaded_composite_observed_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 : valid_state_prop X s) + (m : message) + (Hobserved : composite_has_been_observed s 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 preloaded_messages_received_from_component_of_valid_state_are_valid with s i; assumption. + + apply preloaded_messages_sent_from_component_of_valid_state_are_valid with s i; assumption. + - eapply valid_state_project_preloaded_to_preloaded. + eapply VLSM_incl_valid_state; [|eassumption]. + apply pre_loaded_vlsm_incl_pre_loaded_with_all_messages. + Qed. + End Composite. Lemma composite_has_been_observed_sent_received_iff @@ -2541,9 +2720,7 @@ Context {index_listing : list index} (finite_index : Listing index_listing) (X_HasBeenReceivedCapability : HasBeenReceivedCapability X - := composite_HasBeenReceivedCapability IM finite_index has_been_received_capabilities constraint - ) - . + := composite_HasBeenReceivedCapability IM finite_index has_been_received_capabilities constraint). Existing Instance X_HasBeenReceivedCapability. @@ -2556,45 +2733,6 @@ any message which [has_been_received] for state <> is valid. 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_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 -> - valid_message_prop X m - ) - : 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_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_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_valid_trace_last_pstate in IHtr as Hplst. - repeat split; try assumption. - 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_valid_state_are_valid s (Hs : valid_state_prop X s) diff --git a/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v b/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v index 6e1932dc..ecefacbc 100644 --- a/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v +++ b/theories/VLSM/Core/Equivocation/FixedSetEquivocation.v @@ -117,7 +117,7 @@ Section strong_fixed_equivocation. Definition sent_by_non_equivocating s m := exists i, i ∉ equivocating /\ has_been_sent (IM i) (s i) m. -Local Instance sent_by_non_equivocating_dec : RelDecision sent_by_non_equivocating. +Global Instance sent_by_non_equivocating_dec : RelDecision sent_by_non_equivocating. Proof. intros s m. apply @Decision_iff with (P := Exists (fun i => has_been_sent (IM i) (s i) m) (filter (fun i => i ∉ equivocating) index_listing)). diff --git a/theories/VLSM/Core/EquivocationProjections.v b/theories/VLSM/Core/EquivocationProjections.v index 7d810ce4..c2b0b2ad 100644 --- a/theories/VLSM/Core/EquivocationProjections.v +++ b/theories/VLSM/Core/EquivocationProjections.v @@ -362,12 +362,7 @@ Proof. specialize (preloaded_component_projection IM i) as Hproj. specialize (VLSM_projection_input_valid_transition Hproj (existT i li) li) as Htransition. - spec Htransition. - { clear. unfold composite_project_label. simpl. - case_decide; [|congruence]. - replace H with (@eq_refl index i) by (apply Eqdep_dec.UIP_dec; assumption). - reflexivity. - } + spec Htransition; [apply (composite_project_label_eq IM)|]. apply Htransition in Hemitted. clear Htransition. remember (s0 i) as s0i. clear s0 Heqs0i. remember (s1 i) as s1i. clear s1 Heqs1i. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v index 185bb3c1..8b5024af 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v @@ -765,7 +765,7 @@ Proof. inversion Hpr. subst. clear Hpr. inversion Hpr_pre_item. subst. clear Hpr_pre_item. constructor. reflexivity. - Unshelve. +Unshelve. - assumption. - assumption. Qed. diff --git a/theories/VLSM/Core/ProjectionTraces.v b/theories/VLSM/Core/ProjectionTraces.v index 80cc1c0d..93f6920c 100644 --- a/theories/VLSM/Core/ProjectionTraces.v +++ b/theories/VLSM/Core/ProjectionTraces.v @@ -264,15 +264,27 @@ Definition composite_project_label | _ => None end. +Lemma composite_project_label_eq lj + : composite_project_label (existT j lj) = Some lj. +Proof. + unfold composite_project_label. + cbn. + case_decide as Heqi; [|contradiction]. + replace Heqi with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. +Qed. + +Definition composite_vlsm_induced_projection : VLSM message := + projection_induced_vlsm X (type (IM j)) + composite_project_label (fun s => s j) + (lift_to_composite_label IM j) (lift_to_composite_state IM j). + (** The [composite_vlsm_constraint_projection] is [VLSM_eq]ual (trace-equivalent) to the [projection_induced_vlsm] by the [composite_project_label] and the projection of the state to the component. *) Lemma composite_vlsm_constrained_projection_is_induced - : VLSM_eq Xj - (projection_induced_vlsm X (type (IM j)) - composite_project_label (fun s => s j) - (lift_to_composite_label IM j) (lift_to_composite_state IM j)). + : VLSM_eq Xj composite_vlsm_induced_projection. Proof. apply VLSM_eq_incl_iff. split. @@ -286,11 +298,7 @@ Proof. + intros l s iom [sX [<- Hv]]. exists (existT j l), sX. intuition. - unfold composite_project_label. - case_decide; [|contradiction]. - cbn in H |- *. - replace H with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). - reflexivity. + apply composite_project_label_eq. + intros l s iom s' oom. cbn. unfold lift_to_composite_state at 1. @@ -331,11 +339,7 @@ Lemma component_label_projection_lift (lift_to_composite_label IM j). Proof. intros lj. - unfold composite_project_label. - case_decide as Hj; [|contradiction]. - cbn in Hj |- *. - replace Hj with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). - reflexivity. + apply composite_project_label_eq. Qed. Lemma component_state_projection_lift diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index ea38b81d..b3ed4b97 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -56,6 +56,20 @@ Proof. intros. subst. apply state_update_eq. Qed. +Lemma sub_IM_state_update_neq + (s : composite_state sub_IM) + (i : index) + (ei : sub_index_prop i) + (si : vstate (IM i)) + (j : index) + (ej : sub_index_prop j) + : i <> j -> state_update sub_IM s (dec_exist _ i ei) si (dec_exist _ j ej) = s (dec_exist _ j ej). +Proof. + intro Hneq. + apply state_update_neq. + setoid_rewrite dsig_eq. + simpl. congruence. +Qed. Definition free_sub_vlsm_composition : VLSM message := free_composite_vlsm sub_IM. @@ -335,6 +349,33 @@ Proof. assumption. Qed. +Lemma induced_sub_projection_transition_is_composite l s om + : vtransition induced_sub_projection l (s, om) = composite_transition sub_IM l (s, om). +Proof. + destruct l as (sub_i, li). + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + cbn. + unfold sub_IM, lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi := Hi). + cbn. + destruct (vtransition _ _ _) as (si', om'). + f_equal. + extensionality sub_k. + destruct_dec_sig sub_k k Hk Heqsub_k. + subst. + unfold composite_state_sub_projection. + simpl. + destruct (decide (i = k)). + + subst. + rewrite state_update_eq. + symmetry. + apply sub_IM_state_update_eq. + + setoid_rewrite sub_IM_state_update_neq; [|congruence]. + rewrite state_update_neq by congruence. + apply lift_sub_state_to_eq. +Qed. + End sec_induced_sub_projection. Section induced_sub_projection_subsumption. @@ -456,16 +497,23 @@ Proof. apply Hfinite. Qed. +Program Definition sub_index_list_annotate : list sub_index := + list_annotate _ sub_index_list _. +Next Obligation. + apply Forall_forall. + intuition. +Qed. + Global Instance stdpp_finite_sub_index - {Hfinite : finite.Finite index} : finite.Finite sub_index. Proof. - exists (select_sub_indices (enum index)). - - apply select_sub_indices_nodup. apply Hfinite. - - intro sub_x. destruct_dec_sig sub_x x Hx Heqsub_x. + exists (remove_dups sub_index_list_annotate). + - apply NoDup_remove_dups. + - intro sub_x. apply elem_of_remove_dups. + apply elem_of_list_annotate. + destruct_dec_sig sub_x x Hx Heqsub_x. subst. - apply in_select_sub_indices. - apply Hfinite. + assumption. Qed. Local Instance Sub_Free_HasBeenSentCapability @@ -1173,7 +1221,7 @@ Lemma induced_sub_projection_friendliness Proof. eapply (basic_projection_induces_friendliness (composite_vlsm IM constraint)). assumption. - Unshelve. +Unshelve. - apply induced_sub_projection_transition_consistency_None. - apply composite_label_sub_projection_option_lift. - apply composite_state_sub_projection_lift. @@ -1752,6 +1800,235 @@ Qed. End sub_composition_all. +Section sub_composition_element. +(** ** A component can be lifted to a free subcomposition *) + +Context + {message : Type} + {index : Type} + {IndEqDec : EqDecision index} + (IM : index -> VLSM message) + (indices : set index) + (j : index) + (Hj : j ∈ indices) + . + +Definition sub_element_label (l : vlabel (IM j)) + : composite_label (sub_IM IM indices) := + existT (dexist j Hj) l. + +Definition sub_element_state (s : vstate (IM j)) sub_i + : vstate (sub_IM IM indices sub_i) := + match (decide (` sub_i = j)) with + | left e => + eq_rect_r (λ j : index, vstate (IM j)) s e + | right _ => ` (vs0 (IM (` sub_i))) + end. + +Lemma sub_element_state_eq s H_j + : sub_element_state s (dexist j H_j) = s. +Proof. + unfold sub_element_state. + simpl. + case_decide as Heq_j; [|contradiction]. + replace Heq_j with (@eq_refl index j) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. +Qed. + +Lemma sub_element_state_neq s i Hi + : i <> j -> sub_element_state s (dexist i Hi) = ` (vs0 (IM i)). +Proof. + intros Hij. + unfold sub_element_state. + simpl. + case_decide; [contradiction|reflexivity]. +Qed. + +Lemma preloaded_sub_element_full_projection + (P Q : message -> Prop) + (PimpliesQ : forall m, P m -> Q m) + (PrePXj := pre_loaded_vlsm (IM j) P) + (PreQSubFree := pre_loaded_vlsm (free_composite_vlsm (sub_IM IM indices)) Q) + : VLSM_full_projection PrePXj PreQSubFree sub_element_label sub_element_state. +Proof. + apply basic_VLSM_full_projection_preloaded_with; [assumption|..]. + - intros l s om Hv. + split; [|exact I]. + cbn. + rewrite sub_element_state_eq with (H_j := Hj). + assumption. + - intros l s om s' om'. + cbn. + rewrite sub_element_state_eq with (H_j := Hj). + intro Ht. + replace (vtransition _ _ _) with (s', om'). + f_equal. + extensionality sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + destruct (decide (i = j)). + + subst. + rewrite sub_IM_state_update_eq, sub_element_state_eq. + reflexivity. + + rewrite sub_IM_state_update_neq by congruence. + rewrite !sub_element_state_neq by congruence. + reflexivity. + - intros sj Hsj sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + destruct (decide (i = j)). + + subst. rewrite sub_element_state_eq. assumption. + + rewrite sub_element_state_neq by congruence. + destruct (vs0 (IM i)). + assumption. + - intros m Hm. + exists (dexist j Hj), (exist _ m Hm). + reflexivity. +Qed. + +Lemma valid_preloaded_lifts_can_be_emitted + (P Q : message -> Prop) + (HPvalid : forall dm, P dm -> valid_message_prop (pre_loaded_vlsm (free_composite_vlsm (sub_IM IM indices)) Q) dm) + : forall m, can_emit (pre_loaded_vlsm (IM j) P) m -> + can_emit (pre_loaded_vlsm (free_composite_vlsm (sub_IM IM indices)) Q) m. +Proof. + intros m Hm. + eapply VLSM_incl_can_emit. + { + apply pre_loaded_vlsm_incl_relaxed with (P0 := fun m => Q m \/ P m). + intuition. + } + eapply VLSM_full_projection_can_emit; [|eassumption]. + apply preloaded_sub_element_full_projection. + intuition. +Qed. + +Definition sub_label_element_project + (l : composite_label (sub_IM IM indices)) + : option (vlabel (IM j)) := + match decide (j = ` (projT1 l)) with + | left e => Some (eq_rect_r (fun j => vlabel (IM j)) (projT2 l) e) + | in_right => None + end. + +Definition sub_state_element_project + (s : composite_state (sub_IM IM indices)) + : vstate (IM j) := s (dexist j Hj). + +Lemma sub_transition_element_project_None + : forall lX, sub_label_element_project lX = None -> + forall s om s' om', composite_transition (sub_IM IM indices) lX (s, om) = (s', om') -> + sub_state_element_project s' = sub_state_element_project s. +Proof. + intros (sub_i,li) HlX s om s' om' HtX. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + unfold sub_label_element_project in HlX. + cbn in HlX, HtX. + case_decide as Hij; [congruence|]. + clear HlX. + destruct (vtransition _ _ _) as (si', _om'). + inversion HtX. subst. clear HtX. + unfold sub_state_element_project. + apply sub_IM_state_update_neq. + congruence. +Qed. + +Lemma sub_element_label_project + : forall lY, sub_label_element_project (sub_element_label lY) = Some lY. +Proof. + intros lY. + unfold sub_element_label, sub_label_element_project. + simpl. + case_decide as Heqj; [|contradiction]. + replace Heqj with (eq_refl (A := index) (x := j)); [reflexivity|]. + apply Eqdep_dec.UIP_dec. + assumption. +Qed. + +Lemma sub_element_state_project + : forall sY, sub_state_element_project (sub_element_state sY) = sY. +Proof. + intros sY. + unfold sub_element_state, sub_state_element_project. + simpl. + case_decide as Heqj; [|contradiction]. + replace Heqj with (eq_refl (A := index) (x := j)); [reflexivity|]. + apply Eqdep_dec.UIP_dec. + assumption. +Qed. + +Lemma sub_transition_element_project_Some + : forall lX1 lX2 lY, sub_label_element_project lX1 = Some lY -> sub_label_element_project lX2 = Some lY -> + forall sX1 sX2, sub_state_element_project sX1 = sub_state_element_project sX2 -> + forall iom sX1' oom1, composite_transition (sub_IM IM indices) lX1 (sX1, iom) = (sX1', oom1) -> + forall sX2' oom2, composite_transition (sub_IM IM indices) lX2 (sX2, iom) = (sX2', oom2) -> + sub_state_element_project sX1' = sub_state_element_project sX2' /\ oom1 = oom2. +Proof. + intros (sub_j1, lj1) (sub_j2, lj2) lj. + destruct_dec_sig sub_j1 j1 Hj1 Heqsub_j1. + destruct_dec_sig sub_j2 j2 Hj2 Heqsub_j2. + subst. + unfold sub_label_element_project. + cbn. + case_decide as Heqj; intro HlX; inversion HlX as [Heqlj]. + clear HlX. + subst j1. + cbv in Heqlj. + subst lj1. + case_decide as Heqj; intro HlX; inversion HlX as [Heqlj]. + clear HlX. + subst j2. + cbv in Heqlj. + subst lj2. + unfold sub_state_element_project. + intros sX1 sX2 Hsjeq iom. + replace (sX1 (dec_exist (sub_index_prop indices) j Hj1)) with (sX1 (dexist j Hj)) + by apply sub_IM_state_pi. + replace (sX2 (dec_exist (sub_index_prop indices) j Hj2)) with (sX2 (dexist j Hj)) + by apply sub_IM_state_pi. + rewrite <- Hsjeq. + unfold sub_IM. + cbn. + destruct (vtransition _ _ _) as (si', om'). + intros sX' oom HtX. + inversion HtX. subst oom sX'. clear HtX. + intros sX' oom HtX. + inversion HtX. subst oom sX'. clear HtX. + split; [|reflexivity]. + setoid_rewrite sub_IM_state_update_eq. + reflexivity. +Qed. + +Definition induced_sub_element_projection constraint : VLSM message := + projection_induced_vlsm + (induced_sub_projection IM indices constraint) (type (IM j)) + sub_label_element_project sub_state_element_project + sub_element_label sub_element_state. + +Lemma induced_sub_element_projection_is_projection constraint + : VLSM_projection + (induced_sub_projection IM indices constraint) + (induced_sub_element_projection constraint) + sub_label_element_project sub_state_element_project. +Proof. + apply projection_induced_vlsm_is_projection. + - intros lX HlX s om s' om' [_ Ht]. + apply sub_transition_element_project_None with lX om om'; [assumption|]. + setoid_rewrite <- induced_sub_projection_transition_is_composite. + eassumption. + - apply basic_weak_projection_transition_consistency_Some. + + intro. apply sub_element_label_project. + + intro. apply sub_element_state_project. + + intro. + setoid_rewrite induced_sub_projection_transition_is_composite. + apply sub_transition_element_project_Some. +Unshelve. + exact constraint. +Qed. + +End sub_composition_element. + Section sub_composition_preloaded_lift. Context diff --git a/theories/VLSM/Core/VLSM.v b/theories/VLSM/Core/VLSM.v index 288ab3a2..ae5b0fa7 100644 --- a/theories/VLSM/Core/VLSM.v +++ b/theories/VLSM/Core/VLSM.v @@ -1363,6 +1363,24 @@ traces. (* begin hide *) + Lemma can_produce_from_valid_trace si tr + (Htr: finite_valid_trace_from si tr) + : forall item, item ∈ tr -> + option_can_produce (destination item) (output item). + Proof. + intros item Hitem. + cut (exists l1 l2, tr = l1 ++ item :: l2). + { + intros [l1 [l2 Heq]]. + eexists _,_. + eapply input_valid_transition_to with (tr := tr); [eassumption|]. + simpl. + eassumption. + } + apply elem_of_list_split. + assumption. + Qed. + Lemma can_emit_from_valid_trace (si : state) (m : message) @@ -1371,17 +1389,14 @@ traces. (Hm : trace_has_message (field_selector output) m tr) : can_emit m. Proof. - apply Exists_exists in Hm. - destruct Hm as [x [Hin Houtput]]. - apply elem_of_list_split in Hin. - destruct Hin as [l1 [l2 Hconcat]]. - unfold can_emit. - destruct Htr as [Htr _]. - specialize (input_valid_transition_to _ _ _ _ _ Htr Hconcat). - intros Ht. - simpl in Houtput, Ht. - rewrite Houtput in Ht. - do 3 eexists;exact Ht. + apply can_emit_iff. + revert Hm. + setoid_rewrite Exists_exists. + intros [item [Hitem Houtput]]. + exists (destination item). + unfold can_produce. + replace (Some m) with (output item) by assumption. + eapply can_produce_from_valid_trace; [apply Htr|assumption]. Qed. (* End Hide *) @@ -2173,6 +2188,42 @@ This relation is often used in stating safety and liveness properties.*) eexists. exact Ht. Qed. + Lemma elem_of_trace_in_futures_left is s tr + (Htr : finite_valid_trace_from_to is s tr) + : forall item, item ∈ tr -> in_futures (destination item) s. + Proof. + intros item. + rewrite elem_of_list_In. + intros Hitem. + apply in_split in Hitem as [pre [suf Heqtr]]. + exists suf. + replace (destination item) with (finite_trace_last is (pre ++ [item])) + by apply finite_trace_last_is_last. + eapply finite_valid_trace_from_to_app_split. + rewrite <- app_assoc. + simpl. + rewrite <- Heqtr. + assumption. + Qed. + + Lemma elem_of_trace_in_futures_right is s tr + (Htr : finite_valid_trace_from_to is s tr) + : forall item, item ∈ tr -> in_futures is (destination item). + Proof. + intros item. + rewrite elem_of_list_In. + intros Hitem. + apply in_split in Hitem as [pre [suf Heqtr]]. + exists (pre ++ [item]). + replace (destination item) with (finite_trace_last is (pre ++ [item])) + by apply finite_trace_last_is_last. + eapply finite_valid_trace_from_to_app_split. + rewrite <- app_assoc. + simpl. + rewrite <- Heqtr. + eassumption. + Qed. + Lemma in_futures_witness (first second : state) (Hfutures : in_futures first second) @@ -2763,6 +2814,45 @@ Byzantine fault tolerance analysis. revert Hx. apply preloaded_weaken_input_valid_transition. Qed. + Lemma pre_traces_with_valid_inputs_are_valid is s tr + (Htr : finite_valid_trace_init_to pre_loaded_with_all_messages_vlsm is s tr) + (Hobs : forall m, + trace_has_message (field_selector input) m tr -> + valid_message_prop X m + ) + : 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_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_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]] Ht]. + apply finite_valid_trace_last_pstate in IHtr as Hplst. + repeat split; try assumption. + 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. + End pre_loaded_with_all_messages_vlsm. Lemma non_empty_valid_trace_from_can_produce diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index 177692ab..d6deee64 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -678,13 +678,13 @@ Qed. Lemma VLSM_weak_projection_input_valid : forall lX lY, label_project lX = Some lY -> - forall sX im, - input_valid X lX (sX, im) -> input_valid Y lY (state_project sX, im). + forall s im, input_valid X lX (s, im) -> input_valid Y lY (state_project s, im). Proof. - intros lX lY Hpr sX im HvX. - destruct (vtransition X lX (sX, im)) eqn:HtX. - eapply proj1, VLSM_weak_projection_input_valid_transition, input_valid_can_transition - ; [eassumption|assumption|exact HtX]. + intros lX lY HlX_pr s im Hv. + destruct (vtransition X lX (s, im)) as (s', om') eqn:Ht. + eapply input_valid_transition_valid. + eapply VLSM_weak_projection_input_valid_transition; [eassumption|]. + split; eassumption. Qed. Lemma VLSM_weak_projection_finite_valid_trace_from_to @@ -1046,6 +1046,17 @@ Proof. reflexivity. Qed. +Lemma pre_VLSM_full_projection_finite_trace_last_output + : forall trX, + finite_trace_last_output trX = finite_trace_last_output (pre_VLSM_full_projection_finite_trace_project trX). +Proof. + intros. + destruct_list_last trX trX' lst HtrX + ; [reflexivity|]. + setoid_rewrite map_app. simpl. rewrite !finite_trace_last_output_is_last. + reflexivity. +Qed. + End pre_definitions. Section basic_definitions. @@ -2707,12 +2718,37 @@ Context (X : VLSM message) . +Lemma pre_loaded_vlsm_incl_relaxed + (P Q : message -> Prop) + (PimpliesQorValid : forall m : message, P m -> Q m \/ valid_message_prop (pre_loaded_vlsm X Q) m) + : VLSM_incl (pre_loaded_vlsm X P) (pre_loaded_vlsm X Q). +Proof. + apply basic_VLSM_incl. + - cbv; intuition. + - intros _ _ m _ _ [Him | Hp]. + + apply initial_message_is_valid. left. assumption. + + apply PimpliesQorValid in Hp as [Hq | Hvalid]; [|assumption]. + apply initial_message_is_valid. right. assumption. + - cbv; intuition. + - cbv; intuition. +Qed. + Lemma pre_loaded_vlsm_incl (P Q : message -> Prop) (PimpliesQ : forall m : message, P m -> Q m) : VLSM_incl (pre_loaded_vlsm X P) (pre_loaded_vlsm X Q). Proof. - apply basic_VLSM_incl_preloaded_with; cbv; intuition. + apply pre_loaded_vlsm_incl_relaxed. intuition. +Qed. + +Lemma pre_loaded_vlsm_with_valid_eq + (P Q : message -> Prop) + (QimpliesValid : forall m, Q m -> valid_message_prop (pre_loaded_vlsm X P) m) + : VLSM_eq (pre_loaded_vlsm X (fun m => P m \/ Q m)) (pre_loaded_vlsm X P). +Proof. + apply VLSM_eq_incl_iff; split. + - apply pre_loaded_vlsm_incl_relaxed. intuition. + - cbv; apply pre_loaded_vlsm_incl. intuition. Qed. Lemma pre_loaded_vlsm_idem_l From eb075128b9c4a9b2f0b96a1b9d77e0b29cc54f5c Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:37:01 +0200 Subject: [PATCH 06/18] Revamped Message Dependencies --- theories/VLSM/Core/MessageDependencies.v | 461 ++++++++++++++++++++++- 1 file changed, 452 insertions(+), 9 deletions(-) diff --git a/theories/VLSM/Core/MessageDependencies.v b/theories/VLSM/Core/MessageDependencies.v index 68917ea4..faf19fc0 100644 --- a/theories/VLSM/Core/MessageDependencies.v +++ b/theories/VLSM/Core/MessageDependencies.v @@ -1,7 +1,7 @@ From stdpp Require Import prelude. -From Coq Require Import FinFun. -From VLSM Require Import Lib.Preamble Lib.StdppListSet Lib.Measurable. -From VLSM Require Import Core.VLSM Core.Composition Core.Equivocation. +From Coq Require Import FinFun Relations.Relation_Operators. +From VLSM.Lib Require Import Preamble ListExtras FinFunExtras StdppListSet Measurable. +From VLSM.Core Require Import VLSM VLSMProjections Composition ProjectionTraces SubProjectionTraces Equivocation. (** * VLSM Message Dependencies @@ -10,15 +10,19 @@ Assumes that each message has an associated set of [message_dependencies]. *) -Section message_dependencies. +Section sec_message_dependencies. Context {message : Type} (message_dependencies : message -> set message) (X : VLSM message) - {Hbo : HasBeenObservedCapability X} + {Hbs : HasBeenSentCapability X} + {Hbr : HasBeenReceivedCapability X} + (Hbo : HasBeenObservedCapability X := HasBeenObservedCapability_from_sent_received X) . +Existing Instance Hbo. + (** [MessageDependencies] characterize a <> function through two properties: @@ -51,9 +55,8 @@ observed all of <>'s dependencies. Definition message_dependencies_full_node_condition (s : vstate X) (m : message) - : Prop - := forall dm, dm ∈ message_dependencies m -> - has_been_observed X s dm. + : Prop := + forall dm, dm ∈ message_dependencies m -> has_been_observed X s dm. (** A VLSM has the [message_dependencies_full_node_condition_prop] if the validity of receiving a message in a state implies the @@ -63,4 +66,444 @@ Definition message_dependencies_full_node_condition_prop : Prop := forall l s m, vvalid X l (s, Some m) -> message_dependencies_full_node_condition s m. -End message_dependencies. +(** Membership to the message_dependencies of a message induces a dependency +relation. +*) +Definition msg_dep_rel : relation message := + fun m1 m2 => m1 ∈ message_dependencies m2. + +(** The transitive closure of the [msg_dep_rel]ation is a happens-before +relation. +*) +Definition msg_dep_happens_before : relation message := flip (clos_trans _ (flip msg_dep_rel)). + +(** Unrolling one the [msg_dep_happens_before] relation one step. *) +Lemma msg_dep_happens_before_iff_one x z + : msg_dep_happens_before x z <-> + msg_dep_rel x z \/ exists y, msg_dep_happens_before x y /\ msg_dep_rel y z. +Proof. + split. + - intros Hhb. + apply Operators_Properties.clos_trans_t1n in Hhb. + inversion Hhb; subst. + + left. assumption. + + right. + exists y. + split; [|assumption]. + apply Operators_Properties.clos_t1n_trans. + assumption. + - intros Hhb. + apply Operators_Properties.clos_t1n_trans. + destruct Hhb as [Hone | [y [Hhb Hone]]]. + + apply t1n_step. assumption. + + apply t1n_trans with y; [assumption|]. + apply Operators_Properties.clos_trans_t1n. + assumption. +Qed. + +Global Instance msg_dep_happens_before_transitive : Transitive msg_dep_happens_before. +Proof. + apply flip_Transitive. + intros m1 m2 m3. + apply t_trans. +Qed. + +(** If the [msg_dep_rel]ation reflects a predicate <

    > and the induced +[msg_dep_happens_before] is [well_founded], then if <

    > holds for a message, +it will hold for all its dependencies. *) +Lemma msg_dep_happens_before_reflect + (P : message -> Prop) + (Hreflects : forall dm m, msg_dep_rel dm m -> P m -> P dm) + : forall dm m, msg_dep_happens_before dm m -> P m -> P dm. +Proof. + intros dm m Hdm. + apply Operators_Properties.clos_trans_t1n in Hdm. + induction Hdm; [apply Hreflects; assumption|]. + intro Hpx. + apply IHHdm. + revert Hpx. + apply Hreflects. + assumption. +Qed. + +(** In the absence of initial messages, and if [msg_dep_rel] reflects the +pre-loaded message property, then it also reflects the [valid_message_prop]erty. +*) +Lemma msg_dep_reflects_validity + (HMsgDep : MessageDependencies) + (no_initial_messages_in_X : forall m, ~ vinitial_message_prop X m) + (P : message -> Prop) + (Hreflects : forall dm m, msg_dep_rel dm m -> P m -> P dm) + (Hbr_pre := preloaded_HasBeenReceivedCapability Hbr) + (Hbs_pre := preloaded_HasBeenSentCapability Hbs) + (Hbo_pre := HasBeenObservedCapability_from_sent_received (pre_loaded_vlsm X P)) + : forall dm m, msg_dep_rel dm m -> + valid_message_prop (pre_loaded_vlsm X P) m -> + valid_message_prop (pre_loaded_vlsm X P) dm. +Proof. + intros dm m Hdm. + rewrite emitted_messages_are_valid_iff. + rewrite can_emit_iff. + intros [Hinit | [s Hproduce]]. + - rewrite emitted_messages_are_valid_iff. + left. + right. + apply Hreflects with m; [assumption|]. + destruct Hinit as [Hinit | Hp]; [|assumption]. + elim (no_initial_messages_in_X m). + assumption. + - apply (observed_valid (pre_loaded_vlsm X P)) with (s := s) + ; [exists (Some m); apply can_produce_valid; eassumption|]. + cut (has_been_observed X s dm). + { + intros [Hsent | Hreceived]; [left|right]. + - clear -Hsent. destruct Hbs. assumption. + - clear -Hreceived. destruct Hbr. assumption. + } + apply message_dependencies_are_necessary with m; [|assumption]. + revert Hproduce. + apply VLSM_incl_can_produce. + apply pre_loaded_vlsm_incl_pre_loaded_with_all_messages. +Qed. + +(** Under [MessageDependencies] assumptions, if a message [has_been_sent] +in a state <>, then any of its direct dependencies [has_been_observed]. +*) +Lemma msg_dep_has_been_sent + (HMsgDep : MessageDependencies) + s + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm X) s) + m + (Hsent : has_been_sent X s m) + : forall dm, msg_dep_rel dm m -> has_been_observed X s dm. +Proof. + intros dm Hdm. + cut (exists item, can_produce (pre_loaded_with_all_messages_vlsm X) (destination item) m /\ in_futures (pre_loaded_with_all_messages_vlsm X) (destination item) s). + { + intros [s_dm [Hproduce Hfutures]]. + eapply in_futures_preserving_oracle_from_stepwise; [|eassumption|]. + - apply has_been_observed_from_sent_received_stepwise_props. + - apply message_dependencies_are_necessary with m; assumption. + } + apply proper_sent in Hsent; [|assumption]. + apply valid_state_has_trace in Hs as [is [tr Htr]]. + specialize (Hsent _ _ Htr). + apply Exists_exists in Hsent as [item [Hitem Houtput]]. + exists item. + split; cycle 1. + - eapply elem_of_trace_in_futures_left; [|eassumption]. + apply Htr. + - unfold can_produce. + replace (Some m) with (output item) by assumption. + eapply can_produce_from_valid_trace; [|eassumption]. + eapply valid_trace_forget_last. + apply Htr. +Qed. + +(** If the [valid]ity predicate has the [message_dependencies_full_node_condition_prop]erty, +then if a message [has_been_received] in a state <>, any of its direct +dependencies [has_been_observed]. +*) +Lemma full_node_has_been_received + (Hfull : message_dependencies_full_node_condition_prop) + s + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm X) s) + m + (Hreceived : has_been_received X s m) + : forall dm, msg_dep_rel dm m -> has_been_observed X s dm. +Proof. + intros dm Hdm. + apply proper_received in Hreceived; [|assumption]. + apply valid_state_has_trace in Hs as [is [tr Htr]]. + specialize (Hreceived _ _ Htr). + apply Exists_exists in Hreceived as [item [Hitem Hinput]]. + apply elem_of_list_split in Hitem as [pre [suf Heqtr]]. + eapply in_futures_preserving_oracle_from_stepwise with (s1 := finite_trace_last is pre) + ; [apply has_been_observed_from_sent_received_stepwise_props|..]. + - exists (item :: suf). + eapply finite_valid_trace_from_to_app_split. + rewrite <- Heqtr. + apply Htr. + - eapply Hfull; [|eassumption]. + replace (Some m) with (input item) by assumption. + clear Hinput. + eapply (input_valid_transition_is_valid (pre_loaded_with_all_messages_vlsm X)). + eapply input_valid_transition_to; [|simpl; eassumption]. + eapply valid_trace_forget_last. + apply Htr. +Qed. + +(** By combining Lemmas [msg_dep_has_been_sent] and [full_node_has_been_received], +[msg_dep_rel] reflects the [has_been_observed] predicate. +*) +Lemma msg_dep_full_node_reflects_has_been_observed + (HMsgDep : MessageDependencies) + (Hfull : message_dependencies_full_node_condition_prop) + s + (Hs : valid_state_prop (pre_loaded_with_all_messages_vlsm X) s) + : forall dm m, msg_dep_rel dm m -> + has_been_observed X s m -> + has_been_observed X s dm. +Proof. + intros dm m Hdm [Hsent|Hreceived]. + - eapply msg_dep_has_been_sent; eassumption. + - eapply full_node_has_been_received; eassumption. +Qed. + +End sec_message_dependencies. + +Section sec_composite_message_dependencies. + +Context + {message : Type} + (message_dependencies : message -> set message) + `{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)) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (Free_Hbs := free_composite_HasBeenSentCapability IM (listing_from_finite index) Hbs) + (Free_Hbr := free_composite_HasBeenReceivedCapability IM (listing_from_finite index) Hbr) + . + +Existing Instance Free_Hbs. +Existing Instance Free_Hbr. + +(** If all of the components satisfy the [MessageDependencies] assumptions, +then their free composition will also do so. +*) +Lemma composite_message_dependencies + : MessageDependencies message_dependencies (free_composite_vlsm IM). +Proof. + split. + - intros m s Hproduce dm Hdm. + destruct Hproduce as [(is, iom) [(i, li) Ht]]. + cut (@has_been_observed _ (IM i) (Hbo i) (s i) dm) + ; [intros [Hsent | Hreceived]; [left | right]; exists i; assumption|]. + eapply message_dependencies_are_necessary; [apply HMsgDep| |eassumption]. + exists (is i, iom), li. + revert Ht. + apply + (VLSM_projection_input_valid_transition (preloaded_component_projection IM _)) + with (lY := li). + unfold composite_project_label. + cbn. + case_decide as Heqi; [|contradiction]. + replace Heqi with (@eq_refl index i) by (apply Eqdep_dec.UIP_dec; assumption). + reflexivity. + - intros m Hemit. + apply can_emit_composite_project in Hemit as [j Hemitj]. + eapply message_dependencies_are_sufficient in Hemitj; [| apply HMsgDep]. + revert Hemitj. + eapply VLSM_full_projection_can_emit. + apply lift_to_composite_generalized_preloaded_vlsm_full_projection. + intuition. +Qed. + +Lemma msg_dep_reflects_free_validity + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (X := free_composite_vlsm IM) + : forall dm m, msg_dep_rel message_dependencies dm m -> + valid_message_prop X m -> valid_message_prop X dm. +Proof. + intros dm m Hdm. + rewrite !emitted_messages_are_valid_iff. + intros [[i [[im Him] Heqm]] | Hemit] + ; [elim (no_initial_messages_in_IM i im); assumption|]. + right. + pose proof (vlsm_is_pre_loaded_with_False X) as XeqXFalse. + apply (VLSM_eq_can_emit XeqXFalse) in Hemit. + apply (VLSM_eq_can_emit XeqXFalse). + cut (valid_message_prop (pre_loaded_vlsm X (fun _ => False)) dm). + { + clear -no_initial_messages_in_IM. + rewrite emitted_messages_are_valid_iff. + intros [[[i [[im Him] Heqm]] | Hpreloaded] | Hemit] + ; [elim (no_initial_messages_in_IM i im); assumption|contradiction| assumption]. + } + eapply msg_dep_reflects_validity. + - apply composite_message_dependencies. + - clear -no_initial_messages_in_IM. + intros m [i [[im Him] Heqm]]. + elim (no_initial_messages_in_IM i im). + assumption. + - intuition. + - eassumption. + - apply emitted_messages_are_valid_iff. + right. + assumption. +Qed. + +Lemma msg_dep_reflects_happens_before_free_validity + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (X := free_composite_vlsm IM) + : forall dm m, msg_dep_happens_before message_dependencies dm m -> + valid_message_prop X m -> valid_message_prop X dm. +Proof. + apply msg_dep_happens_before_reflect. + apply msg_dep_reflects_free_validity. + assumption. +Qed. + +Lemma msg_dep_happens_before_composite_no_initial_valid_messages_emitted_by_sender + {validator : Type} + (sender : message -> option validator) + (A : validator -> index) + (Hchannel : channel_authentication_prop IM A sender) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (X := free_composite_vlsm IM) + : forall m, valid_message_prop X m -> + forall dm, msg_dep_happens_before message_dependencies dm m -> + exists v, sender dm = Some v /\ + can_emit (pre_loaded_with_all_messages_vlsm (IM (A v))) dm. +Proof. + intros m Hm dm Hdm. + cut (valid_message_prop X dm). + { + clear Hdm. + revert dm. + apply composite_no_initial_valid_messages_emitted_by_sender; assumption. + } + revert dm m Hdm Hm. + apply msg_dep_reflects_happens_before_free_validity. + assumption. +Qed. + +End sec_composite_message_dependencies. + +Section sec_sub_composite_message_dependencies. + +Context + {message : Type} + (message_dependencies : message -> set message) + `{EqDecision 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)) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (indices : set index) + (SubFree_Hbs := free_composite_HasBeenSentCapability (sub_IM IM indices) (listing_from_finite (sub_index indices)) (sub_has_been_sent_capabilities IM indices Hbs)) + (SubFree_Hbr := free_composite_HasBeenReceivedCapability (sub_IM IM indices) (listing_from_finite (sub_index indices)) (sub_has_been_received_capabilities IM indices Hbr)) + . + +Existing Instance SubFree_Hbs. +Existing Instance SubFree_Hbr. + +Lemma sub_composite_message_dependencies + : MessageDependencies message_dependencies (free_composite_vlsm (sub_IM IM indices)). +Proof. + apply composite_message_dependencies. + intro sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + apply HMsgDep. +Qed. + +Lemma msg_dep_reflects_sub_free_validity + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (P : message -> Prop) + (Hreflects : forall dm m, msg_dep_rel message_dependencies dm m -> P m -> P dm) + (X := free_composite_vlsm (sub_IM IM indices)) + : forall dm m, msg_dep_rel message_dependencies dm m -> + valid_message_prop (pre_loaded_vlsm X P) m -> + valid_message_prop (pre_loaded_vlsm X P) dm. +Proof. + eapply msg_dep_reflects_validity; [..|assumption]. + - apply sub_composite_message_dependencies. + - intros m [sub_i [[im Him] Heqm]]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + elim (no_initial_messages_in_IM i im). + assumption. +Qed. + +End sec_sub_composite_message_dependencies. + +Section sec_full_message_dependencies. + +Context + {message : Type} + . + +Class FullMessageDependencies + (message_dependencies : message -> set message) + (full_message_dependencies : message -> set message) + := + { full_message_dependencies_happens_before + : forall dm m, dm ∈ full_message_dependencies m <-> msg_dep_happens_before message_dependencies dm m + ; full_message_dependencies_irreflexive + : forall m, m ∉ full_message_dependencies m + ; full_message_dependencies_nodups + : forall m, NoDup (full_message_dependencies m) + }. + +End sec_full_message_dependencies. + +Section full_message_dependencies_happens_before. + +Context + {message : Type} + (message_dependencies : message -> set message) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + . + +Global Instance msg_dep_happens_before_irrefl : Irreflexive (msg_dep_happens_before message_dependencies). +Proof. + intros m Hm. + elim (full_message_dependencies_irreflexive m). + eapply full_message_dependencies_happens_before. + assumption. +Qed. + +Global Instance msg_dep_happens_before_strict : StrictOrder (msg_dep_happens_before message_dependencies) := {}. + +Lemma msg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies). +Proof. + cut (forall n m, length (full_message_dependencies m) < n -> Acc (msg_dep_happens_before message_dependencies) m). + { + intros Hn m. + apply (Hn (S (length (full_message_dependencies m)))). + lia. + } + induction n as [|n IHn]; [lia|]. + intros m Hm. + constructor. + intros dm Hdm. + apply IHn. + unfold lt. + transitivity (length (full_message_dependencies m)); [|lia]. + replace (S _) with (length (dm :: full_message_dependencies dm)) + by reflexivity. + apply NoDup_subseteq_length. + - constructor; [apply full_message_dependencies_irreflexive| apply full_message_dependencies_nodups]. + - intros m' Hm'. + inversion Hm'; subst; [apply full_message_dependencies_happens_before; assumption|]. + revert H1. + setoid_rewrite full_message_dependencies_happens_before. + intro Hm'dm. + transitivity dm; assumption. +Qed. + +Lemma FullMessageDependencies_ind + (P : message -> Prop) + m + (IHm : forall dm, dm ∈ full_message_dependencies m -> (forall dm0, dm0 ∈ full_message_dependencies dm -> P dm0) -> P dm) + : forall dm, dm ∈ full_message_dependencies m -> P dm. +Proof. + induction m using (well_founded_ind msg_dep_happens_before_wf). + intros dm Hdm. + apply IHm; [assumption|]. + apply H; [apply full_message_dependencies_happens_before; assumption|]. + intros dm0 Hdm0. + apply IHm. + apply full_message_dependencies_happens_before in Hdm. + apply full_message_dependencies_happens_before in Hdm0. + apply full_message_dependencies_happens_before. + change _ with (msg_dep_happens_before message_dependencies dm0 m). + transitivity dm; assumption. +Qed. + +End full_message_dependencies_happens_before. From 898e801f665c78f4004dc886ccedd6fe57df25c3 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:38:57 +0200 Subject: [PATCH 07/18] FixedSet Equivocation based on message dependencies --- _CoqProject | 1 + .../Equivocation/MsgDepFixedSetEquivocation.v | 570 ++++++++++++++++++ 2 files changed, 571 insertions(+) create mode 100644 theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v diff --git a/_CoqProject b/_CoqProject index 0d4411cd..ad4b4106 100644 --- a/_CoqProject +++ b/_CoqProject @@ -34,6 +34,7 @@ theories/VLSM/Core/Equivocation.v theories/VLSM/Core/EquivocationProjections.v theories/VLSM/Core/Equivocation/NoEquivocation.v theories/VLSM/Core/Equivocation/FixedSetEquivocation.v +theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v theories/VLSM/Core/Equivocation/WitnessedEquivocation.v theories/VLSM/Core/Equivocation/FullNode.v diff --git a/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v b/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v new file mode 100644 index 00000000..a294468a --- /dev/null +++ b/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v @@ -0,0 +1,570 @@ +From stdpp Require Import prelude finite. +From Coq Require Import Relations.Relation_Operators. +From VLSM.Lib Require Import Preamble StdppListSet FinFunExtras. +From VLSM.Core Require Import VLSM MessageDependencies VLSMProjections Composition Equivocation FixedSetEquivocation ProjectionTraces SubProjectionTraces. + +Section msg_dep_fixed_set_equivocation. + +Context + {message : Type} + `{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)) + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (equivocators : set index) + (Hbs_sub : forall sub_i, HasBeenSentCapability (sub_IM IM equivocators sub_i) := sub_has_been_sent_capabilities IM equivocators Hbs) + (Hbr_sub : forall sub_i, HasBeenReceivedCapability (sub_IM IM equivocators sub_i) := sub_has_been_received_capabilities IM equivocators Hbr) + (Hbo_sub := fun sub_i => HasBeenObservedCapability_from_sent_received (sub_IM IM equivocators sub_i)) + . + +Definition equivocator_can_emit (m : message) : Prop := + exists i, i ∈ equivocators /\ can_emit (pre_loaded_with_all_messages_vlsm (IM i)) m. + +Definition dependencies_with_non_equivocating_senders_were_sent s m : Prop := + forall dm, msg_dep_happens_before message_dependencies dm m -> + sent_by_non_equivocating IM _ equivocators s dm \/ equivocator_can_emit dm. + +Definition msg_dep_fixed_set_equivocation (s : composite_state IM) (m : message) := + sent_by_non_equivocating IM _ equivocators s m \/ + equivocator_can_emit m /\ + dependencies_with_non_equivocating_senders_were_sent s m. + +Definition msg_dep_fixed_set_equivocation_constraint + (l : composite_label IM) + (som : composite_state IM * option message) + : Prop := + let (s, om) := som in + from_option (msg_dep_fixed_set_equivocation s) True om. + +Definition msg_dep_fixed_set_equivocation_vlsm : VLSM message := + composite_vlsm IM msg_dep_fixed_set_equivocation_constraint. + +Lemma messages_with_valid_dependences_can_be_emitted s dm + (Hdepm: + forall dm0, msg_dep_rel message_dependencies dm0 dm -> + valid_message_prop (equivocators_composition_for_sent IM Hbs equivocators s) dm0) + (dm_i: index) + (Hdm_i: dm_i ∈ equivocators) + (Hemitted: can_emit (pre_loaded_with_all_messages_vlsm (IM dm_i)) dm) + : can_emit (equivocators_composition_for_sent IM Hbs equivocators s) dm. +Proof. + apply valid_preloaded_lifts_can_be_emitted with dm_i (fun m => m ∈ message_dependencies dm) + ; [assumption|assumption|]. + apply message_dependencies_are_sufficient with (Hbs dm_i) (Hbr dm_i); [|assumption]. + intuition. +Qed. + +Lemma msg_dep_rel_reflects_dependencies_with_non_equivocating_senders_were_sent s + : forall dm m, msg_dep_rel message_dependencies dm m -> + dependencies_with_non_equivocating_senders_were_sent s m -> + dependencies_with_non_equivocating_senders_were_sent s dm. +Proof. + intros dm m Hdm Hdeps dm0 Hdm0. + apply Hdeps. + transitivity dm; [assumption|]. + apply msg_dep_happens_before_iff_one. + left. + assumption. +Qed. + +Lemma dependencies_are_valid s m + (Hmsg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies)) + : dependencies_with_non_equivocating_senders_were_sent s m -> + forall dm, msg_dep_rel message_dependencies dm m -> + valid_message_prop + (equivocators_composition_for_sent IM Hbs equivocators s) dm. +Proof. + induction m as [m Hind] using (well_founded_ind Hmsg_dep_happens_before_wf). + intros Heqv dm Hdm. + apply emitted_messages_are_valid_iff. + assert (Hdm_hb : msg_dep_happens_before message_dependencies dm m) + by (apply msg_dep_happens_before_iff_one; left; assumption). + destruct (Heqv _ Hdm_hb) as [Hsent | [dm_i [Hdm_i Hemitted]]] + ; [left; right; assumption|]. + right. + apply messages_with_valid_dependences_can_be_emitted with dm_i + ; [|assumption|assumption]. + intros dm0 Hdm0. + apply Hind with dm; [assumption| |assumption]. + - clear -Heqv Hdm. + revert dm m Hdm Heqv. + apply msg_dep_rel_reflects_dependencies_with_non_equivocating_senders_were_sent. +Qed. + +Lemma msg_dep_strong_fixed_equivocation_subsumption s m + (Hmsg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies)) + : msg_dep_fixed_set_equivocation s m -> + strong_fixed_equivocation IM Hbs equivocators s m. +Proof. + intros [Hsent | [[i [Hi Hemit]] Heqv]]; [left; assumption|]. + cut (forall dm, msg_dep_rel message_dependencies dm m -> valid_message_prop (equivocators_composition_for_sent IM Hbs equivocators s) dm) + ; [|apply dependencies_are_valid; assumption]. + intro Hdeps. + right. + apply messages_with_valid_dependences_can_be_emitted with i; [|assumption|assumption]. + intros dm0 Hdm0. + apply Hdeps. + assumption. +Qed. + +Lemma msg_dep_strong_fixed_equivocation_constraint_subsumption + (Hmsg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies)) + : strong_constraint_subsumption IM + msg_dep_fixed_set_equivocation_constraint + (strong_fixed_equivocation_constraint IM Hbs equivocators). +Proof. + intros l (s, [m|]) Hc; [|exact I]. + apply msg_dep_strong_fixed_equivocation_subsumption; assumption. +Qed. + +Lemma single_equivocator_projection s j + (Hj : j ∈ equivocators) + : VLSM_projection + (equivocators_composition_for_sent IM Hbs equivocators s) + (pre_loaded_with_all_messages_vlsm (IM j)) + (sub_label_element_project IM equivocators j) + (sub_state_element_project IM equivocators j Hj). +Proof. + apply basic_VLSM_projection. + - intros lX lY HlX_pr sX om [_ [_ [HvX _]]] HsY _. + revert HvX. + destruct lX as [sub_i li]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + unfold sub_label_element_project in HlX_pr. + cbn in HlX_pr. + case_decide as Heqij; [|congruence]. + subst i. + cbv in HlX_pr. + apply Some_inj in HlX_pr. subst li. + unfold sub_IM, sub_state_element_project. + cbn. + match goal with + |- vvalid _ _ (?s,_) -> vvalid _ _ (?s',_) => + replace s with s' + end + ; [exact id | apply sub_IM_state_pi]. + - intros lX lY HlX_pr sX om sX' om' [_ HtX]. + revert HtX. + destruct lX as [sub_i li]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + unfold sub_label_element_project in HlX_pr. + cbn in HlX_pr. + case_decide as Heqij; [|congruence]. + subst i. + cbv in HlX_pr. + apply Some_inj in HlX_pr. subst li. + cbn. + unfold sub_IM, sub_state_element_project. + cbn. + match goal with + |- (let (_, _) := vtransition _ _ (?s,_) in _) = _ -> vtransition _ _ (?s',_) = _ => + replace s with s' + end + ; [| apply sub_IM_state_pi]. + destruct (vtransition _ _ _) as (sj', _om'). + intro Ht. inversion Ht. subst _om' sX'. clear Ht. + f_equal. + symmetry. + apply sub_IM_state_update_eq. + - intros lX HlX_pr sX om sX' om' [_ HtX]. + destruct lX as [sub_i li]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + unfold sub_label_element_project in HlX_pr. + cbn in HlX_pr. + case_decide as Heqij; [congruence|]. + cbn in HtX. + destruct (vtransition _ _ _) as (si', _om'). + inversion HtX. subst _om' sX'. clear HtX. + unfold sub_state_element_project. + rewrite sub_IM_state_update_neq by congruence. + reflexivity. + - intros sX HsX. + specialize (HsX (dexist j Hj)). + assumption. + - intro; intros. apply any_message_is_valid_in_preloaded. +Qed. + +Lemma equivocators_composition_can_emit_sender s m + : can_emit (equivocators_composition_for_sent IM Hbs equivocators s) m -> + equivocator_can_emit m. +Proof. + intros [(sX, iom) [(sub_i, li) [sX' HtX]]]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + exists i. split; [assumption|]. apply can_emit_iff. + apply + (VLSM_projection_input_valid_transition + (single_equivocator_projection s i Hi)) with (lY := li) in HtX + ; [eexists _,_,_; eassumption|]. + unfold sub_label_element_project. cbn. + case_decide as Heqi; [|contradiction]. + replace Heqi with (eq_refl (A := index) (x := i)); [reflexivity|]. + apply Eqdep_dec.UIP_dec. + assumption. +Qed. + +Lemma msg_dep_rel_reflects_strong_fixed_equivocation + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + : forall s, valid_state_prop (composite_vlsm IM (strong_fixed_equivocation_constraint IM Hbs equivocators)) s -> + forall dm m, msg_dep_rel message_dependencies dm m -> + strong_fixed_equivocation IM Hbs equivocators s m -> + strong_fixed_equivocation IM Hbs equivocators s dm. +Proof. + intros s Hs dm m Hdm [Hsent | Hemit]. + - destruct Hsent as [i [Hni Hsent]]. + apply valid_state_has_trace in Hs as Htr. + destruct Htr as [is [tr Htr]]. + eapply VLSM_incl_finite_valid_trace_init_to in Htr as Hpre_tr + ; [|apply constraint_preloaded_free_incl]. + apply + (VLSM_projection_finite_valid_trace_init_to + (preloaded_component_projection IM i)) + in Hpre_tr. + apply valid_trace_last_pstate in Hpre_tr as Hsi. + apply proper_sent in Hsent; [|assumption]. + specialize (Hsent _ _ Hpre_tr). + apply Exists_exists in Hsent as [item [Hitem Hout]]. + apply elem_of_list_In in Hitem. + apply pre_VLSM_projection_trace_project_in_iff in Hitem as HitemX. + destruct HitemX as [itemX [HitemX HitemX_pr]]. + destruct itemX. + unfold pre_VLSM_projection_transition_item_project, composite_project_label in HitemX_pr. + cbn in HitemX_pr. + case_decide as Hi; simpl in HitemX_pr; [|congruence]. + apply Some_inj in HitemX_pr. + apply in_split in Hitem as [pre [suf Htr_pr]]. + change (pre ++ item :: suf) with (pre ++ [item] ++ suf) in Htr_pr. + rewrite app_assoc in Htr_pr. + rewrite Htr_pr in Hpre_tr. + destruct Hpre_tr as [Hpre_tr Hinit]. + apply finite_valid_trace_from_to_app_split in Hpre_tr as [Hpre_tr Hsuf]. + rewrite finite_trace_last_is_last in Hsuf. + apply finite_valid_trace_from_to_app_split in Hpre_tr as Hitem. + apply proj2 in Hitem. + rewrite finite_trace_last_is_last in Hpre_tr, Hitem. + subst. + destruct l as (i, li). + cbn in *. + subst output. + inversion Hitem. + subst. + assert + (Hproduce : + can_produce (pre_loaded_with_all_messages_vlsm (IM i)) (destination i) m) + by (eexists _,_; eassumption). + eapply message_dependencies_are_necessary in Hproduce + ; [|apply HMsgDep|eassumption]. + eapply has_been_observed_sent_received_iff in Hproduce + ; [|revert Ht; apply input_valid_transition_destination]. + destruct Hproduce as [Hreceived | Hsent]; cycle 1. + + left. exists i. split; [assumption|]. + revert Hsent. + eapply in_futures_preserving_oracle_from_stepwise. + * apply has_been_sent_stepwise_from_trace. + * eexists; eassumption. + + apply pre_VLSM_projection_trace_project_app_rev in Htr_pr + as [preX [sufX [Heqtr [HpreX_pr HsufX_pr]]]]. + apply valid_trace_last_pstate in Hpre_tr as Hdestinationi. + apply proper_received in Hreceived; [|assumption]. + specialize (Hreceived _ _ (conj Hpre_tr Hinit)). + apply Exists_exists in Hreceived as [item_dm [Hitem_dm Hreceived]]. + apply elem_of_list_In in Hitem_dm. + rewrite <- HpreX_pr in Hitem_dm. + apply pre_VLSM_projection_trace_project_in_iff in Hitem_dm as HitemX_dm. + destruct HitemX_dm as [itemX_dm [HitemX_dm HitemX_dm_pr]]. + subst tr. + apply proj1 in Htr. + apply + (finite_valid_trace_from_to_app_split (composite_vlsm IM _)) + in Htr as [HpreX HsufX]. + apply in_split in HitemX_dm as [preX_dm [sufX_dm HpreX_dm_pr]]. + change (preX_dm ++ itemX_dm :: sufX_dm) with (preX_dm ++ [itemX_dm] ++ sufX_dm) in HpreX_dm_pr. + rewrite app_assoc in HpreX_dm_pr. + rewrite HpreX_dm_pr in HpreX. + apply + (finite_valid_trace_from_to_app_split (composite_vlsm IM _)) + in HpreX as [HpreX_dm HsufX_dm]. + apply + (finite_valid_trace_from_to_app_split (composite_vlsm IM _)), proj2 + in HpreX_dm. + rewrite finite_trace_last_is_last in HpreX_dm, HsufX_dm. + subst. + assert + (Hfutures : + in_futures (composite_vlsm IM (strong_fixed_equivocation_constraint IM Hbs equivocators)) + (@finite_trace_last _ (composite_type IM) is preX_dm) s). + { clear -HpreX_dm HsufX_dm HsufX. + exists ([itemX_dm] ++ sufX_dm ++ sufX). + apply (finite_valid_trace_from_to_app (composite_vlsm IM _)) with (VLSM.destination itemX_dm) + ; [assumption|]. + eapply (finite_valid_trace_from_to_app (composite_vlsm IM _)); eassumption. + } + destruct itemX_dm. + unfold pre_VLSM_projection_transition_item_project, composite_project_label in HitemX_dm_pr. + cbn in HitemX_dm_pr. + case_decide as Hi; simpl in HitemX_dm_pr; [|congruence]. + apply Some_inj in HitemX_dm_pr. + subst. + destruct l as (i, li_dm). + cbn in *. + subst. + clear -HpreX_dm Hfutures. + inversion HpreX_dm. + subst. + destruct Ht as [[_ [_ [_ Hc]]] _]. + revert Hc. + apply in_futures_preserves_strong_fixed_equivocation. + eapply VLSM_incl_in_futures + ; [apply constraint_preloaded_free_incl with (constraint := (strong_fixed_equivocation_constraint IM Hbs equivocators))|]. + assumption. + - destruct Hemit as [(sX, iom) [(sub_i, li) [sX' HtX]]]. + apply input_valid_transition_destination in HtX as HsX'. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + assert + (Hproduce :can_produce (pre_loaded_with_all_messages_vlsm (IM i)) (sX' (dexist i Hi)) m). + { + apply + (VLSM_projection_input_valid_transition + (single_equivocator_projection s i Hi)) with (lY := li) in HtX + ; [eexists _,_; eassumption|]. + unfold sub_label_element_project. cbn. + case_decide as Heqi; [|contradiction]. + replace Heqi with (eq_refl (A := index) (x := i)); [reflexivity|]. + apply Eqdep_dec.UIP_dec. + assumption. + } + eapply message_dependencies_are_necessary in Hproduce + ; [|apply HMsgDep|eassumption]. + assert (Hsent_comp : composite_has_been_observed (sub_IM IM equivocators) Hbo_sub sX' dm) + by (exists (dexist i Hi); assumption). + eapply preloaded_composite_observed_valid in Hsent_comp; [| apply listing_from_finite| eassumption..]. + apply emitted_messages_are_valid_iff in Hsent_comp as [[[sub_j [[_im Him] Heqim]] | Hsent] | Hemit] + ; [|left; assumption|right; assumption]. + destruct_dec_sig sub_j j Hj Heqsub_j. subst. + elim (no_initial_messages_in_IM j _im); assumption. +Qed. + +Lemma strong_fixed_equivocation_msg_dep_constraint_subsumption + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + : input_valid_constraint_subsumption IM + (strong_fixed_equivocation_constraint IM Hbs equivocators) + msg_dep_fixed_set_equivocation_constraint. +Proof. + intros l (s, [m|]) [Hs [_ [_ Hc]]]; [|exact I]. + cut (dependencies_with_non_equivocating_senders_were_sent s m). + { intros Hassume. + destruct Hc as [Hsent | Hemit]; [left; assumption|]. + right. + split; [|assumption]. + revert Hemit. apply equivocators_composition_can_emit_sender. + } + intros dm Hdm. + cut (strong_fixed_equivocation IM Hbs equivocators s dm). + { + intros [Hsent | Hemit]; [left; assumption|]. + right. + apply equivocators_composition_can_emit_sender with s. + assumption. + } + cbn in Hc. + revert Hdm Hc. + apply msg_dep_happens_before_reflect. + apply msg_dep_rel_reflects_strong_fixed_equivocation; assumption. +Qed. + +Lemma msg_dep_strong_fixed_equivocation_incl + (Hmsg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies)) + : VLSM_incl + (composite_vlsm IM msg_dep_fixed_set_equivocation_constraint) + (composite_vlsm IM (strong_fixed_equivocation_constraint IM Hbs equivocators)). +Proof. + apply constraint_subsumption_incl. + apply preloaded_constraint_subsumption_stronger. + apply strong_constraint_subsumption_strongest. + apply msg_dep_strong_fixed_equivocation_constraint_subsumption. + assumption. +Qed. + +Lemma strong_msg_dep_fixed_equivocation_incl + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + : VLSM_incl + (composite_vlsm IM (strong_fixed_equivocation_constraint IM Hbs equivocators)) + (composite_vlsm IM msg_dep_fixed_set_equivocation_constraint). +Proof. + apply constraint_subsumption_incl + with (constraint1 := strong_fixed_equivocation_constraint IM Hbs equivocators). + apply strong_fixed_equivocation_msg_dep_constraint_subsumption. + assumption. +Qed. + +Lemma msg_dep_strong_fixed_equivocation_eq + (Hmsg_dep_happens_before_wf : well_founded (msg_dep_happens_before message_dependencies)) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + : VLSM_eq + (composite_vlsm IM msg_dep_fixed_set_equivocation_constraint) + (composite_vlsm IM (strong_fixed_equivocation_constraint IM Hbs equivocators)). +Proof. + apply VLSM_eq_incl_iff; split. + - apply msg_dep_strong_fixed_equivocation_incl. assumption. + - apply strong_msg_dep_fixed_equivocation_incl. assumption. +Qed. + +End msg_dep_fixed_set_equivocation. + +Section sec_full_node_fixed_set_equivocation. + +Context + {message : Type} + `{EqDecision index} + (IM : index -> VLSM message) + (Hbs : forall i, HasBeenSentCapability (IM i)) + (message_dependencies : message -> set message) + (equivocators : set index) + {validator : Type} + (A : validator -> index) + (sender : message -> option validator) + . + +Definition has_equivocating_sender (m : message) + := exists v, sender m = Some v /\ A v ∈ equivocators. + + +Definition full_node_fixed_set_equivocation (s : composite_state IM) (m : message) := + sent_by_non_equivocating IM _ equivocators s m \/ has_equivocating_sender m. + +Definition full_node_fixed_set_equivocation_constraint + (l : composite_label IM) + (som : composite_state IM * option message) + : Prop := + let (s, om) := som in + from_option (full_node_fixed_set_equivocation s) True om. + +Lemma msg_dep_full_node_fixed_set_equivocation_constraint_subsumption + (Hchannel : channel_authentication_prop IM A sender) + : strong_constraint_subsumption IM + (msg_dep_fixed_set_equivocation_constraint IM Hbs message_dependencies equivocators) + full_node_fixed_set_equivocation_constraint. +Proof. + intros l (s, [m|]); [|intuition]. + intros [Hsent | [[i [Hi Hemit]] Hdeps]]; [left; assumption|]. + right. + apply Hchannel in Hemit. + cbv in Hemit |- *. + destruct (sender m) as [v|]; [|congruence]. + eexists. split; [reflexivity|]. + replace (A v) with i; [assumption|]. + congruence. +Qed. + +Context + (Hbr : forall i, HasBeenReceivedCapability (IM i)) + (Hbo := fun i => HasBeenObservedCapability_from_sent_received (IM i)) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + . + +Lemma fixed_full_node_equivocation_incl + {finite_index : finite.Finite index} + (Hchannel : channel_authentication_prop IM A sender) + : VLSM_incl + (composite_vlsm IM (fixed_equivocation_constraint IM Hbs Hbr equivocators)) + (composite_vlsm IM full_node_fixed_set_equivocation_constraint). +Proof. + eapply VLSM_incl_trans. + - apply Fixed_incl_StrongFixed with (enum index). + apply listing_from_finite. + - eapply VLSM_incl_trans. + + apply strong_msg_dep_fixed_equivocation_incl with Hbr; eassumption. + + apply constraint_subsumption_incl + with (constraint1 := msg_dep_fixed_set_equivocation_constraint IM Hbs message_dependencies equivocators). + apply preloaded_constraint_subsumption_stronger. + apply strong_constraint_subsumption_strongest. + apply msg_dep_full_node_fixed_set_equivocation_constraint_subsumption. + assumption. +Qed. + +Lemma full_node_fixed_equivocation_constraint_subsumption + (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + (Hsender_safety : sender_safety_alt_prop IM A sender) + : input_valid_constraint_subsumption IM + full_node_fixed_set_equivocation_constraint + (fixed_equivocation_constraint IM Hbs Hbr equivocators). +Proof. + intros l (s, [m|]) [_ [Hm [Hv Hc]]]; [|intuition]. + destruct Hc as [Hsent | Heqv]. + - left. revert Hsent. apply sent_by_non_equivocating_are_observed. + - right. + destruct l as (i, li). + destruct Heqv as [j [Hsender HAj]]. + apply Hfull in Hv. + eapply VLSM_incl_can_emit. + { + apply pre_loaded_vlsm_incl_relaxed + with (P := fun dm => composite_has_been_observed IM Hbo s dm \/ dm ∈ message_dependencies m). + intros m0 [Hsent_m0| Hdep_m0]; [intuition|]. + left. + exists i. + specialize (Hv _ Hdep_m0) as [Hsent | Hreceived] + ; [left | right]; assumption. + } + eapply VLSM_full_projection_can_emit. + { + apply preloaded_sub_element_full_projection with (P := fun dm => dm ∈ message_dependencies m). + intuition. + } + apply message_dependencies_are_sufficient with (Hbs (A j)) (Hbr (A j)); [apply HMsgDep|]. + cut (exists k, can_emit (pre_loaded_with_all_messages_vlsm (IM k)) m). + { + intros [k Hk]. + replace (A j) with k; [assumption|]. + symmetry. + eapply Hsender_safety; eassumption. + } + eapply @can_emit_composite_project + with (constraint := full_node_fixed_set_equivocation_constraint). + apply + (VLSM_incl_can_emit + (vlsm_incl_pre_loaded_with_all_messages_vlsm (composite_vlsm IM _))). + apply emitted_messages_are_valid_iff in Hm as [[k [[im Him] Heqm]] | Hemit] + ; [|assumption]. + elim (no_initial_messages_in_IM k im). + assumption. + Unshelve. + assumption. +Qed. + +Lemma full_node_fixed_equivocation_incl + (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + (Hsender_safety : sender_safety_alt_prop IM A sender) + : VLSM_incl + (composite_vlsm IM full_node_fixed_set_equivocation_constraint) + (composite_vlsm IM (fixed_equivocation_constraint IM Hbs Hbr equivocators)). +Proof. + apply constraint_subsumption_incl. + apply full_node_fixed_equivocation_constraint_subsumption; assumption. +Qed. + +Lemma full_node_fixed_equivocation_eq + {finite_index : finite.Finite index} + (Hchannel : channel_authentication_prop IM A sender) + (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + : VLSM_eq + (composite_vlsm IM full_node_fixed_set_equivocation_constraint) + (composite_vlsm IM (fixed_equivocation_constraint IM Hbs Hbr equivocators)). +Proof. + apply VLSM_eq_incl_iff; split. + - apply full_node_fixed_equivocation_incl; [assumption|]. + apply channel_authentication_sender_safety. + assumption. + - apply fixed_full_node_equivocation_incl. + assumption. +Qed. + +End sec_full_node_fixed_set_equivocation. From 39f7500d8e17a4209fafc1ef95bfe194acd36433 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:40:16 +0200 Subject: [PATCH 08/18] Annotated VLSMs --- _CoqProject | 1 + theories/VLSM/Core/AnnotatedVLSM.v | 195 +++++++++++++++++++++++++++++ 2 files changed, 196 insertions(+) create mode 100644 theories/VLSM/Core/AnnotatedVLSM.v diff --git a/_CoqProject b/_CoqProject index ad4b4106..9a371f91 100644 --- a/_CoqProject +++ b/_CoqProject @@ -29,6 +29,7 @@ theories/VLSM/Core/Decisions.v theories/VLSM/Core/MessageDependencies.v theories/VLSM/Core/ProjectionTraces.v theories/VLSM/Core/SubProjectionTraces.v +theories/VLSM/Core/AnnotatedVLSM.v theories/VLSM/Core/Equivocation.v theories/VLSM/Core/EquivocationProjections.v diff --git a/theories/VLSM/Core/AnnotatedVLSM.v b/theories/VLSM/Core/AnnotatedVLSM.v new file mode 100644 index 00000000..1632dd2c --- /dev/null +++ b/theories/VLSM/Core/AnnotatedVLSM.v @@ -0,0 +1,195 @@ +From stdpp Require Import prelude. +From VLSM.Lib Require Import ListExtras. +From VLSM.Core Require Import VLSM VLSMProjections. + +Section sec_annotated_vlsm. + +Context + {message : Type} + (X : VLSM message) + (annotation : Type) + . + +Record annotated_state := + { original_state : vstate X + ; state_annotation : annotation + }. + +Definition annotated_type : VLSMType message := + {| label := vlabel X; + state := annotated_state + |}. + +Context + (initial_annotation_prop : annotation -> Prop) + `{Inhabited (sig initial_annotation_prop)} + . + +Definition annotated_initial_state_prop (sa : annotated_state) := + vinitial_state_prop X (original_state sa) /\ initial_annotation_prop (state_annotation sa). + +Global Program Instance annotated_initial_state_prop_inhabited + : Inhabited (sig annotated_initial_state_prop) := + populate (exist _ {| original_state := ` (vs0 X); state_annotation := ` inhabitant |} _). +Next Obligation. + split; cbn. + - destruct (vs0 X). assumption. + - destruct inhabitant. assumption. +Qed. + +Instance annotated_sign : VLSMSign annotated_type := + { initial_state_prop := annotated_initial_state_prop + ; initial_message_prop := λ m : message, vinitial_message_prop X m + }. + +Context + (annotated_constraint : @label _ annotated_type -> annotated_state * option message -> Prop) + (annotated_transition_state : @label _ annotated_type -> annotated_state * option message -> annotation) + . + +Definition annotated_valid + (l : @label _ annotated_type) + (som : annotated_state * option message) + : Prop := + vvalid X l (original_state som.1, som.2) /\ + annotated_constraint l som. + +Definition annotated_transition + (l : @label _ annotated_type) + (som : annotated_state * option message) + : annotated_state * option message := + let (s', om') := vtransition X l (original_state som.1, som.2) in + ({| original_state := s'; state_annotation := annotated_transition_state l som |}, om'). + +Instance annotated_vlsm_class : VLSMClass annotated_sign := + { valid := annotated_valid + ; transition := annotated_transition + }. + +Definition annotated_vlsm : VLSM message := mk_vlsm annotated_vlsm_class. + +Definition annotate_trace_item + (item : vtransition_item X) + (k : annotated_state -> list (@transition_item _ annotated_type)) + (sa : annotated_state) + : list (@transition_item _ annotated_type) := + match item with + | {| l := l; input := iom; destination := s'; output := oom |} => + let sa' := {| original_state := s'; state_annotation := annotated_transition_state l (sa, iom) |} in + @Build_transition_item _ annotated_type l iom sa' oom :: k sa' + end. + +Lemma annotate_trace_item_project + (item : vtransition_item X) + (k : annotated_state -> list (@transition_item _ annotated_type)) + (sa : annotated_state) + : pre_VLSM_full_projection_finite_trace_project + annotated_type (type X) id original_state + (annotate_trace_item item k sa) + = item :: + pre_VLSM_full_projection_finite_trace_project + annotated_type (type X) id original_state + (k {| original_state := destination item; state_annotation := annotated_transition_state (l item) (sa, input item) |}). +Proof. + destruct item. reflexivity. +Qed. + +Definition annotate_trace_from (sa : @state _ annotated_type) (tr : list (vtransition_item X)) + : list (@transition_item _ annotated_type) := + fold_right annotate_trace_item (fun sa => []) tr sa. + +Lemma annotate_trace_from_unroll sa item tr + : annotate_trace_from sa (item :: tr) = + let sa' := {| original_state := destination item; state_annotation := annotated_transition_state (l item) (sa, input item) |} in + @Build_transition_item _ annotated_type (l item) (input item) sa' (output item) :: annotate_trace_from sa' tr. +Proof. + destruct item; reflexivity. +Qed. + +Lemma annotate_trace_from_app sa tr1 tr2 + : annotate_trace_from sa (tr1 ++ tr2) = + annotate_trace_from sa tr1 ++ + annotate_trace_from (finite_trace_last sa ( annotate_trace_from sa tr1)) tr2. +Proof. + revert sa. + induction tr1 as [|item tr1]; [reflexivity|]. + intro sa. + change ((item :: tr1) ++ tr2) with (item :: (tr1 ++ tr2)). + rewrite !annotate_trace_from_unroll. + simpl. + rewrite IHtr1. + f_equal. + f_equal. + f_equal. + rewrite finite_trace_last_cons. + reflexivity. +Qed. + +Lemma annotate_trace_from_last_original_state sa tr + : original_state (finite_trace_last sa (annotate_trace_from sa tr)) = + finite_trace_last (original_state sa) tr. +Proof. + destruct_list_last tr tr' item Heqtr; subst; [reflexivity|]. + rewrite annotate_trace_from_app. + destruct item. + cbn. + rewrite! finite_trace_last_is_last. + reflexivity. +Qed. + +Definition annotate_trace (s : vstate X) (tr : list (vtransition_item X)) + : list (@transition_item _ annotated_type) := + annotate_trace_from {| original_state := s; state_annotation := ` inhabitant |} tr. + +Lemma annotate_trace_project is tr + : pre_VLSM_full_projection_finite_trace_project + annotated_type (type X) id original_state + (annotate_trace is tr) + = tr. +Proof. + unfold annotate_trace. + remember {| original_state := is |} as sa. + clear Heqsa. + revert sa. + induction tr as [| item]; [reflexivity|]. + intro sa. + cbn. + rewrite annotate_trace_item_project. + f_equal. + apply IHtr. +Qed. + +End sec_annotated_vlsm. + +Arguments original_state {_ _ _} _ : assert. +Arguments state_annotation {_ _ _} _ : assert. + +Section sec_annotated_vlsm_projections. + +Context + {message : Type} + (X : VLSM message) + {annotation : Type} + (initial_annotation_prop : annotation -> Prop) + `{Inhabited (sig initial_annotation_prop)} + (annotated_constraint : @label _ (annotated_type X annotation) -> annotated_state X annotation * option message -> Prop) + (annotated_transition_state : @label _ (annotated_type X annotation) -> annotated_state X annotation * option message -> annotation) + (AnnotatedX : VLSM message := annotated_vlsm X annotation initial_annotation_prop annotated_constraint annotated_transition_state) + . + +Definition forget_annotations_projection + : VLSM_full_projection AnnotatedX X id original_state. +Proof. + apply basic_VLSM_strong_full_projection. + - cbv; intuition. + - intros l (s,a) om (s', a') om'. + cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (_s', _om'). + inversion 1; reflexivity. + - cbv; intuition. + - cbv; intuition. +Qed. + +End sec_annotated_vlsm_projections. From 41d84cb70d00efdbb84a1ad1308384c0fcd25297 Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:40:48 +0200 Subject: [PATCH 09/18] Limited equivocation based on message dependencies --- _CoqProject | 1 + .../Equivocation/MsgDepLimitedEquivocation.v | 782 ++++++++++++++++++ 2 files changed, 783 insertions(+) create mode 100644 theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v diff --git a/_CoqProject b/_CoqProject index 9a371f91..5006aa87 100644 --- a/_CoqProject +++ b/_CoqProject @@ -40,6 +40,7 @@ theories/VLSM/Core/Equivocation/TraceWiseEquivocation.v theories/VLSM/Core/Equivocation/WitnessedEquivocation.v theories/VLSM/Core/Equivocation/FullNode.v theories/VLSM/Core/Equivocation/LimitedEquivocation.v +theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v theories/VLSM/Core/Equivocators/Common.v theories/VLSM/Core/Equivocators/Projections.v diff --git a/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v new file mode 100644 index 00000000..d8a738a9 --- /dev/null +++ b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v @@ -0,0 +1,782 @@ +From stdpp Require Import prelude finite. +From Coq Require Import Reals. +From VLSM.Lib Require Import Preamble ListExtras StdppListSet ListSetExtras FinFunExtras Measurable. +From VLSM.Core Require Import VLSM AnnotatedVLSM MessageDependencies VLSMProjections Composition ProjectionTraces SubProjectionTraces. +From VLSM.Core Require Import Equivocation Equivocation.FixedSetEquivocation Equivocation.TraceWiseEquivocation Equivocation.LimitedEquivocation Equivocation.MsgDepFixedSetEquivocation. + +Section sec_coequivocating_senders_limited_equivocation. + +Context + {message : Type} + `{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)) + `{EqDecision validator} + `{ReachableThreshold validator} + (A : validator -> index) + (sender : message -> option validator) + (coequivocating_senders : composite_state IM -> message -> set validator) + . + +Definition coeqv_message_equivocators (s : composite_state IM) (m : message) + : set validator := + if (decide (composite_has_been_observed IM Hbo s m)) + then [] (* no additional equivocation *) + else map_option sender [m] ++ coequivocating_senders s m. + (* m itself and all its non-observed dependencies are equivocating. *) + +Definition coeqv_composite_transition_message_equivocators + (l : composite_label IM) + (som : annotated_state (free_composite_vlsm IM) (set validator) * option message) + : set validator := + match som.2 with + | None => state_annotation som.1 + | Some m => + set_union (state_annotation som.1) (coeqv_message_equivocators (original_state som.1) m) + end. + +Definition coeqv_limited_equivocation_constraint + (l : composite_label IM) + (som : annotated_state (free_composite_vlsm IM) (set validator) * option message) + : Prop := + (sum_weights (coeqv_composite_transition_message_equivocators l som) <= proj1_sig threshold)%R. + +Global Instance empty_validators_inhabited : Inhabited {s : set validator | s = empty_set} + := populate (exist _ _ eq_refl). + +Definition coeqv_limited_equivocation_vlsm : VLSM message := + annotated_vlsm (free_composite_vlsm IM) (set validator) (fun s => s = empty_set) + coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. + +Definition coeqv_annotate_trace_with_equivocators := + annotate_trace (free_composite_vlsm IM) (set validator) (fun s => s = empty_set) + coeqv_composite_transition_message_equivocators. + +Lemma coeqv_limited_equivocation_transition_state_annotation_incl [l s iom s' oom] + : vtransition coeqv_limited_equivocation_vlsm l (s, iom) = (s', oom) -> + state_annotation s ⊆ state_annotation s'. +Proof. + cbn. + unfold annotated_transition. + destruct (vtransition _ _ _) as (_s', _om'). + inversion 1. + cbn. + destruct iom as [m|]; [|reflexivity]. + apply set_union_subseteq_left. +Qed. + +Lemma coeqv_limited_equivocation_state_annotation_nodup s + : valid_state_prop coeqv_limited_equivocation_vlsm s -> + NoDup (state_annotation s). +Proof. + induction 1 using valid_state_prop_ind. + - destruct s. + cbn. + replace state_annotation with (@nil validator); [constructor|]. + symmetry. + apply Hs. + - apply proj2 in Ht. + cbn in Ht. + unfold annotated_transition in Ht. + destruct (vtransition _ _ _). + inversion Ht. + destruct om as [m|]; [|assumption]. + cbn. + apply set_union_nodup_left. + assumption. +Qed. + +Lemma coeqv_limited_equivocation_state_not_heavy s + : valid_state_prop coeqv_limited_equivocation_vlsm s -> + (sum_weights (state_annotation s) <= proj1_sig threshold)%R. +Proof. + induction 1 using valid_state_prop_ind. + - destruct s. + cbn. + replace state_annotation with (@nil validator) + by (symmetry; apply Hs). + destruct threshold. simpl. apply Rge_le. assumption. + - destruct Ht as [[_ [_ [_ Hc]]] Ht]. + cbn in Ht. + unfold annotated_transition in Ht. + destruct (vtransition _ _ _). + inversion Ht. + destruct om as [m|]; assumption. +Qed. + +End sec_coequivocating_senders_limited_equivocation. + +Section sec_msg_dep_limited_equivocation. + +Context + {message : Type} + `{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)) + (full_message_dependencies : message -> set message) + `{EqDecision validator} + `{ReachableThreshold validator} + (A : validator -> index) + (sender : message -> option validator) + . + +Definition not_observed_happens_before_dependencies (s : composite_state IM) (m : message) + : set message := + filter (fun dm => ~composite_has_been_observed IM Hbo s dm) (full_message_dependencies m). + +Definition msg_dep_coequivocating_senders (s : composite_state IM) (m : message) + : set validator := + map_option sender (not_observed_happens_before_dependencies s m). + +Definition msg_dep_limited_equivocation_vlsm : VLSM message := + coeqv_limited_equivocation_vlsm IM Hbs Hbr sender msg_dep_coequivocating_senders. + +Definition msg_dep_message_equivocators := + coeqv_message_equivocators IM Hbs Hbr sender msg_dep_coequivocating_senders. + +Definition msg_dep_annotate_trace_with_equivocators := + coeqv_annotate_trace_with_equivocators IM Hbs Hbr sender msg_dep_coequivocating_senders. + +Definition msg_dep_composite_transition_message_equivocators := + coeqv_composite_transition_message_equivocators IM Hbs Hbr sender msg_dep_coequivocating_senders. + +End sec_msg_dep_limited_equivocation. + +Section sec_full_node_limited_equivocation. + +Context + {message : Type} + `{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)) + `{EqDecision validator} + `{ReachableThreshold validator} + (A : validator -> index) + (sender : message -> option validator) + . + +Definition full_node_coequivocating_senders (s : composite_state IM) (m : message) + : set validator := + []. + +Definition full_node_limited_equivocation_vlsm : VLSM message := + coeqv_limited_equivocation_vlsm IM Hbs Hbr sender full_node_coequivocating_senders. + +Definition full_node_composite_transition_message_equivocators := + coeqv_composite_transition_message_equivocators IM Hbs Hbr sender full_node_coequivocating_senders. + +End sec_full_node_limited_equivocation. + +Section sec_full_node_msg_dep_limited_equivocation_equivalence. + +Context + {message : Type} + `{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)) + (full_message_dependencies : message -> set message) + `{EqDecision validator} + `{ReachableThreshold validator} + (A : validator -> index) + (sender : message -> option validator) + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + . + +Lemma full_node_msg_dep_coequivocating_senders s m i li + (Hvalid : input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (s i, Some m)) + : msg_dep_coequivocating_senders IM Hbs Hbr full_message_dependencies sender s m = []. +Proof. + cut (forall v, ~ v ∈ msg_dep_coequivocating_senders IM Hbs Hbr full_message_dependencies sender s m ). + { + intros Hv. + destruct (msg_dep_coequivocating_senders IM Hbs Hbr full_message_dependencies sender s m) + ; [reflexivity|]. + elim (Hv v). + left. + } + intros v. + setoid_rewrite elem_of_map_option. + setoid_rewrite elem_of_list_filter. + intros [dm [[Hnobs Hdm] Hsender]]. + unfold msg_dep_coequivocating_senders. + elim Hnobs. + exists i. + clear -Hbo HMsgDep HFullMsgDep Hfull Hvalid Hdm. + revert dm Hdm. + cut + (forall dm m, msg_dep_rel message_dependencies dm m -> + @has_been_observed _ (IM i) (Hbo i) (s i) m -> + @has_been_observed _ (IM i) (Hbo i) (s i) dm). + { + intros Hdeps dm. + rewrite full_message_dependencies_happens_before. + rewrite msg_dep_happens_before_iff_one. + intros [Hdm | [dm' [Hdm' Hdm]]]; [|]. + - eapply Hfull; [apply Hvalid|eassumption]. + - eapply msg_dep_happens_before_reflect; [eassumption|eassumption|]. + eapply Hfull; [apply Hvalid|eassumption]. + } + clear -HMsgDep Hfull Hvalid. + apply msg_dep_full_node_reflects_has_been_observed; [apply HMsgDep|apply Hfull|]. + apply Hvalid. +Qed. + +Lemma full_node_msg_dep_composite_transition_message_equivocators + iprop `{Inhabited (sig iprop)} constr trans + l (s : @state _ (annotated_type (free_composite_vlsm IM) (set validator))) om + (Hvalid : input_valid (annotated_vlsm (free_composite_vlsm IM) (set validator) iprop constr trans) l (s, om)) + : full_node_composite_transition_message_equivocators IM Hbs Hbr sender l (s, om) = + msg_dep_composite_transition_message_equivocators IM Hbs Hbr full_message_dependencies sender l (s, om). +Proof. + destruct om as [m|]; [|reflexivity]. + cbn. + f_equal. + unfold coeqv_message_equivocators. + case_decide as Hobs; [reflexivity|]. + f_equal. + symmetry. + destruct l as (i, li). + eapply full_node_msg_dep_coequivocating_senders. + eapply + (VLSM_projection_input_valid (preloaded_component_projection IM i)) + with (lX := existT i li) + ; [apply (composite_project_label_eq IM)|]. + apply (VLSM_incl_input_valid (vlsm_incl_pre_loaded_with_all_messages_vlsm (free_composite_vlsm IM))). + revert Hvalid. + apply + (VLSM_full_projection_input_valid + (forget_annotations_projection (free_composite_vlsm IM) _ _ _)). +Qed. + +Lemma msg_dep_full_node_limited_equivocation_vlsm_incl + : VLSM_incl + (msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) + (full_node_limited_equivocation_vlsm IM Hbs Hbr sender). +Proof. + apply basic_VLSM_incl. + - intros s Hs. + assumption. + - intros _ _ m _ _ Hinit. + apply initial_message_is_valid. + assumption. + - intros l s om HvX HsY HomY. + split; [apply HvX|]. + unfold coeqv_limited_equivocation_constraint. + setoid_rewrite full_node_msg_dep_composite_transition_message_equivocators + ; [apply HvX|eassumption]. + - intros (i, li) s iom s' oom [Hv Ht]. + revert Ht. cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (si', om'). + inversion 1. + subst. + clear Ht. + setoid_rewrite full_node_msg_dep_composite_transition_message_equivocators + ; [reflexivity|eassumption]. +Qed. + +Lemma full_node_msg_dep_limited_equivocation_vlsm_incl + : VLSM_incl + (full_node_limited_equivocation_vlsm IM Hbs Hbr sender) + (msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender). +Proof. + apply basic_VLSM_incl. + - intros s Hs. + assumption. + - intros _ _ m _ _ Hinit. + apply initial_message_is_valid. + assumption. + - intros l s om HvX HsY HomY. + split; [apply HvX|]. + unfold coeqv_limited_equivocation_constraint. + setoid_rewrite <- full_node_msg_dep_composite_transition_message_equivocators + ; [apply HvX|eassumption]. + - intros (i, li) s iom s' oom [Hv Ht]. + revert Ht. cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (si', om'). + inversion 1. + subst. + clear Ht. + setoid_rewrite full_node_msg_dep_composite_transition_message_equivocators + ; [reflexivity|eassumption]. +Qed. + +Lemma full_node_msg_dep_limited_equivocation_vlsm_eq + : VLSM_eq + (full_node_limited_equivocation_vlsm IM Hbs Hbr sender) + (msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender). +Proof. + apply VLSM_eq_incl_iff. + split. + - apply full_node_msg_dep_limited_equivocation_vlsm_incl. + - apply msg_dep_full_node_limited_equivocation_vlsm_incl. +Qed. + +End sec_full_node_msg_dep_limited_equivocation_equivalence. + +Section sec_msg_dep_fixed_limited_equivocation. + +Context + {message : Type} + `{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)) + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + `{ReachableThreshold index} + (sender : message -> option index) + (Limited := msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (Hchannel : channel_authentication_prop IM Datatypes.id sender) + (Hsender_safety : sender_safety_alt_prop IM Datatypes.id sender := + channel_authentication_sender_safety _ _ _ Hchannel) + . + +Lemma message_equivocators_can_emit (s : vstate Limited) im + (Hs : + valid_state_prop + (fixed_equivocation_vlsm_composition IM Hbs Hbr (state_annotation s)) + (original_state s)) + (Hnobserved: ¬ composite_has_been_observed IM Hbo (original_state s) im) + (HLemit: can_emit Limited im) + : can_emit + (equivocators_composition_for_observed IM Hbs Hbr + (set_union (state_annotation s) + (msg_dep_message_equivocators IM Hbs Hbr full_message_dependencies sender + (original_state s) im)) + (original_state s)) + im. +Proof. + unfold msg_dep_message_equivocators, coeqv_message_equivocators. + rewrite decide_False by assumption. + remember (map_option sender _ ++ _) as equivocating_senders. + assert + (Hsenders : + forall dm, msg_dep_happens_before message_dependencies dm im -> + composite_has_been_observed IM Hbo (original_state s) dm \/ + exists dm_i, dm_i ∈ equivocating_senders /\ + can_emit (pre_loaded_with_all_messages_vlsm (IM dm_i)) dm). + { + intros dm Hdm. + destruct (decide (composite_has_been_observed IM Hbo (original_state s) dm)) + as [Hobs | Hnobs] + ; [left; assumption|]. + right. + cut + (exists i, sender dm = Some i /\ + can_emit (pre_loaded_with_all_messages_vlsm (IM i)) dm). + { + intros [i [Hsender Hemit]]. + exists i. + split; [|assumption]. + subst. + apply elem_of_app. + right. + apply elem_of_map_option. + exists dm. + split; [|assumption]. + apply elem_of_list_filter. + split; [assumption|]. + rewrite full_message_dependencies_happens_before. + assumption. + } + eapply msg_dep_happens_before_composite_no_initial_valid_messages_emitted_by_sender + ; [eassumption|eassumption|eassumption|eassumption|..|eassumption]. + apply emitted_messages_are_valid_iff. + right. + revert HLemit. + eapply VLSM_full_projection_can_emit. + apply forget_annotations_projection. + } + eapply VLSM_full_projection_can_emit. + { + apply equivocators_composition_for_observed_index_incl_full_projection + with (indices1 := equivocating_senders). + } + assert + (Hemitj : exists j, j ∈ equivocating_senders /\ + can_emit + (pre_loaded_vlsm (IM j) + (fun dm => msg_dep_happens_before message_dependencies dm im)) + im). + { + eapply VLSM_full_projection_can_emit in HLemit + ; [|apply forget_annotations_projection]. + eapply VLSM_incl_can_emit in HLemit + ; [| apply vlsm_incl_pre_loaded_with_all_messages_vlsm]. + apply can_emit_composite_project in HLemit as [j Hemit]. + apply Hchannel in Hemit as Hsender. + unfold channel_authenticated_message in Hsender. + exists j. + subst equivocating_senders. + cbn. + destruct (sender im) as [_j|]; [|inversion Hsender]. + apply Some_inj in Hsender. + cbn in Hsender. + subst _j. + split; [left|]. + apply HMsgDep in Hemit. + revert Hemit. + apply VLSM_incl_can_emit. + apply pre_loaded_vlsm_incl. + intros m Hdm. + apply msg_dep_happens_before_iff_one. + left. + assumption. + } + destruct Hemitj as [j [Heqv_j Hemitj]]. + apply valid_preloaded_lifts_can_be_emitted with j (fun m => msg_dep_happens_before message_dependencies m im) + ; [assumption| |assumption]. + clear Heqequivocating_senders. + induction dm as [dm Hind] + using (well_founded_ind (msg_dep_happens_before_wf _ _ HFullMsgDep)). + intros Hdm. + apply emitted_messages_are_valid_iff. + destruct (Hsenders _ Hdm) as [Hobs_dm | [dm_i [Hdm_i Hemit_dm]]] + ; [left; right; assumption|]. + right. + apply valid_preloaded_lifts_can_be_emitted with dm_i (fun m => m ∈ message_dependencies dm) + ; [assumption|..]; cycle 1. + - eapply message_dependencies_are_sufficient; [apply HMsgDep|assumption]. + - intros dm' Hdm'. + apply Hind. + + apply msg_dep_happens_before_iff_one. left. assumption. + + transitivity dm; [|assumption]. + apply msg_dep_happens_before_iff_one. left. assumption. +Unshelve. + apply set_union_subseteq_right. +Qed. + +Lemma msg_dep_fixed_limited_equivocation is tr + : finite_valid_trace Limited is tr -> + fixed_limited_equivocation_prop IM Hbs Hbr + (original_state is) + (pre_VLSM_full_projection_finite_trace_project + (type Limited) (composite_type IM) Datatypes.id original_state + tr). +Proof. + intro Htr. + exists (state_annotation (finite_trace_last is tr)). + split. + - rewrite set_eq_nodup_sum_weight_eq + with (lv2 := state_annotation (finite_trace_last is tr)). + + apply coeqv_limited_equivocation_state_not_heavy. + apply finite_valid_trace_last_pstate. + apply Htr. + + apply NoDup_remove_dups. + + apply coeqv_limited_equivocation_state_annotation_nodup. + apply finite_valid_trace_last_pstate. + apply Htr. + + apply set_eq_extract_forall. + setoid_rewrite elem_of_remove_dups. + intuition. + - split; [|apply Htr]. + apply valid_trace_add_default_last in Htr. + induction Htr using finite_valid_trace_init_to_rev_ind + ; [constructor; apply initial_state_is_valid; apply Hsi|]. + setoid_rewrite map_app. + apply finite_valid_trace_from_app_iff. + split. + + revert IHHtr. + apply VLSM_incl_finite_valid_trace_from. + apply fixed_equivocation_vlsm_composition_index_incl. + eapply coeqv_limited_equivocation_transition_state_annotation_incl. + apply Ht. + + apply finite_valid_trace_singleton. + unfold input_valid_transition, input_valid. + change (map _ _) with (pre_VLSM_full_projection_finite_trace_project (type Limited) (composite_type IM) Datatypes.id original_state tr). + rewrite <- pre_VLSM_full_projection_finite_trace_last. + assert + (Hs : + valid_state_prop + (fixed_equivocation_vlsm_composition IM Hbs Hbr (state_annotation s)) + (original_state s)). + { + replace s with (finite_trace_last si tr) at 2 + by (apply valid_trace_get_last in Htr; assumption). + rewrite + (pre_VLSM_full_projection_finite_trace_last + (type Limited) (composite_type IM) Datatypes.id original_state + si tr). + apply finite_valid_trace_last_pstate. + assumption. + } + destruct Ht as [[HLs [HLim HLv]] HLt]. + cbn in HLt. + unfold annotated_transition in HLt. + cbn. + replace (finite_trace_last si _) with s + by (apply valid_trace_get_last in Htr; congruence). + cbn in HLt. + destruct l as (i, li). + destruct (vtransition _ _ _) as (si', om'). + inversion HLt. + subst. + clear HLt. + cbn. + repeat split. + * revert Hs. + apply VLSM_incl_valid_state. + apply fixed_equivocation_vlsm_composition_index_incl. + destruct iom as [im|]; [apply set_union_subseteq_left|reflexivity]. + * destruct iom as [im|]; [apply option_valid_message_Some|apply option_valid_message_None]. + destruct (decide (composite_has_been_observed IM Hbo (original_state s) im)) + as [Hobs|Hnobs]. + -- revert Hobs. + apply composite_observed_valid with (enum index). + ++ apply listing_from_finite. + ++ assumption. + ++ assumption. + ++ revert Hs. + apply VLSM_incl_valid_state. + apply fixed_equivocation_vlsm_composition_index_incl. + apply set_union_subseteq_left. + -- revert HLim. + setoid_rewrite emitted_messages_are_valid_iff. + intros [Hinit | Hemit]; [left; assumption|]. + right. + eapply VLSM_weak_full_projection_can_emit. + { + eapply EquivPreloadedBase_Fixed_weak_full_projection + with (base_s := original_state s). + - apply listing_from_finite. + - revert Hs. + apply VLSM_incl_valid_state. + apply fixed_equivocation_vlsm_composition_index_incl. + apply set_union_subseteq_left. + - intros. apply no_initial_messages_in_IM. + } + cbn. + eapply VLSM_incl_can_emit. + { + apply Equivocators_Fixed_Strong_incl with (enum index). + - apply listing_from_finite. + - revert Hs. + apply VLSM_incl_valid_state. + apply fixed_equivocation_vlsm_composition_index_incl. + apply set_union_subseteq_left. + } + apply message_equivocators_can_emit; assumption. + * apply HLv. + * destruct iom as [im|]; [|exact I]. + destruct (decide (composite_has_been_observed IM Hbo (original_state s) im)) + as [Hobs|Hnobs]; + [left; assumption|]. + right. + cbn. + apply message_equivocators_can_emit; [assumption|assumption|]. + cut (vinitial_message_prop Limited im \/ can_emit Limited im); + [|apply emitted_messages_are_valid_iff; assumption]. + intros [[j [[mj Hmj] Heqim]] | Hemit]; [|assumption]. + elim (no_initial_messages_in_IM j mj). + assumption. +Qed. + +Lemma fixed_transition_preserves_annotation_equivocators + equivocators + (is : vstate (free_composite_vlsm IM)) s tr + (Htr1 : + finite_valid_trace_init_to (fixed_equivocation_vlsm_composition IM Hbs Hbr equivocators) + is s tr) + l iom sf oom + (Ht : + input_valid_transition + (fixed_equivocation_vlsm_composition IM Hbs Hbr equivocators) l + (s, iom) (sf, oom)) + (Hsub_equivocators : + state_annotation + (@finite_trace_last _ (type Limited) + {| original_state := is; state_annotation := `inhabitant |} + (msg_dep_annotate_trace_with_equivocators IM Hbs Hbr full_message_dependencies sender is tr)) + ⊆ equivocators) + : msg_dep_composite_transition_message_equivocators IM Hbs Hbr + full_message_dependencies sender l + (@finite_trace_last _ (type Limited) + {| original_state := is; state_annotation := empty_set |} + (annotate_trace_from (free_composite_vlsm IM) + (set index) + (msg_dep_composite_transition_message_equivocators IM Hbs Hbr full_message_dependencies sender) + {| original_state := is; state_annotation := empty_set |} tr), iom) + ⊆ equivocators. +Proof. + destruct iom as [im|]; [|assumption]. + apply set_union_subseteq_iff. + split; [assumption|]. + cbn. + rewrite annotate_trace_from_last_original_state. + cbn. + replace (finite_trace_last _ _) with s + by (apply valid_trace_get_last in Htr1; congruence). + unfold coeqv_message_equivocators. + case_decide as Hnobserved; [apply list_subseteq_nil|]. + destruct Ht as [[Hs [Him [Hv [Hobs | Hemitted]]]] Ht]; [contradiction|]. + intros eqv Heqv. + apply elem_of_app in Heqv as [Heqv|Heqv] + ; apply elem_of_map_option in Heqv as [msg [Hmsg Hsender]]. + - apply elem_of_list_singleton in Hmsg. + subst msg. + eapply VLSM_incl_can_emit in Hemitted + ; [|apply pre_loaded_vlsm_incl_pre_loaded_with_all_messages]. + apply can_emit_composite_project in Hemitted as [sub_eqv Hemitted]. + destruct_dec_sig sub_eqv _eqv H_eqv Heqsub_eqv. + subst. + unfold sub_IM in Hemitted. + cbn in Hemitted. + eapply Hsender_safety in Hemitted; [|eassumption]. + subst. + assumption. + - apply elem_of_list_filter in Hmsg as [Hnobserved_msg Hdep_msg]. + cut (strong_fixed_equivocation IM Hbs equivocators s msg). + { + intros [Hobserved | Hemitted_msg]. + - elim Hnobserved_msg. + eapply sent_by_non_equivocating_are_observed. + eassumption. + - eapply VLSM_incl_can_emit in Hemitted_msg + ; [|apply pre_loaded_vlsm_incl_pre_loaded_with_all_messages]. + apply can_emit_composite_project in Hemitted_msg as [sub_i Hemitted_msg]. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + eapply Hsender_safety in Hemitted_msg; [|eassumption]. + cbn in Hemitted_msg. + subst. + assumption. + } + eapply msg_dep_happens_before_reflect + ; [|eapply full_message_dependencies_happens_before; eassumption|]; cycle 1. + + right. + revert Hemitted. + apply VLSM_incl_can_emit. + apply Equivocators_Fixed_Strong_incl with (enum index) + ; [apply listing_from_finite|assumption]. + + eapply msg_dep_rel_reflects_strong_fixed_equivocation + ; [assumption|assumption|..]. + revert Hs. + apply VLSM_incl_valid_state. + apply Fixed_incl_StrongFixed with (enum index). + apply listing_from_finite. +Qed. + +Lemma msg_dep_limited_fixed_equivocation + (is : vstate (free_composite_vlsm IM)) (tr : list (composite_transition_item IM)) + : fixed_limited_equivocation_prop IM Hbs Hbr is tr -> + finite_valid_trace Limited + {| original_state := is; state_annotation := ` inhabitant |} + (msg_dep_annotate_trace_with_equivocators IM Hbs Hbr full_message_dependencies sender is tr). +Proof. + intros [equivocators [Hlimited Htr]]. + split; [|split; [apply Htr|reflexivity]]. + apply valid_trace_add_default_last in Htr. + match goal with + |- finite_valid_trace_from Limited ?is ?tr => + cut + (finite_valid_trace_from Limited is tr /\ + (state_annotation (@finite_trace_last _ (type Limited) is tr) ⊆ equivocators)) + end + ; [intuition|]. + induction Htr using finite_valid_trace_init_to_rev_strong_ind. + - split. + + constructor. + apply initial_state_is_valid. + split; [assumption|reflexivity]. + + apply list_subseteq_nil. + - setoid_rewrite annotate_trace_from_app. + cbn. + rewrite finite_trace_last_is_last. + cbn. + split. + + apply finite_valid_trace_from_app_iff. + split; [apply IHHtr1|]. + apply finite_valid_trace_singleton. + repeat split. + * apply finite_valid_trace_last_pstate. + apply IHHtr1. + * clear -Heqiom IHHtr2. + apply proj1 in IHHtr2. + unfold empty_initial_message_or_final_output in Heqiom. + destruct_list_last iom_tr iom_tr' iom_item Heqiom_tr + ; [apply option_initial_message_is_valid; assumption|]. + destruct iom as [im|]; [|apply option_valid_message_None]. + eapply valid_trace_output_is_valid; [eassumption|]. + setoid_rewrite annotate_trace_from_app. + apply Exists_app. + right. + destruct iom_item. + apply Exists_exists. + eexists. + split; [left|]. + assumption. + * cbn. + rewrite annotate_trace_from_last_original_state. + destruct l as (i, li). + cbn. + replace (finite_trace_last _ _) with s + by (apply valid_trace_get_last in Htr1; congruence). + apply Ht. + * apply Rle_trans with (sum_weights (remove_dups equivocators)) + ; [|assumption]. + apply sum_weights_subseteq. + -- destruct iom as [im|]; cycle 1. + ++ eapply coeqv_limited_equivocation_state_annotation_nodup. + apply finite_valid_trace_last_pstate. + apply IHHtr1. + ++ apply set_union_nodup_left. + eapply coeqv_limited_equivocation_state_annotation_nodup. + apply finite_valid_trace_last_pstate. + apply IHHtr1. + -- apply NoDup_remove_dups. + -- transitivity equivocators. + ++ eapply fixed_transition_preserves_annotation_equivocators + ; [eassumption|eassumption|apply IHHtr1]. + ++ intro. apply elem_of_remove_dups. + * cbn. + unfold annotated_transition. + cbn. + rewrite !annotate_trace_from_last_original_state. + destruct l as (i, li). + cbn. + replace (finite_trace_last _ _) with s + by (apply valid_trace_get_last in Htr1; congruence). + apply proj2 in Ht. + cbn in Ht. + destruct (vtransition _ _ _) as (si', om'). + inversion Ht. + reflexivity. + + eapply fixed_transition_preserves_annotation_equivocators + ; [eassumption|eassumption|apply IHHtr1]. +Qed. + +Lemma annotated_limited_incl_constrained_limited + {is_equivocating_tracewise_no_has_been_sent_dec : RelDecision (is_equivocating_tracewise_no_has_been_sent IM (fun i => i) sender)} + : VLSM_full_projection + Limited + (limited_equivocation_vlsm_composition IM (listing_from_finite index) sender) + Datatypes.id original_state. +Proof. + constructor. + intros sX trX HtrX. + eapply traces_exhibiting_limited_equivocation_are_valid + ; [apply Hsender_safety|]. + apply msg_dep_fixed_limited_equivocation. + assumption. +Qed. + +End sec_msg_dep_fixed_limited_equivocation. From b2022327574993676685f7fd6930b76b7a4f527e Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Fri, 21 Jan 2022 09:41:29 +0200 Subject: [PATCH 10/18] Equivalence with state equivocation model --- .../LimitedEquivocation/LimitedEquivocation.v | 95 +++++++++++++++---- .../LimitedEquivocationSimulation.v | 72 +++++++++++--- 2 files changed, 132 insertions(+), 35 deletions(-) diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v index e97d2f34..6a883f7c 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v @@ -1,9 +1,9 @@ -From stdpp Require Import prelude. +From stdpp Require Import prelude finite. From Coq Require Import FinFun Lia Reals Lra. -From VLSM.Lib Require Import Preamble ListExtras StdppListSet ListSetExtras FinExtras Measurable. -From VLSM Require Import Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces. +From VLSM.Lib Require Import Preamble ListExtras StdppListSet ListSetExtras FinExtras FinFunExtras Measurable. +From VLSM Require Import Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.AnnotatedVLSM. From VLSM Require Import Core.Equivocation Core.Equivocation.TraceWiseEquivocation. -From VLSM Require Import Core.Equivocation.NoEquivocation Core.Equivocation.LimitedEquivocation. +From VLSM Require Import Core.Equivocation.NoEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.MsgDepLimitedEquivocation. From VLSM Require Import Core.Equivocators.Common Core.Equivocators.Projections. From VLSM Require Import Core.Equivocators.MessageProperties Core.Equivocators.Composition.Common. From VLSM Require Import Core.Equivocators.Composition.Projections Core.MessageDependencies. @@ -48,13 +48,13 @@ Section limited_state_equivocation. Context {message : Type} {index : Type} - {IndEqDec : EqDecision index} + `{finite.Finite index} (IM : index -> VLSM message) (Hbs : forall i : index, HasBeenSentCapability (IM i)) (Hbr : forall i : index, HasBeenReceivedCapability (IM i)) (Free := free_composite_vlsm IM) - {index_listing : list index} - (finite_index : Listing index_listing) + (index_listing : list index := enum index) + (finite_index : Listing index_listing := listing_from_finite index) (equivocator_descriptors := equivocator_descriptors IM) (equivocators_state_project := equivocators_state_project IM) (equivocator_IM := equivocator_IM IM) @@ -107,8 +107,11 @@ Qed. Lemma preloaded_equivocators_limited_equivocations_vlsm_incl_free : VLSM_incl (pre_loaded_with_all_messages_vlsm equivocators_limited_equivocations_vlsm) PreFreeE. Proof. - apply basic_VLSM_incl_preloaded; intro; intros; [assumption| |assumption]. - split; [|exact I]. apply H. + apply basic_VLSM_incl_preloaded. + 1,3: intro; intros; assumption. + intros l s om Hv. + split; [|exact I]. + apply Hv. Qed. (** @@ -156,20 +159,20 @@ Lemma equivocators_limited_valid_trace_is_fixed is s tr (equivocators_fixed_equivocations_vlsm IM Hbs index_listing (equivocating_validators s)) is s tr. Proof. - intro H. - split; [| apply H]. + intro Htr. + split; [| apply Htr]. cut (forall equivocating, equivocating_validators s ⊆ equivocating -> 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_valid_trace_init_to_rev_ind; intros equivocating Hincl. + induction Htr 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_valid_trace_init_to equivocating). - spec IHfinite_valid_trace_init_to. + specialize (IHHtr equivocating). + spec IHHtr. { apply proj2 in Ht. specialize (equivocators_transition_preserves_equivocating_indices IM index_listing _ _ _ _ _ Ht) as Hincl'. @@ -185,7 +188,7 @@ Proof. with s; [assumption|]. 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. + apply valid_trace_last_pstate in IHHtr. destruct Ht as [[_ [_ [Hv [[Hno_equiv _] Hno_heavy]]]] Ht]. repeat split; [assumption| |assumption|assumption| |assumption]. + destruct iom as [m|]; [|apply option_valid_message_None]. @@ -245,6 +248,52 @@ Proof. * intros i Hi. apply elem_of_remove_dups. assumption. Qed. +Section sec_equivocators_projection_annotated_limited. + +Context + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (Hchannel : channel_authentication_prop IM Datatypes.id sender) + . + +(** Projections of valid traces for the composition of equivocators +with limited state-equivocation and no message-equivocation can be +annotated with equivocators to obtain a limited-message equivocation trace. +*) +Lemma equivocators_limited_valid_trace_projects_to_annotated_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_valid_trace equivocators_limited_equivocations_vlsm is tr) + : exists + (trX : list (composite_transition_item IM)) + (initial_descriptors : equivocator_descriptors) + (isX := equivocators_state_project initial_descriptors is) + (final_stateX := finite_trace_last isX trX), + 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_valid_trace (msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) + {| original_state := isX; state_annotation := ` inhabitant |} + (msg_dep_annotate_trace_with_equivocators IM Hbs Hbr full_message_dependencies sender isX trX). +Proof. + eapply equivocators_limited_valid_trace_projects_to_fixed_limited_equivocation in Htr + as [trX [initial_descriptors [Hinitial_descriptors [Hpr [Hlst_pr Hpr_limited]]]]] + ; [|eassumption]. + exists trX, initial_descriptors. + repeat (split; [assumption|]). + eapply msg_dep_limited_fixed_equivocation; eassumption. +Qed. + +End sec_equivocators_projection_annotated_limited. + +Section sec_equivocators_projection_constrained_limited. + Context {is_equivocating_tracewise_no_has_been_sent_dec : RelDecision (is_equivocating_tracewise_no_has_been_sent IM (fun i => i) sender)} (Limited : VLSM message := limited_equivocation_vlsm_composition IM finite_index sender) @@ -321,12 +370,14 @@ above strengthens to a [VLSM_projection]. Lemma limited_equivocators_vlsm_projection : VLSM_projection equivocators_limited_equivocations_vlsm Limited (equivocators_total_label_project IM) (equivocators_total_state_project IM). Proof. - constructor; [constructor|]; intros. - - apply PreFreeE_Free_vlsm_projection_type. - revert H. apply VLSM_incl_finite_valid_trace_from. + constructor; [constructor|]. + - intros sX trX HtrX. + apply PreFreeE_Free_vlsm_projection_type. + revert HtrX. apply VLSM_incl_finite_valid_trace_from. apply equivocators_limited_equivocations_vlsm_incl_preloaded_free. - - assert (Hpre_tr : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). - { revert H. apply VLSM_incl_finite_valid_trace. + - intros sX trX HtrX. + assert (Hpre_tr : finite_valid_trace (pre_loaded_with_all_messages_vlsm FreeE) sX trX). + { revert HtrX. apply VLSM_incl_finite_valid_trace. apply equivocators_limited_equivocations_vlsm_incl_preloaded_free. } specialize @@ -338,11 +389,13 @@ Proof. rewrite (equivocators_total_trace_project_characterization IM (proj1 Hpre_tr)). reflexivity. } - apply Hsim in H. + apply Hsim in HtrX. remember (pre_VLSM_projection_trace_project _ _ _ _ _) as tr. replace tr with (equivocators_total_trace_project IM trX); [assumption|]. subst. symmetry. apply (equivocators_total_VLSM_projection_trace_project IM (proj1 Hpre_tr)). Qed. +End sec_equivocators_projection_constrained_limited. + End limited_state_equivocation. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v index 3dcfe774..59ec8bb6 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocationSimulation.v @@ -1,9 +1,9 @@ -From stdpp Require Import prelude. +From stdpp Require Import prelude finite. From Coq Require Import FinFun Reals. -From VLSM Require Import Lib.StdppListSet. -From VLSM Require Import Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.SubProjectionTraces. +From VLSM Require Import Lib.StdppListSet Lib.FinFunExtras. +From VLSM Require Import Core.VLSM Core.VLSMProjections Core.Composition Core.ProjectionTraces Core.SubProjectionTraces Core.AnnotatedVLSM. From VLSM Require Import Core.Equivocation Core.EquivocationProjections Core.Equivocation.FixedSetEquivocation Core.Equivocation.NoEquivocation. -From VLSM Require Import Lib.Measurable Core.Equivocation.TraceWiseEquivocation Core.Equivocation.LimitedEquivocation. +From VLSM Require Import Lib.Measurable Core.Equivocation.TraceWiseEquivocation Core.Equivocation.LimitedEquivocation Core.Equivocation.MsgDepLimitedEquivocation. From VLSM Require Import MessageDependencies Core.Equivocation.WitnessedEquivocation. From VLSM Require Import Core.Equivocators.Composition.Common Core.Equivocators.Composition.Projections. From VLSM Require Import Core.Equivocators.Composition.LimitedEquivocation.LimitedEquivocation. @@ -26,16 +26,15 @@ Section fixed_limited_state_equivocation. Context {message : Type} - {index : Type} - {IndEqDec : EqDecision index} + `{finite.Finite index} (IM : index -> VLSM message) {i0 : Inhabited index} (Hbs : forall i, HasBeenSentCapability (IM i)) (Hbr : forall i, HasBeenReceivedCapability (IM i)) `{IndThreshold : ReachableThreshold index} - {index_listing : list index} - (finite_index : Listing index_listing) - (Limited : VLSM message := equivocators_limited_equivocations_vlsm IM Hbs finite_index) + (index_listing : list index := enum index) + (finite_index : Listing index_listing := listing_from_finite index) + (Limited : VLSM message := equivocators_limited_equivocations_vlsm IM Hbs) (equivocating : list index) (Fixed : VLSM message := equivocators_fixed_equivocations_vlsm IM Hbs index_listing equivocating) . @@ -74,16 +73,15 @@ Section limited_equivocation_simulation. Context {message : Type} - {index : Type} - {IndEqDec : EqDecision index} + `{finite.Finite index} (IM : index -> VLSM message) {i0 : Inhabited index} (Hbs : forall i, HasBeenSentCapability (IM i)) (Hbr : forall i, HasBeenReceivedCapability (IM i)) `{IndThreshold : ReachableThreshold index} - {index_listing : list index} - (finite_index : Listing index_listing) - (XE : VLSM message := equivocators_limited_equivocations_vlsm IM Hbs finite_index) + (index_listing : list index := enum index) + (finite_index : Listing index_listing := listing_from_finite index) + (XE : VLSM message := equivocators_limited_equivocations_vlsm IM Hbs) . (** If a trace has the [fixed_limited_equivocation_prop]erty, then it can be @@ -120,6 +118,52 @@ Proof. assumption. Qed. +Section sec_equivocators_simulating_annotated_limited. + +Context + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (sender : message -> option index) + (Hchannel : channel_authentication_prop IM Datatypes.id sender) + . + +Lemma equivocators_limited_valid_trace_projects_to_annotated_limited_equivocation_rev + isX sX trX + (HtrX : finite_valid_trace_init_to (msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) isX sX trX) + : exists is, equivocators_total_state_project IM is = original_state isX /\ + exists s, equivocators_total_state_project IM s = original_state sX /\ + exists tr, equivocators_total_trace_project IM tr = + pre_VLSM_full_projection_finite_trace_project + (annotated_type (free_composite_vlsm IM) (set index)) (composite_type IM) Datatypes.id original_state + trX + /\ finite_valid_trace_init_to XE is s tr + /\ finite_trace_last_output trX = finite_trace_last_output tr. +Proof. + apply valid_trace_get_last in HtrX as HeqsX. + apply valid_trace_forget_last in HtrX. + eapply msg_dep_fixed_limited_equivocation in HtrX. + 2-5: eassumption. + apply limited_equivocators_finite_valid_trace_init_to_rev in HtrX + as [is [His_pr [s [Hpr_s [tr [Htr_pr [Htr Houtput]]]]]]] + ; [|assumption]. + eexists; split; [eassumption|]. + subst. + exists s. + rewrite Hpr_s. + erewrite <- pre_VLSM_full_projection_finite_trace_last. + split; [reflexivity|]. + exists tr. + split; [assumption|]. + split; [assumption|]. + rewrite <- Houtput. + apply pre_VLSM_full_projection_finite_trace_last_output. +Qed. + +End sec_equivocators_simulating_annotated_limited. + Context (sender : message -> option index) {is_equivocating_tracewise_no_has_been_sent_dec : RelDecision (is_equivocating_tracewise_no_has_been_sent IM (fun i => i) sender)} From 2a2776e24d2b84de8ba185f283f71b97cd779441 Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Tue, 25 Jan 2022 17:32:08 +0200 Subject: [PATCH 11/18] Byzantine traces result for message-dependency-based limited message-equivocation --- theories/VLSM/Core/AnnotatedVLSM.v | 227 ++++++++++++++- .../ByzantineTraces/LimitedByzantineTraces.v | 271 +++++++++++++++++- .../Equivocation/MsgDepLimitedEquivocation.v | 47 ++- 3 files changed, 536 insertions(+), 9 deletions(-) diff --git a/theories/VLSM/Core/AnnotatedVLSM.v b/theories/VLSM/Core/AnnotatedVLSM.v index 1632dd2c..592145de 100644 --- a/theories/VLSM/Core/AnnotatedVLSM.v +++ b/theories/VLSM/Core/AnnotatedVLSM.v @@ -1,6 +1,14 @@ From stdpp Require Import prelude. From VLSM.Lib Require Import ListExtras. -From VLSM.Core Require Import VLSM VLSMProjections. +From VLSM.Core Require Import VLSM VLSMProjections Validator Composition ProjectionTraces. + +(** * State-annotated VLSMs + +This module describes the VLSM obtained by augmenting the states of an existing +VLSM with annotations and providing additional validity constraints taking into +account the annotations, and a function for updating the annotations following a +transition. +*) Section sec_annotated_vlsm. @@ -193,3 +201,220 @@ Proof. Qed. End sec_annotated_vlsm_projections. + +Section sec_composite_annotated_vlsm_projections. + +(** ** Specializing [projection_validator_prop]erties to annotated compositions. *) + +Context + {message : Type} + `{EqDecision index} + (IM : index -> VLSM message) + (Free := free_composite_vlsm IM) + {annotation : Type} + (initial_annotation_prop : annotation -> Prop) + `{Inhabited (sig initial_annotation_prop)} + (annotated_constraint : @label _ (annotated_type Free annotation) -> annotated_state Free annotation * option message -> Prop) + (annotated_transition_state : @label _ (annotated_type Free annotation) -> annotated_state Free annotation * option message -> annotation) + (AnnotatedFree : VLSM message := annotated_vlsm Free annotation initial_annotation_prop annotated_constraint annotated_transition_state) + (i : index) + . + +Definition annotated_composite_label_project : vlabel AnnotatedFree -> option (vlabel (IM i)) + := composite_project_label IM i. + +Definition annotated_composite_state_project : vstate AnnotatedFree -> vstate (IM i) + := fun s => original_state s i. + +Definition annotated_projection_validator_prop : Prop := + @projection_validator_prop _ AnnotatedFree (IM i) + annotated_composite_label_project annotated_composite_state_project. + +Definition annotated_composite_label_lift : vlabel (IM i) -> vlabel AnnotatedFree + := lift_to_composite_label IM i. + +Definition annotated_composite_state_lift : vstate (IM i) -> vstate AnnotatedFree + := fun si => + @Build_annotated_state _ (free_composite_vlsm IM) _ + (lift_to_composite_state IM i si) (` inhabitant). + +Definition annotated_projection_validator_prop_alt : Prop := + @projection_validator_prop_alt _ AnnotatedFree (IM i) + annotated_composite_label_project annotated_composite_state_project + annotated_composite_label_lift annotated_composite_state_lift. + +Lemma annotated_composite_preloaded_projection + : VLSM_projection AnnotatedFree (pre_loaded_with_all_messages_vlsm (IM i)) + annotated_composite_label_project annotated_composite_state_project. +Proof. + apply basic_VLSM_projection. + - intros (_i, _li) li. + unfold annotated_composite_label_project, composite_project_label. + cbn. + case_decide; [|congruence]. + subst _i. + cbn. + intro Heq_li. + inversion Heq_li. + subst _li. + clear Heq_li. + intros (s, ann) om (_ & _ & [Hv _] & _) _ _. + assumption. + - intros (_i, _li) li. + unfold annotated_composite_label_project, composite_project_label. + cbn. + case_decide; [|congruence]. + subst _i. + cbn. + intro Heq_li. + inversion Heq_li. + subst _li. + clear Heq_li. + intros (s, ann) iom (s', ann') oom. + unfold input_valid_transition. + cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (si', om'). + intros [_ Ht]. + inversion Ht. + f_equal. + symmetry. + apply state_update_eq. + - intros (j, lj). + unfold annotated_composite_label_project, composite_project_label. + cbn. + case_decide as Hij; [congruence|]. + intros _. + intros (s, ann) iom (s', ann') oom. + unfold input_valid_transition. + cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (si', om'). + intros [_ Ht]. + inversion Ht. + apply state_update_neq. + congruence. + - intros (s, ann). cbn. + intros [Hs _]. + apply Hs. + - intro; intros. + apply any_message_is_valid_in_preloaded. +Qed. + +Definition annotated_composite_induced_projection : VLSM message + := projection_induced_vlsm AnnotatedFree (type (IM i)) + annotated_composite_label_project annotated_composite_state_project + annotated_composite_label_lift annotated_composite_state_lift. + +Lemma annotated_composite_induced_projection_transition_None + : weak_projection_transition_consistency_None _ _ annotated_composite_label_project annotated_composite_state_project. +Proof. + intros (j, lj). + unfold annotated_composite_label_project, composite_project_label. + cbn. + case_decide as Hij; [congruence|]. + intros _ sX omX s'X om'X [_ Ht]. + revert Ht. + cbn. + unfold annotated_transition. + cbn. + destruct (vtransition _ _ _) as (si', om'). + inversion 1. + subst om' s'X. + clear Ht. + cbn. + rewrite state_update_neq by congruence. + reflexivity. +Qed. + +Lemma annotated_composite_induced_projection_label_lift + : induced_projection_label_lift_prop _ _ + annotated_composite_label_project annotated_composite_label_lift. +Proof. + apply component_label_projection_lift with (constraint := free_constraint IM). +Qed. + +Lemma annotated_composite_induced_projection_state_lift + : induced_projection_state_lift_prop _ _ + annotated_composite_state_project annotated_composite_state_lift. +Proof. + intros si. + apply state_update_eq. +Qed. + +Lemma annotated_composite_induced_projection_initial_lift + : strong_full_projection_initial_state_preservation (IM i) AnnotatedFree + annotated_composite_state_lift. +Proof. + split; cbn. + - apply lift_to_composite_state_initial; assumption. + - destruct inhabitant; assumption. +Qed. + +Lemma annotated_composite_induced_projection_transition_consistency + : induced_projection_transition_consistency_Some _ _ + annotated_composite_label_project annotated_composite_state_project. +Proof. + intros (i1, li1) (i2, li2) li. + unfold annotated_composite_label_project, composite_project_label. + cbn. + case_decide as Hi1; [|congruence]. + subst i1. + cbn. + intro Hli1. + inversion Hli1. + subst li1. + clear Hli1. + case_decide as Hi2; [|congruence]. + subst i2. + cbn. + intro Hli2. + inversion Hli2. + subst li2. + clear Hli2. + intros (s1, ann1) (s2, ann2). + unfold annotated_transition. + cbn. + intros <- iom sX1' oom1. + destruct (vtransition _ _ _) as (si', om'). + intro Heq. + inversion Heq. + subst. + clear Heq. + intros sX2' oom2. + intro Heq. + inversion Heq. + subst. + clear Heq. + split; [|reflexivity]. + cbn. + rewrite !state_update_eq. + reflexivity. +Qed. + +Definition annotated_composite_induced_projection_transition_Some + := basic_weak_projection_transition_consistency_Some _ _ _ _ _ _ + annotated_composite_induced_projection_label_lift + annotated_composite_induced_projection_state_lift + annotated_composite_induced_projection_transition_consistency. + +Definition annotated_composite_induced_projection_is_projection + := projection_induced_vlsm_is_projection _ _ _ _ _ _ + annotated_composite_induced_projection_transition_None + annotated_composite_induced_projection_transition_Some. + +Lemma annotated_projection_validator_prop_alt_iff + : annotated_projection_validator_prop_alt <-> annotated_projection_validator_prop. +Proof. + apply projection_validator_prop_alt_iff. + - apply annotated_composite_preloaded_projection. + - apply annotated_composite_induced_projection_transition_None. + - apply annotated_composite_induced_projection_label_lift. + - apply annotated_composite_induced_projection_state_lift. + - apply annotated_composite_induced_projection_initial_lift. + - apply annotated_composite_induced_projection_transition_consistency. +Qed. + +End sec_composite_annotated_vlsm_projections. diff --git a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v index 25e1fd8a..d860c50b 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -1,9 +1,9 @@ From stdpp Require Import prelude finite. 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.Lib Require Import Preamble StdppListSet Measurable FinFunExtras StdppExtras ListSetExtras. +From VLSM.Core Require Import VLSM MessageDependencies VLSMProjections Composition ProjectionTraces SubProjectionTraces AnnotatedVLSM. From VLSM Require Import Core.ByzantineTraces Core.ByzantineTraces.FixedSetByzantineTraces. -From VLSM Require Import Core.Validator 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.MsgDepLimitedEquivocation Core.Equivocation.TraceWiseEquivocation. (** * VLSM Compositions with byzantine nodes of limited weight @@ -251,3 +251,268 @@ Proof. Qed. End limited_byzantine_traces. + +Section sec_msg_dep_limited_byzantince_traces. + +Context + {message : Type} + `{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)) + (message_dependencies : message -> set message) + (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + `{ReachableThreshold index} + (sender : message -> option index) + (Limited := msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (Hchannel : channel_authentication_prop IM Datatypes.id sender) + (Hsender_safety : sender_safety_alt_prop IM Datatypes.id sender := + channel_authentication_sender_safety _ _ _ Hchannel) + (Hvalidator: + forall i : index, + msg_dep_limited_equivocation_projection_validator_prop IM Hbs Hbr + full_message_dependencies sender i) + (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + . + +Lemma pre_loaded_fixed_non_byzantine_traces_are_limited + (s: composite_state IM) + (tr: list (composite_transition_item IM)) + (byzantine: set index) + (Hlimited: (sum_weights (remove_dups byzantine) <= `threshold)%R) + (Hbyzantine: + finite_valid_trace + (pre_loaded_fixed_non_byzantine_vlsm IM Hbs byzantine Datatypes.id sender) + (composite_state_sub_projection IM (set_diff (enum index) byzantine) s) + (finite_trace_sub_projection IM (set_diff (enum index) byzantine) tr)) + (s_reset_byzantine := + lift_sub_state IM (set_diff (enum index) byzantine) + (composite_state_sub_projection IM (set_diff (enum index) byzantine) s)) + (bs := Build_annotated_state (free_composite_vlsm IM) (set index) s_reset_byzantine (` inhabitant)) + (btr := + msg_dep_annotate_trace_with_equivocators IM Hbs Hbr full_message_dependencies sender + s_reset_byzantine + (pre_VLSM_full_projection_finite_trace_project _ _ + (lift_sub_label IM (set_diff (enum index) byzantine)) (lift_sub_state IM (set_diff (enum index) byzantine)) + (finite_trace_sub_projection IM (set_diff (enum index) byzantine) tr))) + : finite_valid_trace Limited bs btr /\ + state_annotation (@finite_trace_last _ (type Limited) bs btr) ⊆ byzantine. +Proof. + induction Hbyzantine using finite_valid_trace_rev_ind; [repeat split|]. + - constructor. + apply initial_state_is_valid. + repeat split. + cbn. + apply lift_sub_state_initial. + assumption. + - cbn; apply lift_sub_state_initial; assumption. + - apply list_subseteq_nil. + - subst s_reset_byzantine bs btr. + unfold pre_VLSM_full_projection_finite_trace_project. + rewrite !map_app. + setoid_rewrite annotate_trace_from_app. + cbn. + rewrite finite_trace_last_is_last. + cbn. + destruct l as (sub_i, li). + destruct_dec_sig sub_i i Hi Heqsub_i. + subst sub_i. + destruct IHHbyzantine as [[Htr0_ann Hsi_ann] Htr0_eqv_byzantine]. + apply finite_valid_trace_last_pstate in Htr0_ann as Hlst. + cbn in Htr0_eqv_byzantine, Hlst |- *. + remember + (@finite_trace_last _(annotated_type (free_composite_vlsm IM) (set index)) + _ _) as lst. + assert (Hlsti : original_state lst = lift_sub_state IM (set_diff (enum index) byzantine) (finite_trace_last si tr0)). + { + subst lst. + rewrite annotate_trace_from_last_original_state. + symmetry. + apply + (pre_VLSM_full_projection_finite_trace_last _ _ + (lift_sub_label IM (set_diff (enum index) byzantine)) + (lift_sub_state IM (set_diff (enum index) byzantine))). + } + assert + (Hvi : input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li + (original_state lst i, iom)). + { + specialize (PreSubFree_PreFree_weak_full_projection IM (set_diff (enum index) byzantine) (` inhabitant)) + as Hproj. + spec Hproj. + { clear. apply initial_state_is_valid. destruct inhabitant. assumption. } + specialize (VLSM_weak_full_projection_input_valid Hproj) as Hiv. + apply proj1 in Hx. + eapply VLSM_incl_input_valid in Hx; [|apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages]. + apply Hiv in Hx. + eapply + (VLSM_projection_input_valid (preloaded_component_projection IM i) (existT i li) li) in Hx + ; [|rewrite composite_project_label_eq; reflexivity]. + rewrite Hlsti. + assumption. + } + match goal with + |- _ /\ ?B => cut B + end. + { + intro Heqv_byzantine. + split; [|assumption]. + split; [|assumption]. + apply finite_valid_trace_from_app_iff. + split; [assumption|]. + subst x. + cbn. + apply finite_valid_trace_singleton. + replace (finite_trace_last _ _) with lst. + repeat split. + - assumption. + - apply Hvalidator in Hvi as (_ & _ & _ & _ & _ & Hiom & _). + assumption. + - cbn. + rewrite Hlsti. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + apply Hx. + - apply Rle_trans with (sum_weights (remove_dups byzantine)) + ; [|assumption]. + apply sum_weights_subseteq. + + cut (NoDup (state_annotation lst)). + { intro Hnodup. + destruct iom as [im|]; [|assumption]. + apply set_union_nodup_left; assumption. + } + subst. + eapply coeqv_limited_equivocation_state_annotation_nodup; eassumption. + + apply NoDup_remove_dups. + + intro eqv. + rewrite elem_of_remove_dups. + apply Heqv_byzantine. + - clear -Hx Heqlst Hlsti. + cbn. + unfold annotated_transition. + cbn. + apply proj2 in Hx. + revert Hx. + cbn. + rewrite Hlsti. + unfold lift_sub_state at 1. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + unfold sub_IM at 2. + cbn. + destruct (vtransition _ _ _) as (si', om'). + intro Ht. + inversion Ht. + subst sf om'. + clear Ht. + f_equal. + f_equal. + extensionality j. + destruct (decide (i = j)) as [Hij | Hij]. + + subst j. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + rewrite !state_update_eq. + reflexivity. + + rewrite state_update_neq by congruence. + destruct (decide (j ∈ set_diff (enum index) byzantine)) as [Hj|Hj]. + * unfold lift_sub_state. + rewrite !lift_sub_state_to_eq with (Hi0 := Hj). + rewrite sub_IM_state_update_neq by congruence. + reflexivity. + * unfold lift_sub_state. + rewrite !lift_sub_state_to_neq by assumption. + reflexivity. + } + destruct iom as [im|]; [|assumption]. + apply set_union_subseteq_iff. + split; [assumption|]. + unfold coeqv_message_equivocators. + case_decide as Hnobs; [apply list_subseteq_nil|]. + erewrite full_node_msg_dep_coequivocating_senders + with (i0 := i) (li0 :=li). + 2-5: eassumption. + rewrite app_nil_r. + cbn. + destruct (sender im) as [i_im|] eqn:Hsender + ; [|apply list_subseteq_nil]. + intro _i_im. rewrite elem_of_list_singleton. + intro; subst _i_im. + apply proj1, proj2, proj2, proj2, proj1 in Hx + as [Hsent | [Hsigned _]]. + + elim Hnobs. + destruct Hsent as [sub_i_im Hsent]. + cbn in Hsent. + cbn. + destruct_dec_sig sub_i_im _i_im H_i_im Heqsub_i_im. + subst sub_i_im. + apply composite_has_been_observed_sent_received_iff. + left. + exists _i_im. + rewrite Hlsti. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := H_i_im). + assumption. + + clear -Hsender Hsigned. + destruct Hsigned as (_i_im & H_i_im & Hauth). + unfold channel_authenticated_message in Hauth. + rewrite Hsender in Hauth. + apply Some_inj in Hauth. + subst _i_im. + destruct (decide (i_im ∈ byzantine)) as [Hi_im|Hni_im] + ; [assumption|]. + elim H_i_im. + apply set_diff_intro; [apply elem_of_enum|assumption]. +Qed. + +Lemma msg_dep_validator_limited_non_byzantine_traces_are_limited_non_equivocating s tr + : limited_byzantine_trace_prop IM Hbs sender s tr <-> + exists bs btr selection (selection_complement := set_diff (enum index) selection), + finite_valid_trace Limited bs btr /\ + state_annotation (finite_trace_last bs btr) ⊆ selection /\ + (sum_weights (remove_dups selection) <= `threshold)%R /\ + composite_state_sub_projection IM selection_complement s = + composite_state_sub_projection IM selection_complement (original_state bs) /\ + finite_trace_sub_projection IM selection_complement tr = + finite_trace_sub_projection IM selection_complement + (pre_VLSM_full_projection_finite_trace_project + (type Limited) (composite_type IM) Datatypes.id original_state btr). +Proof. + split. + - intros (byzantine & Hlimited & Hbyzantine). + apply pre_loaded_fixed_non_byzantine_traces_are_limited in Hbyzantine + as [Hbtr Heqv_byzantine] + ; [|assumption]. + eexists _,_, byzantine; repeat (split; [eassumption|]). + split. + + extensionality sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + cbn. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + reflexivity. + + subst Limited. + rewrite msg_dep_annotate_trace_with_equivocators_project. + symmetry. + apply composite_trace_sub_projection_lift. + - intros (bs & btr & byzantine & Hbtr & Heqv_byzantine & Hlimited & His_pr & Htr_pr). + exists byzantine. + split; [assumption|]. + unfold fixed_byzantine_trace_alt_prop. + eapply VLSM_incl_finite_valid_trace + ; [apply fixed_non_equivocating_incl_fixed_non_byzantine; assumption|]. + apply fixed_non_equivocating_traces_char. + symmetry in His_pr, Htr_pr. + eexists _,_; split; [|split; eassumption]. + eapply msg_dep_fixed_limited_equivocation_witnessed, proj2 in Hbtr. + 2-5: eassumption. + revert Hbtr. + apply VLSM_incl_finite_valid_trace. + apply fixed_equivocation_vlsm_composition_index_incl; assumption. +Qed. + +End sec_msg_dep_limited_byzantince_traces. \ No newline at end of file diff --git a/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v index d8a738a9..272e3c51 100644 --- a/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v @@ -106,6 +106,14 @@ Proof. destruct om as [m|]; assumption. Qed. +Definition coeqv_limited_equivocation_projection_validator_prop : index -> Prop := + annotated_projection_validator_prop IM (fun s => s = empty_set) + coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. + +Definition coeqv_limited_equivocation_projection_validator_prop_alt : index -> Prop := + annotated_projection_validator_prop_alt IM (fun s => s = empty_set) + coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. + End sec_coequivocating_senders_limited_equivocation. Section sec_msg_dep_limited_equivocation. @@ -144,6 +152,20 @@ Definition msg_dep_annotate_trace_with_equivocators := Definition msg_dep_composite_transition_message_equivocators := coeqv_composite_transition_message_equivocators IM Hbs Hbr sender msg_dep_coequivocating_senders. +Definition msg_dep_limited_equivocation_projection_validator_prop := + coeqv_limited_equivocation_projection_validator_prop IM Hbs Hbr sender msg_dep_coequivocating_senders. + +Definition msg_dep_limited_equivocation_projection_validator_prop_alt := + coeqv_limited_equivocation_projection_validator_prop_alt IM Hbs Hbr sender msg_dep_coequivocating_senders. + +Lemma msg_dep_annotate_trace_with_equivocators_project s tr + : pre_VLSM_full_projection_finite_trace_project (type msg_dep_limited_equivocation_vlsm) + (composite_type IM) Datatypes.id original_state + (msg_dep_annotate_trace_with_equivocators s tr) = tr. +Proof. + apply (annotate_trace_project (free_composite_vlsm IM) (set validator)). +Qed. + End sec_msg_dep_limited_equivocation. Section sec_full_node_limited_equivocation. @@ -465,16 +487,17 @@ Unshelve. apply set_union_subseteq_right. Qed. -Lemma msg_dep_fixed_limited_equivocation is tr - : finite_valid_trace Limited is tr -> - fixed_limited_equivocation_prop IM Hbs Hbr +Lemma msg_dep_fixed_limited_equivocation_witnessed is tr + (Htr : finite_valid_trace Limited is tr) + (equivocators := state_annotation (finite_trace_last is tr)) + (Fixed := fixed_equivocation_vlsm_composition IM Hbs Hbr equivocators) + : (sum_weights (remove_dups equivocators) <= `threshold)%R /\ + finite_valid_trace Fixed (original_state is) (pre_VLSM_full_projection_finite_trace_project (type Limited) (composite_type IM) Datatypes.id original_state tr). Proof. - intro Htr. - exists (state_annotation (finite_trace_last is tr)). split. - rewrite set_eq_nodup_sum_weight_eq with (lv2 := state_annotation (finite_trace_last is tr)). @@ -590,6 +613,20 @@ Proof. assumption. Qed. +Corollary msg_dep_fixed_limited_equivocation is tr + : finite_valid_trace Limited is tr -> + fixed_limited_equivocation_prop IM Hbs Hbr + (original_state is) + (pre_VLSM_full_projection_finite_trace_project + (type Limited) (composite_type IM) Datatypes.id original_state + tr). +Proof. + intro Htr. + exists (state_annotation (finite_trace_last is tr)). + apply msg_dep_fixed_limited_equivocation_witnessed. + assumption. +Qed. + Lemma fixed_transition_preserves_annotation_equivocators equivocators (is : vstate (free_composite_vlsm IM)) s tr From f4d7121be7b6fddda5d1b07c727c7cc5c4f9113f Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Thu, 27 Jan 2022 11:35:34 +0200 Subject: [PATCH 12/18] Relaxing validator assumption to message-validation --- theories/VLSM/Core/AnnotatedVLSM.v | 3 ++ theories/VLSM/Core/ByzantineTraces.v | 19 +++----- .../ByzantineTraces/FixedSetByzantineTraces.v | 8 ++-- .../ByzantineTraces/LimitedByzantineTraces.v | 12 +++-- .../Equivocation/MsgDepLimitedEquivocation.v | 7 +++ theories/VLSM/Core/Validator.v | 45 +++++++++++++++++++ 6 files changed, 69 insertions(+), 25 deletions(-) diff --git a/theories/VLSM/Core/AnnotatedVLSM.v b/theories/VLSM/Core/AnnotatedVLSM.v index 592145de..f0e9aef7 100644 --- a/theories/VLSM/Core/AnnotatedVLSM.v +++ b/theories/VLSM/Core/AnnotatedVLSM.v @@ -230,6 +230,9 @@ Definition annotated_projection_validator_prop : Prop := @projection_validator_prop _ AnnotatedFree (IM i) annotated_composite_label_project annotated_composite_state_project. +Definition annotated_message_validator_prop : Prop := + @message_validator_prop _ AnnotatedFree (IM i). + Definition annotated_composite_label_lift : vlabel (IM i) -> vlabel AnnotatedFree := lift_to_composite_label IM i. diff --git a/theories/VLSM/Core/ByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces.v index f36fbc96..4eca1a18 100644 --- a/theories/VLSM/Core/ByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces.v @@ -399,7 +399,7 @@ Section composite_validator_byzantine_traces. (X := composite_vlsm IM constraint) (PreLoadedX := pre_loaded_with_all_messages_vlsm X) (FreeX := free_composite_vlsm IM) - (Hvalidator: forall i : index, component_projection_validator_prop IM constraint i) + (Hvalidator: forall i : index, component_message_validator_prop IM constraint i) . (** @@ -438,18 +438,11 @@ included in <> to prove our main result. destruct l as (i, li). simpl in *. specialize (valid_state_project_preloaded_to_preloaded _ IM constraint lst i Hlst) as Hlsti. - 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_valid_initial_state_message (IM i)). - } - apply Hvalidator in Hiv. - clear -Hiv. - destruct Hiv as [_ [_ [_ [Hinput _]]]]. - destruct Hinput as [s Hinput]. - exists s. - revert Hinput. - apply (constraint_subsumption_valid_state_message_preservation IM constraint). - intro. intros. destruct som. apply H. + destruct iom as [im|]; [|apply option_valid_message_None]. + eapply Hvalidator. + split; [eassumption|]. split; [|eassumption]. + eexists _. + apply (pre_loaded_with_all_messages_message_valid_initial_state_message (IM i)). Qed. (** diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index f4839b56..a787c19a 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -550,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 [projection_validator_prop]erty, which basically +abiding nodes satisfy the [projection_message_validator_prop]erty, which basically allows them to locally verify the authenticity of a received message. *) @@ -821,7 +821,7 @@ End assuming_initial_messages_lift. Context (Hvalidator: forall i : index, i ∉ selection -> - component_projection_validator_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) + component_message_validator_prop IM (fixed_equivocation_constraint IM Hbs Hbr selection) i) . Lemma validator_fixed_non_byzantine_vlsm_lift_initial_message @@ -839,9 +839,7 @@ Proof. by exists i, (exist _ m Him). - destruct Hseeded as [Hsigned [i [Hi [li [si Hpre_valid]]]]]. apply set_diff_elim2 in Hi. - specialize (Hvalidator i Hi _ _ Hpre_valid) - as [sX [Heqsi [HsX [Hm HvX]]]]. - assumption. + eapply Hvalidator; eassumption. Qed. (** The VLSM corresponding to the induced projection from a fixed-set byzantine diff --git a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v index d860c50b..c128d690 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -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) - (Hvalidator: forall i : index, component_projection_validator_prop IM limited_constraint i) + (Hvalidator: forall i : index, component_message_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)) @@ -192,9 +192,7 @@ Proof. exists i, (exist _ m Him); intuition. + destruct Hseeded as [Hsigned [i [Hi [li [si Hpre_valid]]]]]. apply set_diff_elim2 in Hi. - specialize (Hvalidator i _ _ Hpre_valid) - as [_ [_ [_ [Hmsg _]]]]. - assumption. + eapply Hvalidator; eassumption. Qed. End fixed_limited_selection. @@ -274,7 +272,7 @@ Context channel_authentication_sender_safety _ _ _ Hchannel) (Hvalidator: forall i : index, - msg_dep_limited_equivocation_projection_validator_prop IM Hbs Hbr + msg_dep_limited_equivocation_message_validator_prop IM Hbs Hbr full_message_dependencies sender i) (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) . @@ -370,8 +368,8 @@ Proof. replace (finite_trace_last _ _) with lst. repeat split. - assumption. - - apply Hvalidator in Hvi as (_ & _ & _ & _ & _ & Hiom & _). - assumption. + - destruct iom as [im|]; [|apply option_valid_message_None]. + eapply Hvalidator; eassumption. - cbn. rewrite Hlsti. unfold lift_sub_state. diff --git a/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v index 272e3c51..6dde878d 100644 --- a/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocation/MsgDepLimitedEquivocation.v @@ -110,6 +110,10 @@ Definition coeqv_limited_equivocation_projection_validator_prop : index -> Prop annotated_projection_validator_prop IM (fun s => s = empty_set) coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. +Definition coeqv_limited_equivocation_message_validator_prop : index -> Prop := + annotated_message_validator_prop IM (fun s => s = empty_set) + coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. + Definition coeqv_limited_equivocation_projection_validator_prop_alt : index -> Prop := annotated_projection_validator_prop_alt IM (fun s => s = empty_set) coeqv_limited_equivocation_constraint coeqv_composite_transition_message_equivocators. @@ -155,6 +159,9 @@ Definition msg_dep_composite_transition_message_equivocators := Definition msg_dep_limited_equivocation_projection_validator_prop := coeqv_limited_equivocation_projection_validator_prop IM Hbs Hbr sender msg_dep_coequivocating_senders. +Definition msg_dep_limited_equivocation_message_validator_prop := + coeqv_limited_equivocation_message_validator_prop IM Hbs Hbr sender msg_dep_coequivocating_senders. + Definition msg_dep_limited_equivocation_projection_validator_prop_alt := coeqv_limited_equivocation_projection_validator_prop_alt IM Hbs Hbr sender msg_dep_coequivocating_senders. diff --git a/theories/VLSM/Core/Validator.v b/theories/VLSM/Core/Validator.v index ea84318d..2748e7a1 100644 --- a/theories/VLSM/Core/Validator.v +++ b/theories/VLSM/Core/Validator.v @@ -43,6 +43,23 @@ Definition projection_validator_prop := input_valid PreY li (si,omi) -> projection_induced_valid X (type Y) label_project state_project li (si,omi). +(** A message validator can check within a component whether the message +is valid for the composition. +*) +Definition message_validator_prop := + forall li si im, + input_valid PreY li (si, Some im) -> + valid_message_prop X im. + +(** The [projection_validator_prop]erty is stronger. *) +Lemma projection_validator_is_message_validator + : projection_validator_prop -> message_validator_prop. +Proof. + intros Hvalidator li si im Hvi. + apply Hvalidator in Hvi as (_ & _ & _ & _ & _ & Him & _). + assumption. +Qed. + (** It is easy to see that the [projection_validator_prop]erty includes the [projection_validator_received_messages_prop]erty. @@ -335,6 +352,34 @@ Proof. apply constraint_preloaded_free_incl with (constraint0 := constraint). Qed. +Lemma component_projection_validator_prop_is_induced + : component_projection_validator_prop <-> + @projection_validator_prop _ X (IM i) (composite_project_label IM i) (fun s => s i). +Proof. + split. + - intros Hvalidator li si omi Hvi. + apply (VLSM_eq_input_valid (composite_vlsm_constrained_projection_is_induced IM constraint i)). + apply projection_valid_input_valid. + apply Hvalidator. + assumption. + - intros Hvalidator li (si, omi) Hvi. + apply (VLSM_eq_input_valid (composite_vlsm_constrained_projection_is_induced IM constraint i)). + revert Hvi. + apply VLSM_incl_input_valid. + apply pre_loaded_with_all_messages_validator_proj_incl. + + apply component_projection_to_preloaded. + + apply component_transition_projection_None. + + apply component_label_projection_lift. + + apply component_state_projection_lift. + + intros isi. + apply (lift_to_composite_state_initial IM). + + apply component_transition_projection_Some. + + assumption. +Qed. + +Definition component_message_validator_prop : Prop := + @message_validator_prop _ X (IM i). + (** Assuming the [component_projection_validator_prop]erty, the component [pre_loaded_with_all_messages_vlsm] is [VLSM_eq]ual (trace-equivalent) with its corresponding [projection_induced_vlsm]. From 6b5ea467e8e9ce860ef98853288f2b9d5cbe114b Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Thu, 27 Jan 2022 15:02:06 +0200 Subject: [PATCH 13/18] Refactoring and documentation --- .../ByzantineTraces/LimitedByzantineTraces.v | 216 ++++++++++-------- theories/VLSM/Core/SubProjectionTraces.v | 26 +++ 2 files changed, 149 insertions(+), 93 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v index c128d690..d52a7e33 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -250,7 +250,7 @@ Qed. End limited_byzantine_traces. -Section sec_msg_dep_limited_byzantince_traces. +Section sec_msg_dep_limited_byzantine_traces. Context {message : Type} @@ -277,29 +277,124 @@ Context (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) . -Lemma pre_loaded_fixed_non_byzantine_traces_are_limited +(** +If the set of byzantine nodes is weight-limited and if an [input_valid_transition] +of the non-byzantine nodes from a state of weight-limited equivocation does not +introduce equivocators from the non-byzantine nodes, then the transition is valid +for weight-limited equivocation. +*) +Lemma lift_pre_loaded_fixed_non_byzantine_valid_transition_to_limited + (byzantine: set index) + (non_byzantine := set_diff (enum index) byzantine) + (Hlimited: (sum_weights (remove_dups byzantine) <= `threshold)%R) + sub_l sub_s iom sub_sf oom + (Ht_sub : input_valid_transition + (pre_loaded_fixed_non_byzantine_vlsm IM Hbs byzantine Datatypes.id sender) + sub_l (sub_s, iom) (sub_sf, oom)) + ann_s + (Hann_s : valid_state_prop Limited ann_s) + (Hann_s_pr : original_state ann_s = lift_sub_state IM non_byzantine sub_s) + (ann' := msg_dep_composite_transition_message_equivocators IM Hbs Hbr full_message_dependencies sender + (lift_sub_label IM non_byzantine sub_l) (ann_s, iom)) + (Heqv_byzantine : ann' ⊆ byzantine) + : input_valid_transition Limited + (lift_sub_label IM non_byzantine sub_l) (ann_s, iom) + (Build_annotated_state (free_composite_vlsm IM) (set index) + (lift_sub_state IM non_byzantine sub_sf) ann', + oom). +Proof. + destruct sub_l as (sub_i, li). + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + repeat split. + - assumption. + - destruct iom as [im|]; [|apply option_valid_message_None]. + eapply Hvalidator, + pre_loaded_sub_composite_input_valid_projection, + Ht_sub. + - cbn. + rewrite Hann_s_pr. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + apply Ht_sub. + - apply Rle_trans with (sum_weights (remove_dups byzantine)) + ; [|assumption]. + apply sum_weights_subseteq. + + cut (NoDup (state_annotation ann_s)). + { intro Hnodup. + destruct iom as [im|]; [|assumption]. + apply set_union_nodup_left; assumption. + } + subst. + eapply coeqv_limited_equivocation_state_annotation_nodup; eassumption. + + apply NoDup_remove_dups. + + intro eqv. + rewrite elem_of_remove_dups. + apply Heqv_byzantine. + - clear -Ht_sub Hann_s_pr. + cbn. + unfold annotated_transition. + cbn. + apply proj2 in Ht_sub. + revert Ht_sub. + cbn. + rewrite Hann_s_pr. + unfold lift_sub_state at 1. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + unfold sub_IM at 2. + cbn. + destruct (vtransition _ _ _) as (si', om'). + intro Ht. + inversion Ht. + subst. + clear Ht. + f_equal. + f_equal. + extensionality j. + destruct (decide (i = j)) as [Hij | Hij]. + + subst j. + unfold lift_sub_state. + rewrite lift_sub_state_to_eq with (Hi0 := Hi). + rewrite !state_update_eq. + reflexivity. + + rewrite state_update_neq by congruence. + destruct (decide (j ∈ set_diff (enum index) byzantine)) as [Hj|Hj]. + * unfold lift_sub_state. + rewrite !lift_sub_state_to_eq with (Hi0 := Hj). + rewrite sub_IM_state_update_neq by congruence. + reflexivity. + * unfold lift_sub_state. + rewrite !lift_sub_state_to_neq by assumption. + reflexivity. +Qed. + +(** Considering a trace with the [fixed_byzantine_trace_alt_prop]erty for a +set <> of indices of bounded weight, its subtrace corresponding to +the non-byzantine nodes is of limited equivocation and its set of equivocators +is included in <>. +*) +Lemma lift_fixed_byzantine_traces_to_limited (s: composite_state IM) (tr: list (composite_transition_item IM)) (byzantine: set index) + (non_byzantine := set_diff (enum index) byzantine) (Hlimited: (sum_weights (remove_dups byzantine) <= `threshold)%R) (Hbyzantine: - finite_valid_trace - (pre_loaded_fixed_non_byzantine_vlsm IM Hbs byzantine Datatypes.id sender) - (composite_state_sub_projection IM (set_diff (enum index) byzantine) s) - (finite_trace_sub_projection IM (set_diff (enum index) byzantine) tr)) + fixed_byzantine_trace_alt_prop IM Hbs byzantine Datatypes.id sender s tr) (s_reset_byzantine := - lift_sub_state IM (set_diff (enum index) byzantine) - (composite_state_sub_projection IM (set_diff (enum index) byzantine) s)) + lift_sub_state IM non_byzantine + (composite_state_sub_projection IM non_byzantine s)) (bs := Build_annotated_state (free_composite_vlsm IM) (set index) s_reset_byzantine (` inhabitant)) (btr := msg_dep_annotate_trace_with_equivocators IM Hbs Hbr full_message_dependencies sender s_reset_byzantine (pre_VLSM_full_projection_finite_trace_project _ _ - (lift_sub_label IM (set_diff (enum index) byzantine)) (lift_sub_state IM (set_diff (enum index) byzantine)) - (finite_trace_sub_projection IM (set_diff (enum index) byzantine) tr))) + (lift_sub_label IM non_byzantine) (lift_sub_state IM (set_diff (enum index) byzantine)) + (finite_trace_sub_projection IM non_byzantine tr))) : finite_valid_trace Limited bs btr /\ state_annotation (@finite_trace_last _ (type Limited) bs btr) ⊆ byzantine. Proof. + subst non_byzantine. induction Hbyzantine using finite_valid_trace_rev_ind; [repeat split|]. - constructor. apply initial_state_is_valid. @@ -320,8 +415,7 @@ Proof. destruct_dec_sig sub_i i Hi Heqsub_i. subst sub_i. destruct IHHbyzantine as [[Htr0_ann Hsi_ann] Htr0_eqv_byzantine]. - apply finite_valid_trace_last_pstate in Htr0_ann as Hlst. - cbn in Htr0_eqv_byzantine, Hlst |- *. + cbn in Htr0_eqv_byzantine |- *. remember (@finite_trace_last _(annotated_type (free_composite_vlsm IM) (set index)) _ _) as lst. @@ -335,24 +429,6 @@ Proof. (lift_sub_label IM (set_diff (enum index) byzantine)) (lift_sub_state IM (set_diff (enum index) byzantine))). } - assert - (Hvi : input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li - (original_state lst i, iom)). - { - specialize (PreSubFree_PreFree_weak_full_projection IM (set_diff (enum index) byzantine) (` inhabitant)) - as Hproj. - spec Hproj. - { clear. apply initial_state_is_valid. destruct inhabitant. assumption. } - specialize (VLSM_weak_full_projection_input_valid Hproj) as Hiv. - apply proj1 in Hx. - eapply VLSM_incl_input_valid in Hx; [|apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages]. - apply Hiv in Hx. - eapply - (VLSM_projection_input_valid (preloaded_component_projection IM i) (existT i li) li) in Hx - ; [|rewrite composite_project_label_eq; reflexivity]. - rewrite Hlsti. - assumption. - } match goal with |- _ /\ ?B => cut B end. @@ -366,73 +442,21 @@ Proof. cbn. apply finite_valid_trace_singleton. replace (finite_trace_last _ _) with lst. - repeat split. - - assumption. - - destruct iom as [im|]; [|apply option_valid_message_None]. - eapply Hvalidator; eassumption. - - cbn. - rewrite Hlsti. - unfold lift_sub_state. - rewrite lift_sub_state_to_eq with (Hi0 := Hi). - apply Hx. - - apply Rle_trans with (sum_weights (remove_dups byzantine)) - ; [|assumption]. - apply sum_weights_subseteq. - + cut (NoDup (state_annotation lst)). - { intro Hnodup. - destruct iom as [im|]; [|assumption]. - apply set_union_nodup_left; assumption. - } - subst. - eapply coeqv_limited_equivocation_state_annotation_nodup; eassumption. - + apply NoDup_remove_dups. - + intro eqv. - rewrite elem_of_remove_dups. - apply Heqv_byzantine. - - clear -Hx Heqlst Hlsti. - cbn. - unfold annotated_transition. - cbn. - apply proj2 in Hx. - revert Hx. - cbn. - rewrite Hlsti. - unfold lift_sub_state at 1. - rewrite lift_sub_state_to_eq with (Hi0 := Hi). - unfold sub_IM at 2. - cbn. - destruct (vtransition _ _ _) as (si', om'). - intro Ht. - inversion Ht. - subst sf om'. - clear Ht. - f_equal. - f_equal. - extensionality j. - destruct (decide (i = j)) as [Hij | Hij]. - + subst j. - unfold lift_sub_state. - rewrite lift_sub_state_to_eq with (Hi0 := Hi). - rewrite !state_update_eq. - reflexivity. - + rewrite state_update_neq by congruence. - destruct (decide (j ∈ set_diff (enum index) byzantine)) as [Hj|Hj]. - * unfold lift_sub_state. - rewrite !lift_sub_state_to_eq with (Hi0 := Hj). - rewrite sub_IM_state_update_neq by congruence. - reflexivity. - * unfold lift_sub_state. - rewrite !lift_sub_state_to_neq by assumption. - reflexivity. + eapply lift_pre_loaded_fixed_non_byzantine_valid_transition_to_limited. + 1,2,4,5: eassumption. + subst lst. + apply finite_valid_trace_last_pstate. + assumption. } destruct iom as [im|]; [|assumption]. apply set_union_subseteq_iff. split; [assumption|]. unfold coeqv_message_equivocators. case_decide as Hnobs; [apply list_subseteq_nil|]. - erewrite full_node_msg_dep_coequivocating_senders - with (i0 := i) (li0 :=li). - 2-5: eassumption. + erewrite full_node_msg_dep_coequivocating_senders with (i0 := i) (li0 :=li). + 2-4: eassumption. + 2: cbn; rewrite Hlsti; + eapply pre_loaded_sub_composite_input_valid_projection, Hx. rewrite app_nil_r. cbn. destruct (sender im) as [i_im|] eqn:Hsender @@ -466,6 +490,12 @@ Proof. apply set_diff_intro; [apply elem_of_enum|assumption]. Qed. +(** +Under full-message dependencies and full node assumptions, if all components are +validators for the [msg_dep_limited_equivocation_vlsm] associated to their +composition, then the traces exposed limited Byzantine behaviour coincide with +the traces exposed to limited equivocation. +*) Lemma msg_dep_validator_limited_non_byzantine_traces_are_limited_non_equivocating s tr : limited_byzantine_trace_prop IM Hbs sender s tr <-> exists bs btr selection (selection_complement := set_diff (enum index) selection), @@ -481,7 +511,7 @@ Lemma msg_dep_validator_limited_non_byzantine_traces_are_limited_non_equivocatin Proof. split. - intros (byzantine & Hlimited & Hbyzantine). - apply pre_loaded_fixed_non_byzantine_traces_are_limited in Hbyzantine + apply lift_fixed_byzantine_traces_to_limited in Hbyzantine as [Hbtr Heqv_byzantine] ; [|assumption]. eexists _,_, byzantine; repeat (split; [eassumption|]). @@ -513,4 +543,4 @@ Proof. apply fixed_equivocation_vlsm_composition_index_incl; assumption. Qed. -End sec_msg_dep_limited_byzantince_traces. \ No newline at end of file +End sec_msg_dep_limited_byzantine_traces. \ No newline at end of file diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index b3ed4b97..b589c801 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -2080,6 +2080,32 @@ Proof. assumption. Qed. +(** Deriving reachable-validity for the component from the input validity +w.r.t. a sub_composition preloaded with messages. +*) +Lemma pre_loaded_sub_composite_input_valid_projection constraint Q + i Hi li sub_s im + : input_valid + (pre_loaded_vlsm (composite_vlsm (sub_IM IM indices) constraint) Q) + (existT (dexist i Hi) li) (sub_s, Some im) -> + input_valid (pre_loaded_with_all_messages_vlsm (IM i)) + li (lift_sub_state IM indices sub_s i, Some im). +Proof. + intro Ht_sub. + eapply + (VLSM_projection_input_valid (preloaded_component_projection IM i) (existT i li) li) + ; [rewrite composite_project_label_eq; reflexivity|]. + cut ( + input_valid + (pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM indices))) + (existT (dexist i Hi) li) (sub_s, Some im)) + ; [eapply + (VLSM_full_projection_input_valid lift_sub_preloaded_free_full_projection)|]. + eapply VLSM_incl_input_valid + ; [apply composite_pre_loaded_vlsm_incl_pre_loaded_with_all_messages|]. + assumption. +Qed. + Lemma can_emit_sub_projection {validator : Type} (A : validator -> index) From 69c618a488c7c981907d548cd0b6231dd54abbf6 Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Thu, 27 Jan 2022 16:03:19 +0200 Subject: [PATCH 14/18] progress --- .../ByzantineTraces/FixedSetByzantineTraces.v | 313 +++++++++++------- .../Equivocation/MsgDepFixedSetEquivocation.v | 6 +- theories/VLSM/Core/EquivocationProjections.v | 4 +- .../LimitedEquivocation/FixedEquivocation.v | 2 +- .../LimitedEquivocation/LimitedEquivocation.v | 2 +- .../Equivocators/Composition/Projections.v | 12 +- theories/VLSM/Core/ProjectionTraces.v | 2 +- theories/VLSM/Core/SubProjectionTraces.v | 58 +++- theories/VLSM/Core/VLSMProjections.v | 164 ++++++--- theories/VLSM/Lib/ListExtras.v | 41 ++- 10 files changed, 404 insertions(+), 200 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index a787c19a..a2115bab 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -60,11 +60,11 @@ Definition emit_any_signed_message_vlsm End emit_any_signed_message_vlsm. -Section fixed_byzantine_traces. +Section sec_fixed_byzantine_traces. Context {message : Type} - `{EqDecision index} + `{finite.Finite index} (IM : index -> VLSM message) (Hbs : forall i : index, HasBeenSentCapability (IM i)) (byzantine : set index) @@ -75,6 +75,15 @@ Context (can_emit_signed : channel_authentication_prop IM A sender) (Hsender_safety : sender_safety_alt_prop IM A sender := channel_authentication_sender_safety IM A sender can_emit_signed) + (non_byzantine : set index := set_diff (enum index) byzantine) + . + +Section sec_fixed_byzantine_replacements. + +Context + (byzantine_replacements_IM : sub_index byzantine -> VLSM message) + (no_initial_messages_in_byzantine_IM : no_initial_messages_in_IM_prop byzantine_replacements_IM) + (can_emit_signed_byzantine : channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender)) . (** Given a collection of nodes indexed by an <>, we define the @@ -82,16 +91,17 @@ associated fixed byzantine collection of nodes by replacing the nodes corresponding to the indices in a given <> with byzantine nodes, i.e., nodes which can emit any (signed) message. *) -Definition fixed_byzantine_IM : index -> VLSM message := - update_IM IM byzantine (fun i => emit_any_signed_message_vlsm A sender (` i)). +Definition fixed_byzantine_IM + : index -> VLSM message := + update_IM IM byzantine byzantine_replacements_IM. Lemma fixed_byzantine_IM_no_initial_messages : forall i m, ~vinitial_message_prop (fixed_byzantine_IM i) m. Proof. unfold fixed_byzantine_IM, update_IM. simpl. intros i m Hm. - case_decide. - - cbn in Hm. assumption. + case_decide as Hi. + - elim (no_initial_messages_in_byzantine_IM (dexist i Hi) m). assumption. - elim (no_initial_messages_in_IM i m). assumption. Qed. @@ -101,11 +111,15 @@ Proof. unfold fixed_byzantine_IM, update_IM. simpl. intros i m Hm. case_decide. - - destruct Hm as [(s0, om) [l [s1 [[_ [_ Hv]] Ht]]]]. - cbn in Hv, Ht. - unfold signed_messages_valid, channel_authenticated_message in Hv. - inversion Ht. - subst; assumption. + - apply can_emit_signed_byzantine in Hm. revert Hm. + unfold channel_authenticated_message, sub_IM_A, sub_IM_sender. + destruct (sender m) as [v|]; [|inversion 1]. + case_decide as HAv; [|inversion 1]. + cbn. + intro Heq. + apply Some_inj, dec_sig_eq_iff in Heq. + simpl in Heq. + congruence. - apply can_emit_signed; assumption. Qed. @@ -114,24 +128,6 @@ Definition fixed_byzantine_IM_sender_safety channel_authentication_sender_safety fixed_byzantine_IM A sender fixed_byzantine_IM_preserves_channel_authentication. -Definition message_as_byzantine_label - (m : message) - (i : index) - (Hi : i ∈ byzantine) - : composite_label fixed_byzantine_IM. -Proof. - exists i. - unfold fixed_byzantine_IM, update_IM. - simpl. - rewrite decide_True by assumption. - cbn. exact m. -Defined. - -Context - {finite_index : finite.Finite index} - (non_byzantine : set index := set_diff (enum index) byzantine) - . - Definition sub_fixed_byzantine_IM_Hbs : forall sub_i : sub_index non_byzantine, HasBeenSentCapability (sub_IM fixed_byzantine_IM non_byzantine sub_i) @@ -143,25 +139,142 @@ Definition non_byzantine_not_equivocating_constraint : composite_label fixed_byzantine_IM -> composite_state fixed_byzantine_IM * option message -> Prop := sub_IM_not_equivocating_constraint fixed_byzantine_IM non_byzantine A sender sub_fixed_byzantine_IM_Hbs. -(** The first definition of the [fixed_byzantine_trace_prop]erty: +Definition fixed_byantine_composite_vlsm : VLSM message := + composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint. -Fixed byzantine traces are projections to the subset of protocol-following 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). +(** +Given that we're only looking at the composition of nodes in the +[non_byzantine], we can define their subset as either a subset of the +[fixed_byzantine_IM] or of the original <>. + +As it turns out, the first definition is easier to prove equivalent to the +induced [fixed_non_byzantine_projection], while the second is easier to work +with in general because it makes no reference to byzantine nodes. + +Therefore we'll first be defining them both below and show that they are +equivalent to each-other using the generic Lemma [same_IM_full_projection]. *) -Definition fixed_byzantine_trace_prop - (is : composite_state (sub_IM fixed_byzantine_IM non_byzantine)) - (tr : list (composite_transition_item (sub_IM fixed_byzantine_IM non_byzantine))) - : Prop := - exists bis btr, - finite_valid_trace (composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint) bis btr /\ - composite_state_sub_projection fixed_byzantine_IM non_byzantine bis = is /\ - finite_trace_sub_projection fixed_byzantine_IM non_byzantine btr = tr. + +Definition pre_loaded_fixed_non_byzantine_vlsm' : VLSM message := + pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM non_byzantine)). + +Lemma pre_loaded_fixed_non_byzantine_vlsm'_projection + : VLSM_projection fixed_byantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm' + (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine) + (composite_state_sub_projection fixed_byzantine_IM non_byzantine). +Proof. + constructor; [apply induced_sub_projection_is_projection|]. + intros sX trX HtrX. + apply + (VLSM_projection_finite_valid_trace + (induced_sub_projection_is_projection fixed_byzantine_IM non_byzantine _)) + in HtrX. + revert HtrX. + apply VLSM_incl_finite_valid_trace. + apply induced_sub_projection_preloaded_free_incl. +Qed. + +Definition pre_loaded_fixed_non_byzantine_vlsm : VLSM message := + pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM non_byzantine)). + +Lemma non_byzantine_nodes_same + : forall sub_i, sub_IM fixed_byzantine_IM non_byzantine sub_i = sub_IM IM non_byzantine sub_i. +Proof. + intro sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + unfold sub_IM, fixed_byzantine_IM, update_IM. + simpl. + apply set_diff_elim2 in Hi. + case_decide; [contradiction|]. + reflexivity. +Qed. + +Lemma non_byzantine_nodes_same_sym + : forall sub_i, sub_IM IM non_byzantine sub_i = sub_IM fixed_byzantine_IM non_byzantine sub_i. +Proof. + intro. symmetry. apply non_byzantine_nodes_same. +Qed. + +Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection + : VLSM_full_projection + pre_loaded_fixed_non_byzantine_vlsm' + pre_loaded_fixed_non_byzantine_vlsm + (same_IM_label_rew non_byzantine_nodes_same) + (same_IM_state_rew non_byzantine_nodes_same). +Proof. + apply same_IM_preloaded_free_full_projection. +Qed. + +Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection' + : VLSM_full_projection + pre_loaded_fixed_non_byzantine_vlsm + pre_loaded_fixed_non_byzantine_vlsm' + (same_IM_label_rew non_byzantine_nodes_same_sym) + (same_IM_state_rew non_byzantine_nodes_same_sym). +Proof. + apply same_IM_preloaded_free_full_projection. +Qed. + +Definition pre_loaded_fixed_non_byzantine_label_projection := + (option_map (same_IM_label_rew non_byzantine_nodes_same)) ∘ + (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine). + +Definition pre_loaded_fixed_non_byzantine_state_projection := + (same_IM_state_rew non_byzantine_nodes_same) ∘ + (composite_state_sub_projection fixed_byzantine_IM non_byzantine). + +Lemma pre_loaded_fixed_non_byzantine_vlsm_projection + : VLSM_projection fixed_byantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm + pre_loaded_fixed_non_byzantine_label_projection + pre_loaded_fixed_non_byzantine_state_projection. +Proof. + specialize + (pre_definitions_projection_relation_left_trace + (composite_type (sub_IM fixed_byzantine_IM non_byzantine)) + (composite_type (sub_IM IM non_byzantine)) + (same_IM_label_rew non_byzantine_nodes_same) + (same_IM_state_rew non_byzantine_nodes_same) + (composite_type fixed_byzantine_IM) + (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine) + (composite_state_sub_projection fixed_byzantine_IM non_byzantine)) + as Hrew. + constructor; [constructor|]; intros sX trX HtrX + ; unfold pre_loaded_fixed_non_byzantine_label_projection, pre_loaded_fixed_non_byzantine_state_projection + ; cbn in Hrew |- * ; rewrite Hrew. + - induction trX using rev_ind; [reflexivity|]. + rewrite finite_trace_last_is_last. + cbn. + rewrite pre_VLSM_projection_finite_trace_project_app. + setoid_rewrite map_app. + rewrite !map_app. + destruct x; destruct l as (i, li); cbn. + unfold pre_VLSM_projection_transition_item_project, pre_loaded_fixed_non_byzantine_label_projection. + cbn. + unfold composite_label_sub_projection_option at 2. + cbn. + case_decide as Hi; cbn; [rewrite last_is_last; reflexivity|]. + apply finite_valid_trace_from_app_iff in HtrX as [HtrX Hx]. + spec IHtrX HtrX. + inversion Hx; subst; clear Hx Htl. + eapply induced_sub_projection_transition_consistency_None + with (sub_index_list := non_byzantine) in Ht. + + rewrite Ht. + cbn in IHtrX. + setoid_rewrite IHtrX. + rewrite app_nil_r. + reflexivity. + + unfold composite_label_sub_projection_option; cbn. + case_decide; [contradiction|reflexivity]. + - cbn. + apply (VLSM_full_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm_full_projection). + revert HtrX. + apply (VLSM_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm'_projection). +Qed. (** ** Byzantine traces characterization as projections. *) -Section fixed_byzantine_traces_as_projections. +Section sec_fixed_non_byzantine_projection. Definition fixed_non_byzantine_projection : VLSM message := induced_sub_projection fixed_byzantine_IM non_byzantine non_byzantine_not_equivocating_constraint. @@ -185,13 +298,14 @@ Qed. Lemma fixed_non_byzantine_projection_incl_preloaded : VLSM_incl fixed_non_byzantine_projection (pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM non_byzantine))). Proof. - apply basic_VLSM_strong_incl; intro; intros. - - apply fixed_non_byzantine_projection_initial_state_preservation. assumption. - - exact I. - - split; [|exact I]. - revert H. + apply basic_VLSM_strong_incl. + - intros is. apply fixed_non_byzantine_projection_initial_state_preservation. + - intros m Hm. exact I. + - intros l s om Hv. + split; [|exact I]. + revert Hv. apply induced_sub_projection_valid_preservation. - - revert H. apply induced_sub_projection_transition_preservation. + - intros l s om s' om'. apply induced_sub_projection_transition_preservation. Qed. (** The induced projection from the composition of [fixed_byzantine_IM] under @@ -215,6 +329,30 @@ Proof. exact (equal_f_dep Heq (dexist (A v) HAv)). Qed. +End sec_fixed_non_byzantine_projection. + +End sec_fixed_byzantine_replacements. + +(** 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 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). +*) +Definition fixed_byzantine_trace_prop + (is : composite_state (sub_IM IM non_byzantine)) + (tr : list (composite_transition_item (sub_IM IM non_byzantine))) + : Prop := + exists + (byzantine_replacements_IM : sub_index byzantine -> VLSM message) + (no_initial_messages_in_byzantine_IM : no_initial_messages_in_IM_prop byzantine_replacements_IM) + (can_emit_signed_byzantine : channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender)), + exists bis btr, + finite_valid_trace (fixed_byantine_composite_vlsm byzantine_replacements_IM) bis btr /\ + pre_loaded_fixed_non_byzantine_state_projection byzantine_replacements_IM bis = is /\ + VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection byzantine_replacements_IM) btr = tr. + (** Characterization result for the first definition: the [fixed_byzantine_trace_prop]erty is equivalent to the [finite_valid_trace_prop]erty of the [induced_sub_projection] of the @@ -223,7 +361,7 @@ 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_valid_trace fixed_non_byzantine_projection is tr. + finite_valid_trace (fixed_non_byzantine_projection byzantine_replacements_IM) is tr. Proof. apply and_comm. apply (projection_friendly_trace_char (induced_sub_projection_is_projection _ _ _)). @@ -245,81 +383,6 @@ 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, 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 -[non_byzantine], we can define their subset as either a subset of the -[fixed_byzantine_IM] or of the original <>. - -As it turns out, the first definition is easier to prove equivalent to the -induced [fixed_non_byzantine_projection], while the second is easier to work -with in general because it makes no reference to byzantine nodes. - -Therefore we'll first be defining them both below and show that they are -equivalent to each-other using the generic Lemma [same_IM_full_projection]. -*) - -Definition pre_loaded_fixed_non_byzantine_vlsm' : VLSM message := - composite_no_equivocation_vlsm_with_pre_loaded (sub_IM fixed_byzantine_IM non_byzantine) (free_constraint _) sub_fixed_byzantine_IM_Hbs fixed_set_signed_message. - -Definition pre_loaded_fixed_non_byzantine_vlsm : VLSM message := - composite_no_equivocation_vlsm_with_pre_loaded (sub_IM IM non_byzantine) (free_constraint _) (sub_has_been_sent_capabilities IM non_byzantine Hbs) fixed_set_signed_message. - -Lemma non_byzantine_nodes_same - : forall sub_i, sub_IM fixed_byzantine_IM non_byzantine sub_i = sub_IM IM non_byzantine sub_i. -Proof. - intro sub_i. - destruct_dec_sig sub_i i Hi Heqsub_i. - subst. - unfold sub_IM, fixed_byzantine_IM, update_IM. - simpl. - apply set_diff_elim2 in Hi. - rewrite decide_False by assumption. - reflexivity. -Qed. - -Lemma non_byzantine_nodes_same_sym - : forall sub_i, sub_IM IM non_byzantine sub_i = sub_IM fixed_byzantine_IM non_byzantine sub_i. -Proof. - intro. symmetry. apply non_byzantine_nodes_same. -Qed. - -Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection - : VLSM_full_projection - pre_loaded_fixed_non_byzantine_vlsm' - pre_loaded_fixed_non_byzantine_vlsm - (same_IM_label_rew non_byzantine_nodes_same) - (same_IM_state_rew non_byzantine_nodes_same). -Proof. - apply same_IM_full_projection. - intros s1 H l1 om H0. - apply proj1 in H0. - split; [|exact I]. - destruct om as [m|]; [|exact I]. - destruct H0 as [Hsent | Hseeded]; [|right; assumption]. - cbn. left. - revert Hsent. - apply same_IM_composite_has_been_sent_preservation; [|assumption]. - apply stdpp_finite_sub_index; assumption. -Qed. - -Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection' - : VLSM_full_projection - pre_loaded_fixed_non_byzantine_vlsm - pre_loaded_fixed_non_byzantine_vlsm' - (same_IM_label_rew non_byzantine_nodes_same_sym) - (same_IM_state_rew non_byzantine_nodes_same_sym). -Proof. - apply same_IM_full_projection. - intros s1 H l1 om H0. - apply proj1 in H0. - split; [|exact I]. - destruct om as [m|]; [|exact I]. - destruct H0 as [Hsent | Hseeded]; [|right; assumption]. - cbn. left. - revert Hsent. - apply same_IM_composite_has_been_sent_preservation; [|assumption]. - apply stdpp_finite_sub_index; assumption. -Qed. (** We next show that the [fixed_non_byzantine_projection] and diff --git a/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v b/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v index a294468a..4a95e73a 100644 --- a/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v +++ b/theories/VLSM/Core/Equivocation/MsgDepFixedSetEquivocation.v @@ -231,7 +231,7 @@ Proof. specialize (Hsent _ _ Hpre_tr). apply Exists_exists in Hsent as [item [Hitem Hout]]. apply elem_of_list_In in Hitem. - apply pre_VLSM_projection_trace_project_in_iff in Hitem as HitemX. + apply pre_VLSM_projection_finite_trace_project_in_iff in Hitem as HitemX. destruct HitemX as [itemX [HitemX HitemX_pr]]. destruct itemX. unfold pre_VLSM_projection_transition_item_project, composite_project_label in HitemX_pr. @@ -268,7 +268,7 @@ Proof. eapply in_futures_preserving_oracle_from_stepwise. * apply has_been_sent_stepwise_from_trace. * eexists; eassumption. - + apply pre_VLSM_projection_trace_project_app_rev in Htr_pr + + apply pre_VLSM_projection_finite_trace_project_app_rev in Htr_pr as [preX [sufX [Heqtr [HpreX_pr HsufX_pr]]]]. apply valid_trace_last_pstate in Hpre_tr as Hdestinationi. apply proper_received in Hreceived; [|assumption]. @@ -276,7 +276,7 @@ Proof. apply Exists_exists in Hreceived as [item_dm [Hitem_dm Hreceived]]. apply elem_of_list_In in Hitem_dm. rewrite <- HpreX_pr in Hitem_dm. - apply pre_VLSM_projection_trace_project_in_iff in Hitem_dm as HitemX_dm. + apply pre_VLSM_projection_finite_trace_project_in_iff in Hitem_dm as HitemX_dm. destruct HitemX_dm as [itemX_dm [HitemX_dm HitemX_dm_pr]]. subst tr. apply proj1 in Htr. diff --git a/theories/VLSM/Core/EquivocationProjections.v b/theories/VLSM/Core/EquivocationProjections.v index c2b0b2ad..ed7cae78 100644 --- a/theories/VLSM/Core/EquivocationProjections.v +++ b/theories/VLSM/Core/EquivocationProjections.v @@ -53,7 +53,7 @@ Proof. specialize (Hm _ _ HtrX). apply Exists_exists in Hm as [itemY [HitemY Hm]]. apply elem_of_list_In in HitemY. - apply pre_VLSM_projection_trace_project_in_iff in HitemY + apply pre_VLSM_projection_finite_trace_project_in_iff in HitemY as [itemX [HitemX Hpr]]. apply Exists_exists. apply elem_of_list_In in HitemX. @@ -156,7 +156,7 @@ Proof. subst tr. setoid_rewrite map_app. apply List.Exists_app. right. simpl. left. - remember (pre_VLSM_full_projection_trace_item_project _ _ _ _ _) as itemY. + remember (pre_VLSM_full_projection_transition_item_project _ _ _ _ _) as itemY. specialize (Hselector item itemY). subst. destruct item. apply (Hselector eq_refl eq_refl). assumption. diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v index 8b5024af..8cc547a0 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/FixedEquivocation.v @@ -1404,7 +1404,7 @@ Proof. reflexivity. } apply Hsim in H. - remember (pre_VLSM_projection_trace_project _ _ _ _ _) as tr. + remember (pre_VLSM_projection_finite_trace_project _ _ _ _ _) as tr. replace tr with (equivocators_total_trace_project IM trX); [assumption|]. subst. symmetry. apply (equivocators_total_VLSM_projection_trace_project IM (proj1 Hpre_tr)). diff --git a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v index 6a883f7c..f0d59a81 100644 --- a/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v +++ b/theories/VLSM/Core/Equivocators/Composition/LimitedEquivocation/LimitedEquivocation.v @@ -390,7 +390,7 @@ Proof. reflexivity. } apply Hsim in HtrX. - remember (pre_VLSM_projection_trace_project _ _ _ _ _) as tr. + remember (pre_VLSM_projection_finite_trace_project _ _ _ _ _) as tr. replace tr with (equivocators_total_trace_project IM trX); [assumption|]. subst. symmetry. apply (equivocators_total_VLSM_projection_trace_project IM (proj1 Hpre_tr)). diff --git a/theories/VLSM/Core/Equivocators/Composition/Projections.v b/theories/VLSM/Core/Equivocators/Composition/Projections.v index 048df4ce..63d4fbd3 100644 --- a/theories/VLSM/Core/Equivocators/Composition/Projections.v +++ b/theories/VLSM/Core/Equivocators/Composition/Projections.v @@ -898,7 +898,7 @@ Proof. destruct Hproject_tr as [trX0 [HtrX0 HtrX]]. specialize (IHtr _ _ HtrX0). unfold finite_trace_projection_list in Hproject_tri. - rewrite @pre_VLSM_projection_trace_project_app in Hproject_tri. + rewrite @pre_VLSM_projection_finite_trace_project_app in Hproject_tri. apply equivocator_vlsm_trace_project_app in Hproject_tri. destruct Hproject_tri as [eqv_final' [trXi' [project_xi [HtrXi' [Hproject_xi HeqtrXi]]]]]. assert (Hfinal'i : final_descriptors' i = eqv_final' /\ finite_trace_projection_list IM i projectx = project_xi). @@ -1536,12 +1536,12 @@ Qed. Lemma equivocators_total_VLSM_projection_trace_project {s tr} (Hpre_tr : finite_valid_trace_from PreFreeE s tr) - : @pre_VLSM_projection_trace_project _ (type PreFreeE) _ equivocators_total_label_project + : @pre_VLSM_projection_finite_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. + rewrite @pre_VLSM_projection_finite_trace_project_app. apply finite_valid_trace_from_app_iff in Hpre_tr as [Hpre_tr Hpre_x]. specialize (IHtr Hpre_tr). rewrite IHtr. @@ -2145,7 +2145,7 @@ Proof. simpl. unfold pre_VLSM_projection_transition_item_project, composite_label_sub_projection_option, - pre_VLSM_full_projection_trace_item_project. + pre_VLSM_full_projection_transition_item_project. simpl. case_decide. - f_equal; [|assumption]. @@ -2290,7 +2290,7 @@ Proof. reflexivity. } apply Hsim in H. - remember (pre_VLSM_projection_trace_project _ _ _ _ _) as tr. + remember (pre_VLSM_projection_finite_trace_project _ _ _ _ _) as tr. replace tr with (equivocators_total_trace_project IM trX); [assumption|]. subst. symmetry. apply (equivocators_total_VLSM_projection_trace_project IM (proj1 Hpre_tr)). @@ -2313,7 +2313,7 @@ Proof. reflexivity. } apply Hsim in H as Hpr. - remember (pre_VLSM_projection_trace_project _ _ _ _ _) as tr. + remember (pre_VLSM_projection_finite_trace_project _ _ _ _ _) as tr. replace tr with (equivocators_total_trace_project IM trX); [assumption|]. subst. symmetry. apply (equivocators_total_VLSM_projection_trace_project IM (proj1 H)). diff --git a/theories/VLSM/Core/ProjectionTraces.v b/theories/VLSM/Core/ProjectionTraces.v index 93f6920c..52b9ee9d 100644 --- a/theories/VLSM/Core/ProjectionTraces.v +++ b/theories/VLSM/Core/ProjectionTraces.v @@ -566,7 +566,7 @@ Qed. Definition finite_trace_projection_list (tr : list (composite_transition_item IM)) : list (vtransition_item (IM j)) := - @pre_VLSM_projection_trace_project _ (composite_type IM) _ + @pre_VLSM_projection_finite_trace_project _ (composite_type IM) _ (composite_project_label IM j) (fun s => s j) tr. Lemma preloaded_valid_state_projection diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index b589c801..b76da3b2 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -248,7 +248,7 @@ Qed. Lemma composite_trace_sub_projection_lift (tr : list (composite_transition_item sub_IM)) - : @pre_VLSM_projection_trace_project _ (composite_type IM) _ + : @pre_VLSM_projection_finite_trace_project _ (composite_type IM) _ composite_label_sub_projection_option composite_state_sub_projection (pre_VLSM_full_projection_finite_trace_project _ _ lift_sub_label lift_sub_state tr) = tr. @@ -321,14 +321,11 @@ Proof. - 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, input_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, om). +Lemma induced_sub_projection_valid_projection_strong i Hi li s om + : vvalid induced_sub_projection (existT (dexist i Hi) li) (s, om) -> + input_valid (pre_loaded_with_all_messages_vlsm (IM i)) li (s (dexist i Hi), om). Proof. - destruct l as (sub_i, li). - destruct Hv as [lX [sX [HlX [Heqs [HsX [Hom Hv]]]]]]. - destruct lX as [i _li]. + intros ((_i, _li) & sX & HlX & Heqs & HsX & Hom & Hv). unfold composite_label_sub_projection_option in HlX. simpl in HlX. case_decide; [|congruence]. @@ -336,12 +333,11 @@ Proof. simpl in HlX. apply Some_inj in HlX. inversion HlX. subst. - simpl_existT. subst. - exists i. - split; [assumption|]. + apply + (@dec_sig_sigT_eq_rev _ _ _ sub_index_prop_dec (fun i => vlabel (IM i))) + in HlX as ->. apply proj1 in Hv. cbn in Hv. - exists li, (sX i). 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))). @@ -349,6 +345,20 @@ Proof. assumption. 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, input_valid (pre_loaded_with_all_messages_vlsm (IM i)) l (s, om). +Proof. + destruct l as (sub_i, li). + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + exists i. + split; [assumption|]. + eexists _,_. + eapply induced_sub_projection_valid_projection_strong; eassumption. +Qed. + Lemma induced_sub_projection_transition_is_composite l s om : vtransition induced_sub_projection l (s, om) = composite_transition sub_IM l (s, om). Proof. @@ -376,6 +386,26 @@ Proof. apply lift_sub_state_to_eq. Qed. +Lemma induced_sub_projection_preloaded_free_incl + : VLSM_incl induced_sub_projection (pre_loaded_with_all_messages_vlsm (free_composite_vlsm sub_IM)). +Proof. + apply basic_VLSM_strong_incl. + 2: cbv; intuition. + - intros s (sX & <- & HsX) sub_i. + destruct_dec_sig sub_i i Hi Heqsub_i. + subst. + apply (HsX i). + - intros (sub_i, li) s om Hv. + split; [|exact I]. + cbn. + destruct_dec_sig sub_i i Hi Heqsub_i; subst. + apply induced_sub_projection_valid_projection_strong. + assumption. + - intros l s om s' om'; cbn; intros <-. + symmetry. + apply induced_sub_projection_transition_is_composite. +Qed. + End sec_induced_sub_projection. Section induced_sub_projection_subsumption. @@ -403,7 +433,7 @@ Definition from_sub_projection : composite_transition_item IM -> Prop := Definition finite_trace_sub_projection : list (composite_transition_item IM) -> list (composite_transition_item sub_IM) := - @pre_VLSM_projection_trace_project _ (composite_type IM) _ composite_label_sub_projection_option composite_state_sub_projection. + @pre_VLSM_projection_finite_trace_project _ (composite_type IM) _ composite_label_sub_projection_option composite_state_sub_projection. Section sub_projection_with_no_equivocation_constraints. @@ -526,7 +556,7 @@ Definition finite_trace_sub_projection_app : finite_trace_sub_projection (tr1 ++ tr2) = finite_trace_sub_projection tr1 ++ finite_trace_sub_projection tr2 := - @pre_VLSM_projection_trace_project_app _ (composite_type IM) _ composite_label_sub_projection_option composite_state_sub_projection tr1 tr2. + @pre_VLSM_projection_finite_trace_project_app _ (composite_type IM) _ composite_label_sub_projection_option composite_state_sub_projection tr1 tr2. Lemma X_incl_Pre : VLSM_incl X Pre. Proof. diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index d6deee64..0ed47abe 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -1,4 +1,5 @@ From stdpp Require Import prelude. +From Coq Require Import FunctionalExtensionality. From VLSM Require Import Core.VLSM. From VLSM.Lib Require Import Preamble ListExtras StreamExtras StreamFilters. @@ -312,7 +313,7 @@ Proof. apply pre_VLSM_projection_transition_item_project_is_Some_rev. Qed. -Definition pre_VLSM_projection_trace_project +Definition pre_VLSM_projection_finite_trace_project : list (@transition_item _ TX) -> list (@transition_item _ TY) := map_option pre_VLSM_projection_transition_item_project. @@ -328,28 +329,28 @@ Definition pre_VLSM_projection_infinite_finite_trace_project (s : Streams.Stream (@transition_item _ TX)) (Hs : FinitelyManyBound pre_VLSM_projection_in_projection s) : list (@transition_item _ TY) := - pre_VLSM_projection_trace_project (stream_prefix s (proj1_sig Hs)). + pre_VLSM_projection_finite_trace_project (stream_prefix s (proj1_sig Hs)). -Definition pre_VLSM_projection_trace_project_app - : forall l1 l2, pre_VLSM_projection_trace_project (l1 ++ l2) = - pre_VLSM_projection_trace_project l1 ++ pre_VLSM_projection_trace_project l2 +Definition pre_VLSM_projection_finite_trace_project_app + : forall l1 l2, pre_VLSM_projection_finite_trace_project (l1 ++ l2) = + pre_VLSM_projection_finite_trace_project l1 ++ pre_VLSM_projection_finite_trace_project l2 := map_option_app _. -Definition pre_VLSM_projection_trace_project_app_rev - : forall l l1' l2', pre_VLSM_projection_trace_project l = l1' ++ l2' -> +Definition pre_VLSM_projection_finite_trace_project_app_rev + : forall l l1' l2', pre_VLSM_projection_finite_trace_project l = l1' ++ l2' -> exists l1 l2, l = l1 ++ l2 /\ - pre_VLSM_projection_trace_project l1 = l1' /\ - pre_VLSM_projection_trace_project l2 = l2' + pre_VLSM_projection_finite_trace_project l1 = l1' /\ + pre_VLSM_projection_finite_trace_project l2 = l2' := map_option_app_rev _. -Definition pre_VLSM_projection_trace_project_in_iff - : forall trX itemY, In itemY (pre_VLSM_projection_trace_project trX) <-> +Definition pre_VLSM_projection_finite_trace_project_in_iff + : forall trX itemY, In itemY (pre_VLSM_projection_finite_trace_project trX) <-> exists itemX, In itemX trX /\ pre_VLSM_projection_transition_item_project itemX = Some itemY := in_map_option _. -Definition pre_VLSM_projection_trace_project_in +Definition pre_VLSM_projection_finite_trace_project_in : forall itemX itemY, pre_VLSM_projection_transition_item_project itemX = Some itemY -> - forall trX, In itemX trX -> In itemY (pre_VLSM_projection_trace_project trX) + forall trX, In itemX trX -> In itemY (pre_VLSM_projection_finite_trace_project trX) := in_map_option_rev _. End pre_definitions. @@ -360,7 +361,7 @@ Record VLSM_projection_type (TY : VLSMType message) (label_project : vlabel X -> option (@label _ TY)) (state_project : vstate X -> @state _ TY) - (trace_project := pre_VLSM_projection_trace_project (type X) TY label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project (type X) TY label_project state_project) := { final_state_project : forall sX trX, @@ -376,7 +377,7 @@ Definition VLSM_partial_trace_project_from_projection {TY : VLSMType message} (label_project : vlabel X -> option (@label _ TY)) (state_project : vstate X -> @state _ TY) - (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project _ _ label_project state_project) := fun str : vstate X * list (vtransition_item X) => let (s, tr) := str in Some (state_project s, trace_project tr). @@ -385,7 +386,7 @@ Context {X Y : VLSM message} {label_project : vlabel X -> option (vlabel Y)} {state_project : vstate X -> vstate Y} - (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project _ _ label_project state_project) (Hsimul : VLSM_projection_type X (type Y) label_project state_project) . @@ -397,7 +398,7 @@ Lemma VLSM_partial_projection_type_from_projection Proof. split; intros; inversion H; subst; clear H. exists (state_project s'X), (trace_project preX). split. - + simpl. f_equal. f_equal. apply pre_VLSM_projection_trace_project_app. + + simpl. f_equal. f_equal. apply pre_VLSM_projection_finite_trace_project_app. + symmetry. apply (final_state_project _ _ _ _ Hsimul). apply (finite_valid_trace_from_app_iff X) in H1. apply H1. @@ -413,7 +414,7 @@ Context (TY : VLSMType message) (label_project : vlabel X -> option (@label _ TY)) (state_project : vstate X -> @state _ TY) - (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project _ _ label_project state_project) . (** When a label cannot be projected, and thus the transition will not be @@ -448,7 +449,7 @@ Context (X Y : VLSM message) (label_project : vlabel X -> option (vlabel Y)) (state_project : vstate X -> vstate Y) - (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project _ _ label_project state_project) . (** Similarly to the [VLSM_partial_projection] case we distinguish two types of @@ -541,7 +542,7 @@ Definition VLSM_weak_projection_trace_project {state_project : vstate X -> vstate Y} (Hsimul : VLSM_weak_projection X Y label_project state_project) : list (vtransition_item X) -> list (vtransition_item Y) - := pre_VLSM_projection_trace_project _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project _ _ label_project state_project. Definition VLSM_weak_projection_in {message : Type} @@ -584,14 +585,14 @@ Context Definition VLSM_weak_projection_trace_project_app : forall l1 l2, VLSM_weak_projection_trace_project Hsimul (l1 ++ l2) = VLSM_weak_projection_trace_project Hsimul l1 ++ VLSM_weak_projection_trace_project Hsimul l2 - := pre_VLSM_projection_trace_project_app _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project_app _ _ label_project state_project. Definition VLSM_weak_projection_trace_project_app_rev : forall l l1' l2', VLSM_weak_projection_trace_project Hsimul l = l1' ++ l2' -> exists l1 l2, l = l1 ++ l2 /\ VLSM_weak_projection_trace_project Hsimul l1 = l1' /\ VLSM_weak_projection_trace_project Hsimul l2 = l2' - := pre_VLSM_projection_trace_project_app_rev _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project_app_rev _ _ label_project state_project. Definition VLSM_weak_projection_finite_trace_last : forall sX trX, @@ -723,7 +724,7 @@ Definition VLSM_projection_trace_project {state_project : vstate X -> vstate Y} (Hsimul : VLSM_projection X Y label_project state_project) : list (vtransition_item X) -> list (vtransition_item Y) - := pre_VLSM_projection_trace_project _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project _ _ label_project state_project. Definition VLSM_projection_in {message : Type} @@ -766,19 +767,19 @@ Context Definition VLSM_projection_trace_project_app : forall l1 l2, VLSM_projection_trace_project Hsimul (l1 ++ l2) = VLSM_projection_trace_project Hsimul l1 ++ VLSM_projection_trace_project Hsimul l2 - := pre_VLSM_projection_trace_project_app _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project_app _ _ label_project state_project. Definition VLSM_projection_trace_project_app_rev : forall l l1' l2', VLSM_projection_trace_project Hsimul l = l1' ++ l2' -> exists l1 l2, l = l1 ++ l2 /\ VLSM_projection_trace_project Hsimul l1 = l1' /\ VLSM_projection_trace_project Hsimul l2 = l2' - := pre_VLSM_projection_trace_project_app_rev _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project_app_rev _ _ label_project state_project. Definition VLSM_projection_trace_project_in : forall itemX itemY, pre_VLSM_projection_transition_item_project _ _ label_project state_project itemX = Some itemY -> forall trX, In itemX trX -> In itemY (VLSM_projection_trace_project Hsimul trX) - := pre_VLSM_projection_trace_project_in _ _ label_project state_project. + := pre_VLSM_projection_finite_trace_project_in _ _ label_project state_project. Definition VLSM_projection_finite_trace_last : forall sX trX, @@ -1002,7 +1003,7 @@ Context (state_project : @state _ TX -> @state _ TY) . -Definition pre_VLSM_full_projection_trace_item_project +Definition pre_VLSM_full_projection_transition_item_project : @transition_item _ TX -> @transition_item _ TY := fun item => @@ -1014,11 +1015,11 @@ Definition pre_VLSM_full_projection_trace_item_project Definition pre_VLSM_full_projection_finite_trace_project : list (@transition_item _ TX) -> list (@transition_item _ TY) - := map pre_VLSM_full_projection_trace_item_project. + := map pre_VLSM_full_projection_transition_item_project. Definition pre_VLSM_full_projection_infinite_trace_project : Streams.Stream (@transition_item _ TX) -> Streams.Stream (@transition_item _ TY) - := Streams.map pre_VLSM_full_projection_trace_item_project. + := Streams.map pre_VLSM_full_projection_transition_item_project. Lemma pre_VLSM_full_projection_infinite_trace_project_infinitely_often : forall s, InfinitelyOften (is_Some ∘ (Some ∘ label_project ∘ l)) s. @@ -1059,6 +1060,77 @@ Qed. End pre_definitions. +Section pre_definitions_projection_relation. + +Context + {message : Type} + (TX TY : VLSMType message) + (label_project : @label _ TX -> @label _ TY) + (state_project : @state _ TX -> @state _ TY) + . + +Section pre_definitions_projection_relation_right. + +Context + (TZ : VLSMType message) + (label_projectYZ : @label _ TY -> option (@label _ TZ)) + (state_projectYZ : @state _ TY -> @state _ TZ) + . + +Lemma pre_definitions_projection_relation_right_item + : pre_VLSM_projection_transition_item_project TX TZ (label_projectYZ ∘ label_project) (state_projectYZ ∘ state_project) = + pre_VLSM_projection_transition_item_project TY TZ label_projectYZ state_projectYZ ∘ + pre_VLSM_full_projection_transition_item_project TX TY label_project state_project. +Proof. + extensionality item. + reflexivity. +Qed. + +Lemma pre_definitions_projection_relation_right_trace + : pre_VLSM_projection_finite_trace_project TX TZ (label_projectYZ ∘ label_project) (state_projectYZ ∘ state_project) = + pre_VLSM_projection_finite_trace_project TY TZ label_projectYZ state_projectYZ ∘ + pre_VLSM_full_projection_finite_trace_project TX TY label_project state_project. +Proof. + unfold pre_VLSM_projection_finite_trace_project at 1. + rewrite pre_definitions_projection_relation_right_item. + apply map_option_composed_right. +Qed. + +End pre_definitions_projection_relation_right. + +Section pre_definitions_projection_relation_left. + +Context + (TW : VLSMType message) + (label_projectWX : @label _ TW -> option (@label _ TX)) + (state_projectWX : @state _ TW -> @state _ TX) + . + +Lemma pre_definitions_projection_relation_left_item + : pre_VLSM_projection_transition_item_project TW TY (option_map label_project ∘ label_projectWX) (state_project ∘ state_projectWX) = + option_map (pre_VLSM_full_projection_transition_item_project TX TY label_project state_project) ∘ + pre_VLSM_projection_transition_item_project TW TX label_projectWX state_projectWX. +Proof. + extensionality item. + destruct item. + unfold pre_VLSM_projection_transition_item_project; cbn. + destruct (label_projectWX l) as [lX|]; cbn; reflexivity. +Qed. + +Lemma pre_definitions_projection_relation_left_trace + : pre_VLSM_projection_finite_trace_project TW TY (option_map label_project ∘ label_projectWX) (state_project ∘ state_projectWX) = + pre_VLSM_full_projection_finite_trace_project TX TY label_project state_project ∘ + pre_VLSM_projection_finite_trace_project TW TX label_projectWX state_projectWX. +Proof. + unfold pre_VLSM_projection_finite_trace_project at 1. + rewrite pre_definitions_projection_relation_left_item. + apply map_option_composed_left. +Qed. + +End pre_definitions_projection_relation_left. + +End pre_definitions_projection_relation. + Section basic_definitions. Context @@ -1212,13 +1284,13 @@ Qed. End basic_definitions. -Definition VLSM_full_projection_trace_item_project +Definition VLSM_full_projection_transition_item_project {message : Type} {X Y : VLSM message} {label_project : vlabel X -> vlabel Y} {state_project : vstate X -> vstate Y} (Hsimul : VLSM_full_projection X Y label_project state_project) - := pre_VLSM_full_projection_trace_item_project _ _ label_project state_project + := pre_VLSM_full_projection_transition_item_project _ _ label_project state_project . Definition VLSM_full_projection_finite_trace_project @@ -2159,7 +2231,7 @@ Proof. intros is tr Htr. induction Htr using finite_valid_trace_from_rev_ind ; [reflexivity|]. - rewrite (pre_VLSM_projection_trace_project_app _ _ label_project state_project). + rewrite (pre_VLSM_projection_finite_trace_project_app _ _ label_project state_project). rewrite finite_trace_last_is_last. rewrite finite_trace_last_app, <- IHHtr. clear IHHtr. @@ -2198,11 +2270,11 @@ Context Local Lemma basic_VLSM_projection_finite_valid_trace_init_to is s 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). + : finite_valid_trace_from_to Y (state_project is) (state_project s) (pre_VLSM_projection_finite_trace_project _ _ label_project state_project tr). Proof. induction Htr using finite_valid_trace_init_to_rev_strong_ind. - constructor. apply Hstate. assumption. - - unfold pre_VLSM_projection_trace_project. + - unfold pre_VLSM_projection_finite_trace_project. rewrite map_option_app. apply finite_valid_trace_from_to_app with (state_project s) ; [assumption|]. @@ -2227,7 +2299,7 @@ Local Lemma basic_VLSM_projection_finite_valid_trace_from (s : state) (ls : list transition_item) (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). + : finite_valid_trace_from Y (state_project s) (pre_VLSM_projection_finite_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. @@ -2236,7 +2308,7 @@ Proof. 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. + rewrite (pre_VLSM_projection_finite_trace_project_app _ _ label_project state_project) in Happ_pr. 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. @@ -2314,7 +2386,7 @@ Proof. intros is tr Htr. 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 (@pre_VLSM_projection_finite_trace_project_app _ (type (pre_loaded_with_all_messages_vlsm X)) (type Y) label_project state_project). rewrite finite_trace_last_is_last. rewrite finite_trace_last_app, <- IHHtr. clear IHHtr. @@ -2344,7 +2416,7 @@ Proof. 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). + - rewrite (@pre_VLSM_projection_finite_trace_project_app _ (type (pre_loaded_with_all_messages_vlsm X)) (type Y) label_project state_project). apply (finite_valid_trace_from_app_iff (pre_loaded_with_all_messages_vlsm Y)). split; [assumption|]. simpl. unfold pre_VLSM_projection_transition_item_project. @@ -2377,7 +2449,7 @@ Proof. intros is tr Htr. 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 (@pre_VLSM_projection_finite_trace_project_app _ (type (pre_loaded_vlsm X P)) (type Y) label_project state_project). rewrite finite_trace_last_is_last. rewrite finite_trace_last_app, <- IHHtr. clear IHHtr. @@ -2409,7 +2481,7 @@ Proof. 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). + - rewrite (@pre_VLSM_projection_finite_trace_project_app _ (type (pre_loaded_vlsm X P)) (type Y) label_project state_project). apply (finite_valid_trace_from_app_iff (pre_loaded_vlsm Y Q)). split; [assumption|]. simpl. unfold pre_VLSM_projection_transition_item_project. @@ -2904,7 +2976,7 @@ Context (TY : VLSMType message) (label_project : vlabel X -> option (@label _ TY)) (state_project : vstate X -> @state _ TY) - (trace_project := pre_VLSM_projection_trace_project _ _ label_project state_project) + (trace_project := pre_VLSM_projection_finite_trace_project _ _ label_project state_project) (label_lift : @label _ TY -> vlabel X) (state_lift : @state _ TY -> vstate X) . @@ -3083,11 +3155,11 @@ 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) + (pre_VLSM_full_projection_transition_item_project _ _ label_lift state_lift item) = Some item. Proof. destruct item. - unfold pre_VLSM_full_projection_trace_item_project, + unfold pre_VLSM_full_projection_transition_item_project, pre_VLSM_projection_transition_item_project. simpl. rewrite Hlabel_lift. @@ -3097,7 +3169,7 @@ Qed. Lemma induced_projection_trace_lift (tr : list (@transition_item _ TY)) - : @pre_VLSM_projection_trace_project _ (type X) _ + : @pre_VLSM_projection_finite_trace_project _ (type X) _ label_project state_project (pre_VLSM_full_projection_finite_trace_project _ _ label_lift state_lift tr) = tr. @@ -3142,7 +3214,7 @@ Lemma projection_induced_vlsm_incl (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) + (trace_project := pre_VLSM_projection_finite_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) @@ -3203,7 +3275,7 @@ Lemma projection_induced_vlsm_eq (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) + (trace_project := pre_VLSM_projection_finite_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) diff --git a/theories/VLSM/Lib/ListExtras.v b/theories/VLSM/Lib/ListExtras.v index b034d934..4800ac78 100644 --- a/theories/VLSM/Lib/ListExtras.v +++ b/theories/VLSM/Lib/ListExtras.v @@ -1,9 +1,22 @@ From stdpp Require Import prelude finite. -From Coq Require Import FinFun. +From Coq Require Import FinFun FunctionalExtensionality. From VLSM Require Import Lib.Preamble. (** * Utility lemmas about lists *) +Lemma map_composed + [A B C : Type] + (f : A -> B) + (g : B -> C) + : map (g ∘ f) = map g ∘ map f. +Proof. + extensionality l. + induction l; [reflexivity|]; cbn. + rewrite IHl. + reflexivity. +Qed. + + (** A list is empty if it has no members *) Lemma empty_nil [X:Type] (l:list X) : (forall v, ~In v l) -> l = []. @@ -1268,6 +1281,32 @@ Definition map_option ) []. +Lemma map_option_composed_right + [A B C : Type] + (f : A -> B) + (g : B -> option C) + : map_option (g ∘ f) = map_option g ∘ map f. +Proof. + extensionality l. + induction l; [reflexivity|]. + cbn. + rewrite IHl. + reflexivity. +Qed. + +Lemma map_option_composed_left + [A B C : Type] + (f : A -> option B) + (g : B -> C) + : map_option (option_map g ∘ f) = map g ∘ map_option f. +Proof. + extensionality l. + induction l; [reflexivity|]. + cbn. + rewrite IHl. + destruct (f a) as [b|]; reflexivity. +Qed. + Lemma map_option_app {A B : Type} (f : A -> option B) From d1a4a3f0fa2370b322e5b01ac16986194af7248e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Traian=20Florin=20=C8=98erb=C4=83nu=C8=9B=C4=83?= Date: Mon, 31 Jan 2022 14:06:47 +0200 Subject: [PATCH 15/18] Apply suggestions from code review MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Wojciech Kołowski --- .../ByzantineTraces/FixedSetByzantineTraces.v | 2 +- theories/VLSM/Core/SubProjectionTraces.v | 8 +++---- theories/VLSM/Core/VLSMProjections.v | 4 ++-- theories/VLSM/Lib/ListExtras.v | 22 ++++++------------- 4 files changed, 13 insertions(+), 23 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index a2115bab..06c0e6f0 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -300,7 +300,7 @@ Lemma fixed_non_byzantine_projection_incl_preloaded Proof. apply basic_VLSM_strong_incl. - intros is. apply fixed_non_byzantine_projection_initial_state_preservation. - - intros m Hm. exact I. + - intros m Hm. compute. trivial. - intros l s om Hv. split; [|exact I]. revert Hv. diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index b76da3b2..01aaada8 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -390,14 +390,12 @@ Lemma induced_sub_projection_preloaded_free_incl : VLSM_incl induced_sub_projection (pre_loaded_with_all_messages_vlsm (free_composite_vlsm sub_IM)). Proof. apply basic_VLSM_strong_incl. - 2: cbv; intuition. - intros s (sX & <- & HsX) sub_i. destruct_dec_sig sub_i i Hi Heqsub_i. - subst. - apply (HsX i). + subst. apply HsX. + - cbv; intuition. - intros (sub_i, li) s om Hv. - split; [|exact I]. - cbn. + split; [cbn |exact I]. destruct_dec_sig sub_i i Hi Heqsub_i; subst. apply induced_sub_projection_valid_projection_strong. assumption. diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index 0ed47abe..5aa3c83f 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -1093,7 +1093,7 @@ Lemma pre_definitions_projection_relation_right_trace Proof. unfold pre_VLSM_projection_finite_trace_project at 1. rewrite pre_definitions_projection_relation_right_item. - apply map_option_composed_right. + apply map_option_comp_r. Qed. End pre_definitions_projection_relation_right. @@ -1124,7 +1124,7 @@ Lemma pre_definitions_projection_relation_left_trace Proof. unfold pre_VLSM_projection_finite_trace_project at 1. rewrite pre_definitions_projection_relation_left_item. - apply map_option_composed_left. + apply map_option_comp_l. Qed. End pre_definitions_projection_relation_left. diff --git a/theories/VLSM/Lib/ListExtras.v b/theories/VLSM/Lib/ListExtras.v index 4800ac78..d3cd206a 100644 --- a/theories/VLSM/Lib/ListExtras.v +++ b/theories/VLSM/Lib/ListExtras.v @@ -4,19 +4,16 @@ From VLSM Require Import Lib.Preamble. (** * Utility lemmas about lists *) -Lemma map_composed +Lemma map_comp [A B C : Type] (f : A -> B) (g : B -> C) : map (g ∘ f) = map g ∘ map f. Proof. extensionality l. - induction l; [reflexivity|]; cbn. - rewrite IHl. - reflexivity. + induction l; cbn; rewrite ?IHl; reflexivity. Qed. - (** A list is empty if it has no members *) Lemma empty_nil [X:Type] (l:list X) : (forall v, ~In v l) -> l = []. @@ -1281,20 +1278,15 @@ Definition map_option ) []. -Lemma map_option_composed_right - [A B C : Type] - (f : A -> B) - (g : B -> option C) - : map_option (g ∘ f) = map_option g ∘ map f. +Lemma map_option_comp_r + [A B C : Type] (f : A -> B) (g : B -> option C) : + map_option (g ∘ f) = map_option g ∘ map f. Proof. extensionality l. - induction l; [reflexivity|]. - cbn. - rewrite IHl. - reflexivity. + induction l; cbn; rewrite ?IHl; reflexivity. Qed. -Lemma map_option_composed_left +Lemma map_option_comp_l [A B C : Type] (f : A -> option B) (g : B -> C) From 04af695bfaf0c31ddcefadb453908e8fa74fd0ad Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Tue, 8 Feb 2022 19:14:23 +0200 Subject: [PATCH 16/18] progress --- .../ByzantineTraces/FixedSetByzantineTraces.v | 337 +++++++++++------- theories/VLSM/Core/Composition.v | 73 +++- theories/VLSM/Core/SubProjectionTraces.v | 14 + theories/VLSM/Core/VLSM.v | 12 + theories/VLSM/Core/VLSMProjections.v | 12 + 5 files changed, 314 insertions(+), 134 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index 06c0e6f0..3dcaad63 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -101,8 +101,8 @@ Proof. unfold fixed_byzantine_IM, update_IM. simpl. intros i m Hm. case_decide as Hi. - - elim (no_initial_messages_in_byzantine_IM (dexist i Hi) m). assumption. - - elim (no_initial_messages_in_IM i m). assumption. + - eapply no_initial_messages_in_byzantine_IM; eassumption. + - eapply no_initial_messages_in_IM; eassumption. Qed. Lemma fixed_byzantine_IM_preserves_channel_authentication @@ -116,9 +116,7 @@ Proof. destruct (sender m) as [v|]; [|inversion 1]. case_decide as HAv; [|inversion 1]. cbn. - intro Heq. - apply Some_inj, dec_sig_eq_iff in Heq. - simpl in Heq. + intro Heq; apply Some_inj, dec_sig_eq_iff in Heq; cbn in *. congruence. - apply can_emit_signed; assumption. Qed. @@ -139,7 +137,7 @@ Definition non_byzantine_not_equivocating_constraint : composite_label fixed_byzantine_IM -> composite_state fixed_byzantine_IM * option message -> Prop := sub_IM_not_equivocating_constraint fixed_byzantine_IM non_byzantine A sender sub_fixed_byzantine_IM_Hbs. -Definition fixed_byantine_composite_vlsm : VLSM message := +Definition fixed_byzantine_composite_vlsm : VLSM message := composite_vlsm fixed_byzantine_IM non_byzantine_not_equivocating_constraint. (** @@ -159,7 +157,7 @@ Definition pre_loaded_fixed_non_byzantine_vlsm' : VLSM message := pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM non_byzantine)). Lemma pre_loaded_fixed_non_byzantine_vlsm'_projection - : VLSM_projection fixed_byantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm' + : VLSM_projection fixed_byzantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm' (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine) (composite_state_sub_projection fixed_byzantine_IM non_byzantine). Proof. @@ -177,24 +175,16 @@ Qed. Definition pre_loaded_fixed_non_byzantine_vlsm : VLSM message := pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM IM non_byzantine)). -Lemma non_byzantine_nodes_same - : forall sub_i, sub_IM fixed_byzantine_IM non_byzantine sub_i = sub_IM IM non_byzantine sub_i. +Definition non_byzantine_nodes_same sub_i + : sub_IM fixed_byzantine_IM non_byzantine sub_i = sub_IM IM non_byzantine sub_i. Proof. - intro sub_i. - destruct_dec_sig sub_i i Hi Heqsub_i. - subst. - unfold sub_IM, fixed_byzantine_IM, update_IM. - simpl. - apply set_diff_elim2 in Hi. - case_decide; [contradiction|]. - reflexivity. -Qed. + unfold sub_IM, fixed_byzantine_IM, update_IM; cbn. + case_decide as Hi; [|reflexivity]. + exfalso. + eapply selection_complement_index_not_in_selection; eassumption. +Defined. -Lemma non_byzantine_nodes_same_sym - : forall sub_i, sub_IM IM non_byzantine sub_i = sub_IM fixed_byzantine_IM non_byzantine sub_i. -Proof. - intro. symmetry. apply non_byzantine_nodes_same. -Qed. +Definition non_byzantine_nodes_same_sym := same_IM_sym non_byzantine_nodes_same. Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection : VLSM_full_projection @@ -225,7 +215,7 @@ Definition pre_loaded_fixed_non_byzantine_state_projection := (composite_state_sub_projection fixed_byzantine_IM non_byzantine). Lemma pre_loaded_fixed_non_byzantine_vlsm_projection - : VLSM_projection fixed_byantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm + : VLSM_projection fixed_byzantine_composite_vlsm pre_loaded_fixed_non_byzantine_vlsm pre_loaded_fixed_non_byzantine_label_projection pre_loaded_fixed_non_byzantine_state_projection. Proof. @@ -243,33 +233,26 @@ Proof. ; unfold pre_loaded_fixed_non_byzantine_label_projection, pre_loaded_fixed_non_byzantine_state_projection ; cbn in Hrew |- * ; rewrite Hrew. - induction trX using rev_ind; [reflexivity|]. - rewrite finite_trace_last_is_last. - cbn. + rewrite finite_trace_last_is_last; cbn. rewrite pre_VLSM_projection_finite_trace_project_app. - setoid_rewrite map_app. - rewrite !map_app. - destruct x; destruct l as (i, li); cbn. - unfold pre_VLSM_projection_transition_item_project, pre_loaded_fixed_non_byzantine_label_projection. - cbn. - unfold composite_label_sub_projection_option at 2. - cbn. + unfold pre_VLSM_full_projection_finite_trace_project; rewrite !map_app. + destruct x, l as (i, li); cbn. + unfold pre_VLSM_projection_transition_item_project, pre_loaded_fixed_non_byzantine_label_projection; cbn. + unfold composite_label_sub_projection_option at 2; cbn. case_decide as Hi; cbn; [rewrite last_is_last; reflexivity|]. apply finite_valid_trace_from_app_iff in HtrX as [HtrX Hx]. spec IHtrX HtrX. inversion Hx; subst; clear Hx Htl. eapply induced_sub_projection_transition_consistency_None with (sub_index_list := non_byzantine) in Ht. - + rewrite Ht. - cbn in IHtrX. - setoid_rewrite IHtrX. - rewrite app_nil_r. + + rewrite Ht; cbn in *. + rewrite IHtrX, app_nil_r. reflexivity. + unfold composite_label_sub_projection_option; cbn. case_decide; [contradiction|reflexivity]. - - cbn. - apply (VLSM_full_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm_full_projection). - revert HtrX. - apply (VLSM_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm'_projection). + - apply (VLSM_full_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm_full_projection). + eapply (VLSM_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm'_projection). + eassumption. Qed. (** ** Byzantine traces characterization as projections. *) @@ -284,28 +267,23 @@ Lemma fixed_non_byzantine_projection_initial_state_preservation composite_initial_state_prop (sub_IM fixed_byzantine_IM non_byzantine) s. Proof. split. - - intros Hs sub_i. - destruct Hs as [sX [HeqsX Hinitial]]. - subst. + - intros (sX & HeqsX & Hinitial) sub_i; subst. apply Hinitial. - intros Hs. - exists (lift_sub_state fixed_byzantine_IM non_byzantine s). - split. + eexists; split. + apply composite_state_sub_projection_lift_to. - + apply (lift_sub_state_initial fixed_byzantine_IM non_byzantine); assumption. + + apply (lift_sub_state_initial fixed_byzantine_IM); assumption. Qed. Lemma fixed_non_byzantine_projection_incl_preloaded : VLSM_incl fixed_non_byzantine_projection (pre_loaded_with_all_messages_vlsm (free_composite_vlsm (sub_IM fixed_byzantine_IM non_byzantine))). Proof. - apply basic_VLSM_strong_incl. - - intros is. apply fixed_non_byzantine_projection_initial_state_preservation. - - intros m Hm. compute. trivial. - - intros l s om Hv. - split; [|exact I]. - revert Hv. - apply induced_sub_projection_valid_preservation. - - intros l s om s' om'. apply induced_sub_projection_transition_preservation. + apply basic_VLSM_strong_incl; intros ? *. + - apply fixed_non_byzantine_projection_initial_state_preservation. + - intros _; compute; trivial. + - split; [|exact I]. + eapply induced_sub_projection_valid_preservation; eassumption. + - apply induced_sub_projection_transition_preservation. Qed. (** The induced projection from the composition of [fixed_byzantine_IM] under @@ -317,16 +295,12 @@ Lemma fixed_non_byzantine_projection_friendliness (induced_sub_projection_is_projection fixed_byzantine_IM non_byzantine non_byzantine_not_equivocating_constraint). Proof. - 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. + apply induced_sub_projection_friendliness, induced_sub_projection_lift. + intros s1 s2 Heq l [m|]; [|intuition]; cbn. + destruct (sender m) as [v|]; [|intuition]; cbn. case_decide as HAv; [|intuition]. replace (s2 (A v)) with (s1 (A v)); [intuition|]. - exact (equal_f_dep Heq (dexist (A v) HAv)). + apply (equal_f_dep Heq (dexist (A v) HAv)). Qed. End sec_fixed_non_byzantine_projection. @@ -340,35 +314,190 @@ 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). *) +Record fixed_byzantine_trace_witness + (is : composite_state (sub_IM IM non_byzantine)) + (tr : list (composite_transition_item (sub_IM IM non_byzantine))) + (byzantine_replacements_IM : sub_index byzantine -> VLSM message) + (byzantine_is : vstate (fixed_byzantine_composite_vlsm byzantine_replacements_IM)) + (byzantine_tr : list (vtransition_item (fixed_byzantine_composite_vlsm byzantine_replacements_IM))) + : Prop := + { + no_initial_messages_in_byzantine_IM : no_initial_messages_in_IM_prop byzantine_replacements_IM; + can_emit_signed_byzantine : channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender); + byzantine_tr_valid : finite_valid_trace (fixed_byzantine_composite_vlsm byzantine_replacements_IM) + byzantine_is byzantine_tr; + byzantine_is_pr : pre_loaded_fixed_non_byzantine_state_projection byzantine_replacements_IM + byzantine_is = is; + byzantine_tr_pr : VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection byzantine_replacements_IM) + byzantine_tr = tr + }. + Definition fixed_byzantine_trace_prop (is : composite_state (sub_IM IM non_byzantine)) (tr : list (composite_transition_item (sub_IM IM non_byzantine))) : Prop := - exists - (byzantine_replacements_IM : sub_index byzantine -> VLSM message) - (no_initial_messages_in_byzantine_IM : no_initial_messages_in_IM_prop byzantine_replacements_IM) - (can_emit_signed_byzantine : channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender)), - exists bis btr, - finite_valid_trace (fixed_byantine_composite_vlsm byzantine_replacements_IM) bis btr /\ - pre_loaded_fixed_non_byzantine_state_projection byzantine_replacements_IM bis = is /\ - VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection byzantine_replacements_IM) btr = tr. - -(** Characterization result for the first definition: -the [fixed_byzantine_trace_prop]erty is equivalent to 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_valid_trace (fixed_non_byzantine_projection byzantine_replacements_IM) is tr. + exists byzantine_replacements_IM bis btr, + fixed_byzantine_trace_witness is tr byzantine_replacements_IM bis btr. + +End sec_fixed_byzantine_traces. + +Section fixed_non_equivocating_vs_byzantine. + +Context + {message : Type} + `{finite.Finite index} + (IM : index -> VLSM message) + `{Hbs : forall i : index, HasBeenSentCapability (IM i)} + `{Hbr : forall i : index, HasBeenReceivedCapability (IM i)} + (selection : set index) + {validator : Type} + (A : validator -> index) + (sender : message -> option validator) + (selection_complement := set_diff (enum index) selection) + (Fixed : VLSM message := fixed_equivocation_vlsm_composition IM Hbs Hbr selection) + (StrongFixed := strong_fixed_equivocation_vlsm_composition IM Hbs 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 := + channel_authentication_sender_safety IM A sender can_emit_signed) + . + +Definition byzantine_id_replacements_same_IM + : forall i, fixed_byzantine_IM IM selection (fun sub_i => IM (` sub_i)) i = IM i + := update_IM_id_same_IM IM selection. + +Definition byzantine_id_replacements_same_IM_sym := same_IM_sym byzantine_id_replacements_same_IM. + +Lemma byzantine_id_replacements_full_projection + : VLSM_full_projection + (fixed_byzantine_composite_vlsm IM Hbs selection A sender (fun sub_i => IM (` sub_i))) + (composite_vlsm IM (sub_IM_not_equivocating_constraint IM selection_complement A sender (sub_has_been_sent_capabilities IM selection_complement Hbs))) + (same_IM_label_rew byzantine_id_replacements_same_IM) + (same_IM_state_rew byzantine_id_replacements_same_IM). +Proof. + apply same_IM_full_projection; intros s Hs l [m|]; [|trivial]. + subst selection_complement + ; unfold non_byzantine_not_equivocating_constraint, + sub_IM_not_equivocating_constraint. + destruct (sender m) as [v|]; [|trivial]; cbn. + case_decide as HAv; [|trivial]. + unfold sub_IM, fixed_byzantine_IM; cbn; intro Hsent. + eapply VLSM_full_projection_has_been_sent in Hsent + ; [eassumption|apply same_VLSM_preloaded_full_projection|]. + eapply VLSM_projection_valid_state in Hs + ; [| apply preloaded_component_projection with (j := A v)]. + assumption. +Qed. + +Lemma byzantine_id_replacements_full_projection_rev + : VLSM_full_projection + (composite_vlsm IM (sub_IM_not_equivocating_constraint IM selection_complement A sender (sub_has_been_sent_capabilities IM selection_complement Hbs))) + (fixed_byzantine_composite_vlsm IM Hbs selection A sender (fun sub_i => IM (` sub_i))) + (same_IM_label_rew byzantine_id_replacements_same_IM_sym) + (same_IM_state_rew byzantine_id_replacements_same_IM_sym). +Proof. + apply same_IM_full_projection; intros s Hs l [m|]; [|trivial]. + subst selection_complement + ; unfold non_byzantine_not_equivocating_constraint, + sub_IM_not_equivocating_constraint. + destruct (sender m) as [v|]; [|trivial]; cbn. + case_decide as HAv; [|trivial]. + unfold sub_IM, fixed_byzantine_IM; cbn; intro Hsent. + eapply VLSM_full_projection_has_been_sent in Hsent + ; [eassumption|apply same_VLSM_preloaded_full_projection|]. + eapply VLSM_projection_valid_state in Hs + ; [| apply preloaded_component_projection with (j := A v)]. + assumption. +Qed. + +Lemma fixed_sub_IM_non_equivocating_subsumption + : input_valid_constraint_subsumption IM + (fixed_equivocation_constraint IM Hbs Hbr selection) + (sub_IM_not_equivocating_constraint IM selection_complement A sender + (sub_has_been_sent_capabilities IM selection_complement Hbs)). +Proof. + intros l (s, [m|]) Hv; cbn; [|trivial]. + apply (fixed_strong_equivocation_subsumption IM Hbs Hbr (listing_from_finite index) selection) + in Hv as Hstrong_v. + destruct Hv as (Hs & _). + destruct (sender m) as [v|] eqn:Hsender; cbn; [|trivial]. + case_decide as HAv; [|trivial]. + unfold sub_IM; cbn. + destruct Hstrong_v as [(i & Hi & Hsent) | Hemitted]; cycle 1. + - exfalso; apply set_diff_elim2 in HAv; contradict HAv. + eapply (sub_can_emit_sender IM); eassumption. + - eapply VLSM_incl_valid_state in Hs; [|apply constraint_free_incl]. + eapply VLSM_incl_valid_state in Hs; [|apply vlsm_incl_pre_loaded_with_all_messages_vlsm]. + cut (@has_been_sent _ _ (Hbs (A v)) (s (A v)) m). + { eapply has_been_sent_irrelevance, + valid_state_project_preloaded_to_preloaded, Hs. + } + eapply valid_state_has_trace in Hs as (is & tr & Htr). + eapply (has_been_sent_iff_by_sender IM). 1-3: eassumption. + eexists; eassumption. +Qed. + +Lemma fixed_non_equivocating_incl_sub_non_equivocating + : VLSM_incl FixedNonEquivocating + (induced_sub_projection IM (set_diff (enum index) selection) + (sub_IM_not_equivocating_constraint IM + (set_diff (enum index) selection) A sender + (sub_has_been_sent_capabilities IM + (set_diff (enum index) selection) Hbs))). +Proof. + apply induced_sub_projection_constraint_subsumption_incl, + fixed_sub_IM_non_equivocating_subsumption. +Qed. + +Lemma Fixed_sub_IM_non_equivocating_incl + : VLSM_incl Fixed (composite_vlsm IM (sub_IM_not_equivocating_constraint IM selection_complement A sender (sub_has_been_sent_capabilities IM selection_complement Hbs))). +Proof. + apply constraint_subsumption_incl, fixed_sub_IM_non_equivocating_subsumption. +Qed. + +Lemma fixed_equivocating_traces_are_byzantine is tr + : finite_valid_trace Fixed is tr -> + fixed_byzantine_trace_prop IM Hbs selection A sender + (composite_state_sub_projection IM selection_complement is) + (finite_trace_sub_projection IM selection_complement tr). +Proof. + intros Htr. + exists (fun sub_i => IM (` sub_i)), + (same_IM_state_rew byzantine_id_replacements_same_IM_sym is), + (VLSM_full_projection_finite_trace_project byzantine_id_replacements_full_projection_rev tr). + constructor. + - apply sub_IM_preserves_no_initial_messages; assumption. + - apply sub_IM_preserves_channel_authentication; assumption. + - apply (VLSM_full_projection_finite_valid_trace byzantine_id_replacements_full_projection_rev). + revert Htr; apply VLSM_incl_finite_valid_trace. + apply Fixed_sub_IM_non_equivocating_incl. + - unfold pre_loaded_fixed_non_byzantine_state_projection, + composite_state_sub_projection, same_IM_state_rew. + extensionality sub_i; cbn. + rewrite <- same_VLSM_state_rew_21 + with (Heq := non_byzantine_nodes_same IM selection (fun sub_i => IM (`sub_i)) sub_i). + f_equal; destruct_dec_sig sub_i i Hi Heqsub_i; subst; cbn. + f_equal. unfold byzantine_id_replacements_same_IM_sym, same_IM_sym. + f_equal. + unfold byzantine_id_replacements_same_IM, non_byzantine_nodes_same, + update_IM_id_same_IM, fixed_byzantine_IM, update_IM; cbn. + apply set_diff_elim2 in Hi as Hni. + case_decide; [contradiction|reflexivity]. + - admit. +Admitted. + +Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection' + : VLSM_full_projection + pre_loaded_fixed_non_byzantine_vlsm + pre_loaded_fixed_non_byzantine_vlsm' + (same_IM_label_rew non_byzantine_nodes_same_sym) + (same_IM_state_rew non_byzantine_nodes_same_sym). Proof. - apply and_comm. - apply (projection_friendly_trace_char (induced_sub_projection_is_projection _ _ _)). - apply fixed_non_byzantine_projection_friendliness. + apply same_IM_preloaded_free_full_projection. Qed. -End fixed_byzantine_traces_as_projections. + (** ** Fixed Byzantine traces as traces pre-loaded with signed messages @@ -641,44 +770,6 @@ Context channel_authentication_sender_safety IM A sender can_emit_signed) . -Lemma fixed_non_equivocating_incl_sub_non_equivocating - : VLSM_incl FixedNonEquivocating - (induced_sub_projection IM (set_diff (enum index) selection) - (sub_IM_not_equivocating_constraint IM - (set_diff (enum index) selection) A sender - (sub_has_been_sent_capabilities IM - (set_diff (enum index) selection) Hbs))). -Proof. - apply induced_sub_projection_constraint_subsumption_incl. - intros l (s, om) Hv. - apply (fixed_strong_equivocation_subsumption IM Hbs Hbr (listing_from_finite index) selection) - in Hv as Hstrong_v. - destruct Hv as [Hs [Hom [Hv Hc]]]. - destruct om; [|exact I]. - simpl. - destruct (sender m) as [v|] eqn:Hsender; [|exact I]. - simpl. - case_decide; [|exact I]. - unfold sub_IM. simpl. - 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 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 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. - apply (sub_can_emit_sender IM selection A sender Hsender_safety _ _ _ Hsender Hemitted). -Qed. - Lemma fixed_non_equivocating_incl_fixed_non_byzantine : VLSM_incl FixedNonEquivocating PreNonByzantine. Proof. diff --git a/theories/VLSM/Core/Composition.v b/theories/VLSM/Core/Composition.v index a3e015dd..13f7c626 100644 --- a/theories/VLSM/Core/Composition.v +++ b/theories/VLSM/Core/Composition.v @@ -1429,7 +1429,9 @@ Definition same_IM_state_rew : composite_state IM2 := fun i => same_VLSM_state_rew (Heq i) (s1 i). -Section pre_loaded_constrained. +Definition same_IM_sym : forall i, IM2 i = IM1 i := fun i => eq_sym (Heq i). + +Section sec_same_IM_constrained. Context (constraint1 : composite_label IM1 -> composite_state IM1 * option message -> Prop) @@ -1438,10 +1440,15 @@ Context : 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)) + . + +Section sec_same_IM_preloaded_full_projection. + +Context (seed : message -> Prop) . -Lemma same_IM_full_projection +Lemma same_IM_preloaded_full_projection : VLSM_full_projection (pre_loaded_vlsm (composite_vlsm IM1 constraint1) seed) (pre_loaded_vlsm (composite_vlsm IM2 constraint2) seed) @@ -1478,7 +1485,28 @@ Proof. + exists (exist _ m Hm). reflexivity. Qed. -End pre_loaded_constrained. +End sec_same_IM_preloaded_full_projection. + +Lemma same_IM_full_projection + : VLSM_full_projection + (composite_vlsm IM1 constraint1) + (composite_vlsm IM2 constraint2) + same_IM_label_rew + same_IM_state_rew. +Proof. + constructor. + intros s1 tr1 Htr1. + apply + (VLSM_eq_finite_valid_trace + (vlsm_is_pre_loaded_with_False (composite_vlsm IM2 constraint2))), + (VLSM_full_projection_finite_valid_trace + (same_IM_preloaded_full_projection (fun _ => False))), + (VLSM_eq_finite_valid_trace + (vlsm_is_pre_loaded_with_False (composite_vlsm IM1 constraint1))). + assumption. +Qed. + +End sec_same_IM_constrained. Lemma same_IM_preloaded_free_full_projection : VLSM_full_projection @@ -1489,16 +1517,16 @@ Lemma same_IM_preloaded_free_full_projection 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_valid_trace Heq1) in Htr1. - clear Heq1. - specialize (same_IM_full_projection (free_constraint IM1) (free_constraint IM2)) + specialize (same_IM_preloaded_full_projection (free_constraint IM1) (free_constraint IM2)) as Hproj. - spec Hproj. { intros. exact I. } + spec Hproj; [intros; exact I|]. specialize (Hproj (fun _ => True)). - 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_valid_trace Heq2). + apply + (VLSM_eq_finite_valid_trace + (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (free_composite_vlsm IM2))), + (VLSM_full_projection_finite_valid_trace Hproj), + (VLSM_eq_finite_valid_trace + (pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True (free_composite_vlsm IM1))). assumption. Qed. @@ -1506,3 +1534,26 @@ End sec_same_IM_full_projection. Arguments same_IM_label_rew {_ _ _ _} _ _ : assert. Arguments same_IM_state_rew {_ _ _ _} _ _ _ : assert. +Arguments same_IM_sym {_ _ _ _} _ _ : assert. + +Lemma same_IM_state_rew_12 + {message : Type} + `{EqDecision index} + {IM1 IM2 : index -> VLSM message} + (Heq : forall i, IM1 i = IM2 i) + : forall s : composite_state IM1, + same_IM_state_rew (same_IM_sym Heq) (same_IM_state_rew Heq s) = s. +Proof. + intro; extensionality i; apply same_VLSM_state_rew_12. +Qed. + +Lemma same_IM_state_rew_21 + {message : Type} + `{EqDecision index} + {IM1 IM2 : index -> VLSM message} + (Heq : forall i, IM1 i = IM2 i) + : forall s : composite_state IM2, + same_IM_state_rew Heq (same_IM_state_rew (same_IM_sym Heq) s) = s. +Proof. + intro; extensionality i; apply same_VLSM_state_rew_21. +Qed. diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index 01aaada8..8b219056 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -2229,6 +2229,12 @@ Definition update_IM for fixed-set equivocation model, similar to the one for byzantine traces. *) +Definition update_IM_id_same_IM + : forall i, update_IM (fun sub_i => IM (` sub_i)) i = IM i. +Proof. + intro; unfold update_IM; case_decide; reflexivity. +Defined. + Context (replacement_IM : sub_index selection -> VLSM message) (updated_IM := update_IM replacement_IM) @@ -2236,6 +2242,14 @@ Context (selection_complement : set index := set_diff (enum index) selection) . +Lemma selection_complement_index_not_in_selection + (sub_i : sub_index selection_complement) + : ~ (`sub_i ∈ selection). +Proof. + destruct_dec_sig sub_i i Hi Heqsub_i; subst; cbn in *. + eapply set_diff_elim2; eassumption. +Qed. + Lemma update_IM_complement_Hbs (Hbs : forall i : index, HasBeenSentCapability (IM i)) : forall sub_i : sub_index selection_complement, diff --git a/theories/VLSM/Core/VLSM.v b/theories/VLSM/Core/VLSM.v index ae5b0fa7..0fa3a1bf 100644 --- a/theories/VLSM/Core/VLSM.v +++ b/theories/VLSM/Core/VLSM.v @@ -2927,6 +2927,18 @@ Context (Heq : X1 = X2) . +Lemma same_VLSM_state_rew_12 (s1 : vstate X1) + : same_VLSM_state_rew (eq_sym Heq) (same_VLSM_state_rew Heq s1) = s1. +Proof. + subst; reflexivity. +Qed. + +Lemma same_VLSM_state_rew_21 (s2 : vstate X2) + : same_VLSM_state_rew Heq (same_VLSM_state_rew (eq_sym Heq) s2) = s2. +Proof. + subst; reflexivity. +Qed. + Lemma same_VLSM_valid_preservation l1 s1 om : vvalid X1 l1 (s1, om) -> vvalid X2 (same_VLSM_label_rew Heq l1) (same_VLSM_state_rew Heq s1, om). diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index 5aa3c83f..7d8d8077 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -3316,4 +3316,16 @@ Proof. assumption. Qed. +Lemma same_VLSM_preloaded_full_projection + : VLSM_full_projection + (pre_loaded_with_all_messages_vlsm X1) (pre_loaded_with_all_messages_vlsm X2) + (same_VLSM_label_rew Heq) (same_VLSM_state_rew Heq). +Proof. + apply basic_VLSM_strong_full_projection. + - cbv; apply same_VLSM_valid_preservation. + - cbv; apply same_VLSM_transition_preservation. + - cbv; apply same_VLSM_initial_state_preservation. + - cbv; trivial. +Qed. + End same_VLSM_full_projection. From 480f1b609b999aa3eb9522589c4b259563d0fb27 Mon Sep 17 00:00:00 2001 From: Traian Florin Serbanuta Date: Thu, 10 Feb 2022 12:42:52 +0200 Subject: [PATCH 17/18] defining types --- .../ByzantineTraces/FixedSetByzantineTraces.v | 225 ++++++++++-------- .../ByzantineTraces/LimitedByzantineTraces.v | 69 +++--- theories/VLSM/Core/MessageDependencies.v | 15 ++ theories/VLSM/Core/VLSM.v | 12 + 4 files changed, 195 insertions(+), 126 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index 3dcaad63..2efe8db2 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -1,8 +1,8 @@ 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.Validator Core.Equivocation Core.EquivocationProjections Core.Equivocation.NoEquivocation Core.Equivocation.FixedSetEquivocation. +From VLSM.Core Require Import VLSM MessageDependencies VLSMProjections Composition ProjectionTraces SubProjectionTraces ByzantineTraces Validator Equivocation EquivocationProjections. +From VLSM.Core.Equivocation Require Import NoEquivocation FixedSetEquivocation MsgDepFixedSetEquivocation. (** * VLSM Compositions with a fixed set of byzantine nodes @@ -22,44 +22,6 @@ nodes and equivocating ones. Therefore, when analyzing the security of a protocol it suffices to consider equivocating nodes. *) -Section emit_any_signed_message_vlsm. - -(** ** A machine which can emit any valid message for a given node - -This VLSM is similar to the [emit_any_message_vlsm] with the exception of the -validity predicate which requires that the sender of the emitted message -corresponds to the given node. -*) - -Context - {message : Type} - {index : Type} - {validator : Type} - (A : validator -> index) - (sender : message -> option validator) - (node_idx : index) - . - -(** The [valid]ity predicate allows sending only signed messages -*) -Definition signed_messages_valid - (l : @label message all_messages_type) - (som : @state message all_messages_type * option message) - : Prop := - channel_authenticated_message A sender node_idx l. - -Definition emit_any_signed_message_vlsm_machine - : VLSMClass all_messages_sig - := - {| transition := all_messages_transition - ; valid := signed_messages_valid - |}. - -Definition emit_any_signed_message_vlsm - := mk_vlsm emit_any_signed_message_vlsm_machine. - -End emit_any_signed_message_vlsm. - Section sec_fixed_byzantine_traces. Context @@ -307,6 +269,13 @@ End sec_fixed_non_byzantine_projection. End sec_fixed_byzantine_replacements. + +Definition byzantine_replacements : Type := + { byzantine_replacements_IM : sub_index byzantine -> VLSM message | + no_initial_messages_in_IM_prop byzantine_replacements_IM /\ + channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender) + }. + (** The first definition of the [fixed_byzantine_trace_prop]erty: Fixed byzantine traces are projections to the subset of protocol-following nodes @@ -314,30 +283,20 @@ 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). *) -Record fixed_byzantine_trace_witness - (is : composite_state (sub_IM IM non_byzantine)) - (tr : list (composite_transition_item (sub_IM IM non_byzantine))) - (byzantine_replacements_IM : sub_index byzantine -> VLSM message) - (byzantine_is : vstate (fixed_byzantine_composite_vlsm byzantine_replacements_IM)) - (byzantine_tr : list (vtransition_item (fixed_byzantine_composite_vlsm byzantine_replacements_IM))) - : Prop := - { - no_initial_messages_in_byzantine_IM : no_initial_messages_in_IM_prop byzantine_replacements_IM; - can_emit_signed_byzantine : channel_authentication_prop byzantine_replacements_IM (sub_IM_A byzantine A) (sub_IM_sender byzantine A sender); - byzantine_tr_valid : finite_valid_trace (fixed_byzantine_composite_vlsm byzantine_replacements_IM) - byzantine_is byzantine_tr; - byzantine_is_pr : pre_loaded_fixed_non_byzantine_state_projection byzantine_replacements_IM - byzantine_is = is; - byzantine_tr_pr : VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection byzantine_replacements_IM) - byzantine_tr = tr - }. +Definition fixed_byzantine_trace : Type := + { byz_IM : byzantine_replacements & { istr | + finite_valid_trace (fixed_byzantine_composite_vlsm (` byz_IM)) istr.1 istr.2}}. -Definition fixed_byzantine_trace_prop +Definition trace_exposed_to_fixed_byzantine_behavior_prop (is : composite_state (sub_IM IM non_byzantine)) (tr : list (composite_transition_item (sub_IM IM non_byzantine))) : Prop := - exists byzantine_replacements_IM bis btr, - fixed_byzantine_trace_witness is tr byzantine_replacements_IM bis btr. + exists byz_tr : fixed_byzantine_trace, + match byz_tr with + existT (exist _ byz_IM _) (exist _ (byzantine_is, byzantine_tr) _) => + pre_loaded_fixed_non_byzantine_state_projection byz_IM byzantine_is = is /\ + VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection byz_IM) byzantine_tr = tr + end. End sec_fixed_byzantine_traces. @@ -456,48 +415,125 @@ Proof. apply constraint_subsumption_incl, fixed_sub_IM_non_equivocating_subsumption. Qed. -Lemma fixed_equivocating_traces_are_byzantine is tr - : finite_valid_trace Fixed is tr -> - fixed_byzantine_trace_prop IM Hbs selection A sender - (composite_state_sub_projection IM selection_complement is) - (finite_trace_sub_projection IM selection_complement tr). +Lemma fixed_equivocating_traces_are_byzantine_component i (Hi : i ∈ set_diff (enum index) selection) + : byzantine_id_replacements_same_IM_sym i + = + non_byzantine_nodes_same_sym IM selection (fun sub_i => IM (`sub_i)) (dexist i Hi). Proof. - intros Htr. - exists (fun sub_i => IM (` sub_i)), - (same_IM_state_rew byzantine_id_replacements_same_IM_sym is), - (VLSM_full_projection_finite_trace_project byzantine_id_replacements_full_projection_rev tr). - constructor. + unfold byzantine_id_replacements_same_IM_sym, + non_byzantine_nodes_same_sym, same_IM_sym + ; f_equal. + unfold byzantine_id_replacements_same_IM, non_byzantine_nodes_same, + update_IM_id_same_IM, fixed_byzantine_IM, update_IM; cbn. + apply set_diff_elim2 in Hi as Hni. + case_decide; [contradiction|reflexivity]. +Qed. + +Lemma fixed_equivocating_traces_are_byzantine_label l + : pre_loaded_fixed_non_byzantine_label_projection IM selection (fun sub_i => IM (`sub_i)) + (same_IM_label_rew byzantine_id_replacements_same_IM_sym l) + = + composite_label_sub_projection_option IM selection_complement l. +Proof. + destruct l as (i, li). + unfold pre_loaded_fixed_non_byzantine_label_projection, + composite_label_sub_projection_option, composite_label_sub_projection, same_IM_label_rew + ; subst selection_complement; cbn. + case_decide as Hi; [|reflexivity]; cbn. + do 2 f_equal. + rewrite <- same_VLSM_label_rew_21 + with (Heq := non_byzantine_nodes_same IM selection (fun sub_i => IM (`sub_i)) (dexist i Hi)). + do 2 f_equal. + apply fixed_equivocating_traces_are_byzantine_component. +Qed. + +Lemma fixed_equivocating_traces_are_byzantine_state s + : pre_loaded_fixed_non_byzantine_state_projection IM selection (fun sub_i => IM (`sub_i)) + (same_IM_state_rew byzantine_id_replacements_same_IM_sym s) + = + composite_state_sub_projection IM selection_complement s. +Proof. + unfold pre_loaded_fixed_non_byzantine_state_projection, + composite_state_sub_projection, same_IM_state_rew. + extensionality sub_i; cbn. + rewrite <- same_VLSM_state_rew_21 + with (Heq := non_byzantine_nodes_same IM selection (fun sub_i => IM (`sub_i)) sub_i). + f_equal; destruct_dec_sig sub_i i Hi Heqsub_i; subst; cbn; f_equal. + apply fixed_equivocating_traces_are_byzantine_component. +Qed. + +Lemma fixed_equivocating_traces_are_byzantine_trace tr +: VLSM_projection_trace_project + (pre_loaded_fixed_non_byzantine_vlsm_projection IM Hbs selection A sender (fun sub_i => IM (`sub_i))) + (VLSM_full_projection_finite_trace_project byzantine_id_replacements_full_projection_rev tr) += + finite_trace_sub_projection IM selection_complement tr. +Proof. + induction tr; [reflexivity|]; cbn; rewrite IHtr. + replace (pre_VLSM_projection_transition_item_project _ _ _ _ _) with + (pre_VLSM_projection_transition_item_project (composite_type IM) + (composite_type (sub_IM IM selection_complement)) + (composite_label_sub_projection_option IM selection_complement) + (composite_state_sub_projection IM selection_complement) a) + ; [reflexivity|]. + destruct a; unfold pre_VLSM_projection_transition_item_project; cbn. + rewrite fixed_equivocating_traces_are_byzantine_state, + fixed_equivocating_traces_are_byzantine_label. + reflexivity. +Qed. + +Program Definition fixed_equivocating_traces_are_byzantine_replacements + : byzantine_replacements selection A sender := + exist _ (fun sub_i => IM (` sub_i)) _. +Next Obligation. + split. - apply sub_IM_preserves_no_initial_messages; assumption. - apply sub_IM_preserves_channel_authentication; assumption. - - apply (VLSM_full_projection_finite_valid_trace byzantine_id_replacements_full_projection_rev). - revert Htr; apply VLSM_incl_finite_valid_trace. - apply Fixed_sub_IM_non_equivocating_incl. - - unfold pre_loaded_fixed_non_byzantine_state_projection, - composite_state_sub_projection, same_IM_state_rew. - extensionality sub_i; cbn. - rewrite <- same_VLSM_state_rew_21 - with (Heq := non_byzantine_nodes_same IM selection (fun sub_i => IM (`sub_i)) sub_i). - f_equal; destruct_dec_sig sub_i i Hi Heqsub_i; subst; cbn. - f_equal. unfold byzantine_id_replacements_same_IM_sym, same_IM_sym. - f_equal. - unfold byzantine_id_replacements_same_IM, non_byzantine_nodes_same, - update_IM_id_same_IM, fixed_byzantine_IM, update_IM; cbn. - apply set_diff_elim2 in Hi as Hni. - case_decide; [contradiction|reflexivity]. - - admit. -Admitted. +Qed. -Lemma pre_loaded_fixed_non_byzantine_vlsm_full_projection' - : VLSM_full_projection - pre_loaded_fixed_non_byzantine_vlsm - pre_loaded_fixed_non_byzantine_vlsm' - (same_IM_label_rew non_byzantine_nodes_same_sym) - (same_IM_state_rew non_byzantine_nodes_same_sym). +Program Definition fixed_equivocating_traces_are_byzantine_traces eis etr + (Htr : finite_valid_trace Fixed eis etr) + : fixed_byzantine_trace IM Hbs selection A sender := + (existT fixed_equivocating_traces_are_byzantine_replacements + (exist _ + (same_IM_state_rew byzantine_id_replacements_same_IM_sym eis, + VLSM_full_projection_finite_trace_project byzantine_id_replacements_full_projection_rev etr) + _)). +Next Obligation. + intros; cbn. + apply (VLSM_full_projection_finite_valid_trace byzantine_id_replacements_full_projection_rev). + revert Htr; apply VLSM_incl_finite_valid_trace. + apply Fixed_sub_IM_non_equivocating_incl. +Qed. + +Lemma fixed_equivocating_traces_are_byzantine eis etr + : finite_valid_trace Fixed eis etr -> + trace_exposed_to_fixed_byzantine_behavior_prop IM Hbs selection A sender + (composite_state_sub_projection IM selection_complement eis) + (finite_trace_sub_projection IM selection_complement etr). Proof. - apply same_IM_preloaded_free_full_projection. + intros Htr. + exists (fixed_equivocating_traces_are_byzantine_traces _ _ Htr); cbn; split. + - apply fixed_equivocating_traces_are_byzantine_state. + - apply fixed_equivocating_traces_are_byzantine_trace. Qed. +Lemma fixed_byzantine_traces_are_equivocating + (byz_tr : fixed_byzantine_trace IM Hbs selection A sender) + : match byz_tr with + existT (exist _ byz_IM _) (exist _ (byzantine_is, byzantine_tr) _) => + exists eis, composite_state_sub_projection IM selection_complement eis = pre_loaded_fixed_non_byzantine_state_projection IM selection byz_IM byzantine_is /\ + exists etr, finite_trace_sub_projection IM selection_complement etr = VLSM_projection_trace_project (pre_loaded_fixed_non_byzantine_vlsm_projection IM Hbs selection A sender byz_IM) byzantine_tr /\ + finite_valid_trace Fixed eis etr + end. +Proof. + destruct byz_tr as ((byz_IM & no_initial_byz & can_emit_byz) & (byz_is, byz_tr) & Hbyz_tr); cbn in *. + eexists; split; [apply composite_state_sub_projection_lift|]. + eexists; split; [apply composite_trace_sub_projection_lift|]. + admit. +Admitted. +(* (** ** Fixed Byzantine traces as traces pre-loaded with signed messages @@ -1037,5 +1073,6 @@ Proof. Qed. 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 d52a7e33..70df4662 100644 --- a/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/LimitedByzantineTraces.v @@ -18,54 +18,58 @@ projections of traces of the composition of the regular nodes under a composition constraint allowing only a limited amount of equivocation. *) -Section limited_byzantine_traces. +Section sec_limited_byzantine_traces. Context {message : Type} - `{EqDecision index} + `{finite.Finite index} (IM : index -> VLSM message) (Hbs : forall i : index, HasBeenSentCapability (IM i)) (Hbr : forall i : index, HasBeenReceivedCapability (IM i)) (sender : message -> option index) - {finite_index : finite.Finite index} `{ReachableThreshold index} . -(** -We define the [limited_byzantine_trace_prop]erty in two steps. First, we -leverage the [fixed_byzantine_trace_alt_prop]erty by assuming a fixed selection -of <> nodes whose added weight is below the [ReachableThreshold]. -*) -Definition fixed_limited_byzantine_trace_prop - (s : composite_state IM) - (tr : list (composite_transition_item IM)) - (byzantine : set index) - : Prop - := (sum_weights (remove_dups byzantine) <= `threshold)%R /\ - fixed_byzantine_trace_alt_prop IM Hbs byzantine (fun i => i) sender s tr. +Definition limited_byzantine_set : Type := + { byzantine : set index | (sum_weights (remove_dups byzantine) <= `threshold)%R }. -(** The union of traces with the [fixed_limited_byzantine_trace_prop]erty over -all possible selections of (limited) byzantine nodes. -*) -Definition limited_byzantine_trace_prop - (s : composite_state IM) - (tr : list (composite_transition_item IM)) - : Prop := - exists byzantine, fixed_limited_byzantine_trace_prop s tr byzantine. +Definition limited_byzantine_trace : Type := + { limited_byzantine : limited_byzantine_set & fixed_byzantine_trace IM Hbs (` limited_byzantine) Datatypes.id sender}. 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) - (Hvalidator: forall i : index, component_message_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)) (message_dependencies : message -> set message) (HMsgDep : forall i, MessageDependencies message_dependencies (IM i)) - (Hfull : forall i, message_dependencies_full_node_condition_prop message_dependencies (IM i)) + (full_message_dependencies : message -> set message) + (HFullMsgDep : FullMessageDependencies message_dependencies full_message_dependencies) + (Limited := msg_dep_limited_equivocation_vlsm IM Hbs Hbr full_message_dependencies sender) + (no_initial_messages_in_IM : no_initial_messages_in_IM_prop IM) + (Hchannel : channel_authentication_prop IM Datatypes.id sender) + (Hsender_safety : sender_safety_alt_prop IM Datatypes.id sender := + channel_authentication_sender_safety _ _ _ Hchannel) . +Program Definition limited_equivocating_traces_are_byzantine is tr + (Htr : finite_valid_trace Limited is tr) + (equivocators := state_annotation (finite_trace_last is tr)) + : limited_byzantine_trace := + existT (exist _ equivocators _) + (fixed_equivocating_traces_are_byzantine_traces IM equivocators Datatypes.id sender no_initial_messages_in_IM Hchannel (original_state is) (pre_VLSM_full_projection_finite_trace_project + (type Limited) (composite_type IM) Datatypes.id original_state + tr) _). +Next Obligation. + intros; cbn. + eapply msg_dep_fixed_limited_equivocation_witnessed; eassumption. +Qed. +Next Obligation. + intros; cbn. + eapply msg_dep_fixed_limited_equivocation_witnessed; eassumption. +Qed. + + +End sec_limited_byzantine_traces. + +(* + (** ** Assuming the byzantine nodes are known We will first fix a selection of <> nodes of limited weight and analyze traces with the [fixed_limited_byzantine_trace_prop]erty w.r.t. that @@ -543,4 +547,5 @@ Proof. apply fixed_equivocation_vlsm_composition_index_incl; assumption. Qed. -End sec_msg_dep_limited_byzantine_traces. \ No newline at end of file +End sec_msg_dep_limited_byzantine_traces. +*) \ No newline at end of file diff --git a/theories/VLSM/Core/MessageDependencies.v b/theories/VLSM/Core/MessageDependencies.v index faf19fc0..9f1bb63c 100644 --- a/theories/VLSM/Core/MessageDependencies.v +++ b/theories/VLSM/Core/MessageDependencies.v @@ -77,6 +77,21 @@ relation. *) Definition msg_dep_happens_before : relation message := flip (clos_trans _ (flip msg_dep_rel)). +(** The relaxed (local) full node condition for a given <> +function requires that a state (receiving the message) is aware of all +of <>'s dependencies. +*) +Definition relaxed_message_dependencies_full_node_condition + (s : vstate X) + (m : message) + : Prop := + forall dm, dm ∈ message_dependencies m -> + exists obs_m, has_been_observed X s obs_m /\ msg_dep_happens_before dm obs_m. + +Definition relaxed_message_dependencies_full_node_condition_prop : Prop := + forall l s m, + vvalid X l (s, Some m) -> relaxed_message_dependencies_full_node_condition s m. + (** Unrolling one the [msg_dep_happens_before] relation one step. *) Lemma msg_dep_happens_before_iff_one x z : msg_dep_happens_before x z <-> diff --git a/theories/VLSM/Core/VLSM.v b/theories/VLSM/Core/VLSM.v index 0fa3a1bf..352c1eab 100644 --- a/theories/VLSM/Core/VLSM.v +++ b/theories/VLSM/Core/VLSM.v @@ -2927,6 +2927,18 @@ Context (Heq : X1 = X2) . +Lemma same_VLSM_label_rew_12 (l1 : vlabel X1) + : same_VLSM_label_rew (eq_sym Heq) (same_VLSM_label_rew Heq l1) = l1. +Proof. + subst; reflexivity. +Qed. + +Lemma same_VLSM_label_rew_21 (l2 : vlabel X2) + : same_VLSM_label_rew Heq (same_VLSM_label_rew (eq_sym Heq) l2) = l2. +Proof. + subst; reflexivity. +Qed. + Lemma same_VLSM_state_rew_12 (s1 : vstate X1) : same_VLSM_state_rew (eq_sym Heq) (same_VLSM_state_rew Heq s1) = s1. Proof. From 07c34a231d27da1cf323f0c0fb33d51dde070c5c Mon Sep 17 00:00:00 2001 From: Traian-Florin Serbanuta Date: Tue, 22 Feb 2022 21:30:53 +0200 Subject: [PATCH 18/18] progress --- .../ByzantineTraces/FixedSetByzantineTraces.v | 95 ++++++++++++------- theories/VLSM/Core/SubProjectionTraces.v | 66 +++++++++++++ theories/VLSM/Core/VLSMProjections.v | 56 ++++++++++- theories/VLSM/Lib/ListExtras.v | 11 +++ 4 files changed, 192 insertions(+), 36 deletions(-) diff --git a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v index 14f110ab..ed4de5f5 100644 --- a/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v +++ b/theories/VLSM/Core/ByzantineTraces/FixedSetByzantineTraces.v @@ -164,7 +164,7 @@ Proof. Qed. Definition pre_loaded_fixed_non_byzantine_label_projection := - (option_map (same_IM_label_rew non_byzantine_nodes_same)) ∘ + (mbind (Some ∘ (same_IM_label_rew non_byzantine_nodes_same))) ∘ (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine). Definition pre_loaded_fixed_non_byzantine_state_projection := @@ -176,40 +176,24 @@ Lemma pre_loaded_fixed_non_byzantine_vlsm_projection pre_loaded_fixed_non_byzantine_label_projection pre_loaded_fixed_non_byzantine_state_projection. Proof. - specialize - (pre_definitions_projection_relation_left_trace - (composite_type (sub_IM fixed_byzantine_IM non_byzantine)) - (composite_type (sub_IM IM non_byzantine)) - (same_IM_label_rew non_byzantine_nodes_same) - (same_IM_state_rew non_byzantine_nodes_same) - (composite_type fixed_byzantine_IM) - (composite_label_sub_projection_option fixed_byzantine_IM non_byzantine) - (composite_state_sub_projection fixed_byzantine_IM non_byzantine)) - as Hrew. - constructor; [constructor|]; intros sX trX HtrX - ; unfold pre_loaded_fixed_non_byzantine_label_projection, pre_loaded_fixed_non_byzantine_state_projection - ; cbn in Hrew |- * ; rewrite Hrew. - - induction trX using rev_ind; [reflexivity|]. - rewrite finite_trace_last_is_last; cbn. - rewrite pre_VLSM_projection_finite_trace_project_app. - unfold pre_VLSM_full_projection_finite_trace_project; rewrite !map_app. - destruct x, l as (i, li); cbn. - unfold pre_VLSM_projection_transition_item_project, pre_loaded_fixed_non_byzantine_label_projection; cbn. - unfold composite_label_sub_projection_option at 2; cbn. - case_decide as Hi; cbn; [rewrite last_is_last; reflexivity|]. - apply finite_valid_trace_from_app_iff in HtrX as [HtrX Hx]. - spec IHtrX HtrX. - inversion Hx; subst; clear Hx Htl. - eapply induced_sub_projection_transition_consistency_None - with (sub_index_list := non_byzantine) in Ht. - + rewrite Ht; cbn in *. - rewrite IHtrX, app_nil_r. - reflexivity. - + unfold composite_label_sub_projection_option; cbn. - case_decide; [contradiction|reflexivity]. - - apply (VLSM_full_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm_full_projection). - eapply (VLSM_projection_finite_valid_trace pre_loaded_fixed_non_byzantine_vlsm'_projection). - eassumption. + apply (@VLSM_projection_composition _ _ _ pre_loaded_fixed_non_byzantine_vlsm _ _ pre_loaded_fixed_non_byzantine_vlsm'_projection), + @VLSM_full_projection_is_projection, pre_loaded_fixed_non_byzantine_vlsm_full_projection. +Qed. + +Definition pre_loaded_fixed_non_byzantine_label_component_projection i := + (mbind (sub_label_element_project IM non_byzantine i)) ∘ + pre_loaded_fixed_non_byzantine_label_projection. + +Definition pre_loaded_fixed_non_byzantine_state_component_projection i (Hi : i ∈ non_byzantine) := + sub_state_element_project IM non_byzantine i Hi ∘ + pre_loaded_fixed_non_byzantine_state_projection. + +Lemma pre_loaded_fixed_non_byzantine_vlsm_component_projection i (Hi : i ∈ non_byzantine) + : VLSM_projection fixed_byzantine_composite_vlsm (pre_loaded_with_all_messages_vlsm (IM i)) + (pre_loaded_fixed_non_byzantine_label_component_projection i) + (pre_loaded_fixed_non_byzantine_state_component_projection i Hi). +Proof. + apply (VLSM_projection_composition pre_loaded_fixed_non_byzantine_vlsm_projection (sub_preloaded_all_component_projection IM non_byzantine i Hi _)). Qed. (** ** Byzantine traces characterization as projections. *) @@ -531,6 +515,47 @@ Proof. destruct byz_tr as ((byz_IM & no_initial_byz & can_emit_byz) & (byz_is, byz_tr) & Hbyz_tr); cbn in *. eexists; split; [apply composite_state_sub_projection_lift|]. eexists; split; [apply composite_trace_sub_projection_lift|]. + assert + (Hinit : composite_initial_state_prop IM + (lift_sub_state IM selection_complement + (pre_loaded_fixed_non_byzantine_state_projection IM selection byz_IM byz_is))). + { destruct Hbyz_tr as [_ Hinit] + ; eapply VLSM_projection_initial_state in Hinit + ; [| apply pre_loaded_fixed_non_byzantine_vlsm_projection]. + apply lift_sub_state_initial; assumption. + } + split; [|assumption]. + induction Hbyz_tr using finite_valid_trace_rev_ind + ; [constructor; apply initial_state_is_valid; assumption|]. + specialize (IHHbyz_tr Hinit). + unfold VLSM_projection_trace_project, pre_VLSM_projection_finite_trace_project; + rewrite map_option_app. + unfold pre_VLSM_full_projection_finite_trace_project; rewrite map_app. + apply finite_valid_trace_from_app_iff; split; [assumption|]; cbn. + remember (finite_trace_last (lift_sub_state _ _ _) _) as lst. + assert (Hlst : valid_state_prop Fixed lst) + by (subst; apply finite_valid_trace_last_pstate; assumption). + destruct l as (i, li); unfold pre_VLSM_projection_transition_item_project, pre_loaded_fixed_non_byzantine_label_projection, composite_label_sub_projection_option; cbn. + case_decide as Hi; [|constructor; assumption]; cbn. + unfold composite_label_sub_projection; cbn. + apply finite_valid_trace_singleton; cbn. + repeat split. + - assumption. + - destruct iom as [im|]; [|apply option_valid_message_None]. + apply set_diff_elim2 in Hi as Hni. + specialize (VLSM_projection_input_valid_transition (pre_loaded_fixed_non_byzantine_vlsm_component_projection IM selection A sender byz_IM i Hi) (existT i li)) as Hproj. + remember (pre_loaded_fixed_non_byzantine_label_component_projection _ _ _ _ _) as somelY. + unfold pre_loaded_fixed_non_byzantine_label_component_projection, pre_loaded_fixed_non_byzantine_label_projection, composite_label_sub_projection_option in HeqsomelY; cbn in HeqsomelY. + case_decide as H_i; [|contradiction]; unfold sub_label_element_project in HeqsomelY; cbn in HeqsomelY. + rewrite decide_True_pi with eq_refl in HeqsomelY; cbn in HeqsomelY; subst somelY. + specialize (Hproj _ eq_refl _ _ _ _ Hx) as [Hv _]. + eapply (Hvalidator i); [assumption|eassumption]. + - destruct Hx as [(_ & _ & Hv & _) _]. + revert Hv; cbn. + clear -Heqlst Hi. + subst lst. revert li. + unfold same_VLSM_label_rew, non_byzantine_nodes_same, fixed_byzantine_IM, update_IM; cbn. + admit. Admitted. diff --git a/theories/VLSM/Core/SubProjectionTraces.v b/theories/VLSM/Core/SubProjectionTraces.v index 0ba9e523..9978ed4a 100644 --- a/theories/VLSM/Core/SubProjectionTraces.v +++ b/theories/VLSM/Core/SubProjectionTraces.v @@ -101,6 +101,15 @@ Definition composite_label_sub_projection := existT (dexist i e) (projT2 l). +Lemma composite_label_sub_projection_pi + (l : composite_label IM) + (i := projT1 l) + (e e' : sub_index_prop i) + : composite_label_sub_projection l e = composite_label_sub_projection l e'. +Proof. + apply (@dec_sig_sigT_eq _ _ sub_index_prop_dec (fun i => vlabel (IM i))); reflexivity. +Qed. + Definition lift_sub_label (l : composite_label sub_IM) : composite_label IM @@ -1844,6 +1853,54 @@ Proof. apply sub_transition_element_project_Some. Qed. +Lemma sub_preloaded_component_projection constraint P + : VLSM_projection + (pre_loaded_vlsm (composite_vlsm (sub_IM IM indices) constraint) P) + (pre_loaded_with_all_messages_vlsm (IM j)) + sub_label_element_project sub_state_element_project. +Proof. + apply basic_VLSM_projection. + - intros (sub_i, lXi) lY HlX_pr; destruct_dec_sig sub_i i Hi Heqsub_i; subst. + unfold sub_label_element_project in HlX_pr; cbn in HlX_pr. + case_decide as Hij; [|congruence]; apply Some_inj in HlX_pr; subst; cbn. + intros s om (_ & _ & Hv & _) _ _; revert Hv; cbn. + replace (s (dexist _ _)) with (sub_state_element_project s) + by apply sub_IM_state_pi. + auto. + - intros (sub_i, lXi) lY HlX_pr; destruct_dec_sig sub_i i Hi Heqsub_i; subst. + unfold sub_label_element_project in HlX_pr; cbn in HlX_pr. + case_decide as Hij; [|congruence]; apply Some_inj in HlX_pr; subst; cbn. + intros * [_ Ht]; revert Ht; cbn. + replace (s (dexist _ _)) with (sub_state_element_project s) + by apply sub_IM_state_pi. + destruct (vtransition _ _ _); inversion_clear 1; f_equal. + symmetry; apply sub_IM_state_update_eq. + - intros (sub_i, lXi) HlX_pr; destruct_dec_sig sub_i i Hi Heqsub_i; subst. + unfold sub_label_element_project in HlX_pr; cbn in HlX_pr. + case_decide as Hij; [congruence|]; clear HlX_pr. + intros * [_ Ht]; revert Ht; cbn. + destruct (vtransition _ _ _); inversion_clear 1. + apply sub_IM_state_update_neq; congruence. + - intros s Hs; apply (Hs (dexist j Hj)). + - intro; intros; apply any_message_is_valid_in_preloaded. +Qed. + +Lemma sub_preloaded_all_component_projection constraint + : VLSM_projection + (pre_loaded_with_all_messages_vlsm (composite_vlsm (sub_IM IM indices) constraint)) + (pre_loaded_with_all_messages_vlsm (IM j)) + sub_label_element_project sub_state_element_project. +Proof. + constructor; [constructor|]; cbn; intros sX trX HtrX. + - apply (@final_state_project message (pre_loaded_vlsm (composite_vlsm (sub_IM IM indices) constraint) (fun m => true))). + + apply (projection_type _ _ _ _ (sub_preloaded_component_projection constraint (fun m => True))). + + revert HtrX; apply VLSM_incl_finite_valid_trace_from. + apply pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True_l. + - apply (VLSM_projection_finite_valid_trace (sub_preloaded_component_projection constraint (fun m => True))). + revert HtrX; apply VLSM_incl_finite_valid_trace. + apply pre_loaded_with_all_messages_vlsm_is_pre_loaded_with_True_l. +Qed. + End sub_composition_element. Section sub_composition_preloaded_lift. @@ -2020,6 +2077,15 @@ Context (selection_complement : set index := set_diff (enum index) selection) . +Lemma selection_complement_index_not_in_selection + (sub_i : sub_index selection_complement) + : ~ (`sub_i ∈ selection). +Proof. + destruct_dec_sig sub_i i Hi Heqsub_i; subst; cbn in *. + eapply set_diff_elim2; eassumption. +Qed. + + Global Instance update_IM_complement_Hbs `{forall i : index, HasBeenSentCapability (IM i)} : forall sub_i : sub_index selection_complement, diff --git a/theories/VLSM/Core/VLSMProjections.v b/theories/VLSM/Core/VLSMProjections.v index c6a48466..4f32e57d 100644 --- a/theories/VLSM/Core/VLSMProjections.v +++ b/theories/VLSM/Core/VLSMProjections.v @@ -1468,7 +1468,7 @@ Definition VLSM_full_projection_finite_valid_trace to lift to VLSM full projections the generic results proved about VLSM projections. *) Lemma VLSM_full_projection_is_projection - : VLSM_projection X Y (fun l => Some (label_project l)) state_project. + : VLSM_projection X Y (Some ∘ label_project) state_project. Proof. split. - apply VLSM_full_projection_projection_type. @@ -3298,3 +3298,57 @@ Proof. Qed. End same_VLSM_full_projection. + +Section VLSM_projection_composition. + +Context + {message} + [X Y Z : VLSM message] + `(HXY : VLSM_projection X Y xy_label xy_state) + `(HYZ : VLSM_projection Y Z yz_label yz_state) + . + +Lemma pre_VLSM_projection_composition_transition_item_project + : pre_VLSM_projection_transition_item_project (type X) (type Z) + (mbind yz_label ∘ xy_label) (yz_state ∘ xy_state) + = + mbind (pre_VLSM_projection_transition_item_project (type Y) (type Z) yz_label yz_state) + ∘ (pre_VLSM_projection_transition_item_project (type X) (type Y) xy_label xy_state). +Proof. + extensionality item; destruct item; cbn. + unfold pre_VLSM_projection_transition_item_project; cbn. + destruct (xy_label l) as [lY|]; reflexivity. +Qed. + +Lemma pre_VLSM_projection_composition_finite_trace_project + : pre_VLSM_projection_finite_trace_project (type X) (type Z) + (mbind yz_label ∘ xy_label) (yz_state ∘ xy_state) + = + pre_VLSM_projection_finite_trace_project (type Y) (type Z) yz_label yz_state + ∘ pre_VLSM_projection_finite_trace_project (type X) (type Y) xy_label xy_state. +Proof. + unfold pre_VLSM_projection_finite_trace_project. + setoid_rewrite pre_VLSM_projection_composition_transition_item_project. + setoid_rewrite map_option_comp_bind. + reflexivity. +Qed. + +Lemma VLSM_projection_composition + : VLSM_projection X Z (mbind yz_label ∘ xy_label) (yz_state ∘ xy_state). +Proof. + constructor; [constructor|]; intros sX trX HtrX + ; replace (pre_VLSM_projection_finite_trace_project _ _ _ _) + with + ((pre_VLSM_projection_finite_trace_project (type Y) (type Z) yz_label yz_state + ∘ pre_VLSM_projection_finite_trace_project (type X) (type Y) xy_label xy_state)) + by (symmetry; apply pre_VLSM_projection_composition_finite_trace_project). + - unfold compose. + erewrite final_state_project; [|eapply projection_type; eassumption|assumption]. + eapply final_state_project; [apply projection_type; assumption|]. + apply VLSM_projection_finite_valid_trace_from with (Hsimul := HXY); assumption. + - apply VLSM_projection_finite_valid_trace with (Hsimul := HYZ), + VLSM_projection_finite_valid_trace with (Hsimul := HXY). + assumption. +Qed. + +End VLSM_projection_composition. \ No newline at end of file diff --git a/theories/VLSM/Lib/ListExtras.v b/theories/VLSM/Lib/ListExtras.v index 60bbc145..807304bf 100644 --- a/theories/VLSM/Lib/ListExtras.v +++ b/theories/VLSM/Lib/ListExtras.v @@ -1261,6 +1261,17 @@ Proof. destruct (f a) as [b|]; reflexivity. Qed. +Lemma map_option_comp_bind + [A B C : Type] + (f : A -> option B) + (g : B -> option C) + : map_option (mbind g ∘ f) = map_option g ∘ map_option f. +Proof. + extensionality l. + induction l; [reflexivity|]; cbn; rewrite IHl. + destruct (f a) as [b|]; reflexivity. +Qed. + Lemma map_option_app {A B : Type} (f : A -> option B)