diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index a8d9ef50d2ec668b231cba6cd03d7f039d2526ff..755316c24590ce0b380c3e0f524da4aeb72a7a16 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -9,7 +9,8 @@ stages: stage: build image: coqorg/${CI_JOB_NAME} before_script: - - if [ -n "${COMPILER_EDGE}" ]; then opam switch ${COMPILER_EDGE} && eval $(opam env); fi + - if [ -n "${COMPILER_EDGE}" ]; then opam switch ${COMPILER_EDGE}; fi + - eval $(opam env) - opam update -y - opam config list - opam repo list @@ -20,8 +21,11 @@ stages: - make -j "$NJOBS" # - make install -coq:8.12: +coq:8.18: extends: .build -coq:8.13: +coq:8.19: + extends: .build + +coq:8.20: extends: .build diff --git a/CREDIT.md b/CREDIT.md new file mode 100644 index 0000000000000000000000000000000000000000..dfd9aa608b6e2deaffa29524ca38057fcb4aa21e --- /dev/null +++ b/CREDIT.md @@ -0,0 +1,11 @@ +The following people have contrinuted to the Pactole library. See also +the LICENSE file. + +Cédric Auger (2013) +Thibaut Balabonski (2016-2017) +Sebastien Bouchard (2021-) +Pierre Courtieu (2013-) +Robin Pelle (2016-2020) +Lionel Rieg (2014-) +Xavier Urbain (2013-) + diff --git a/CaseStudies/Convergence/Algorithm_noB.v b/CaseStudies/Convergence/Algorithm_noB.v index d814c44f7380968b988f79c98ec875ac91ca5723..48c4ed1bce14c65c481e954e9cf24f8ed308c233 100644 --- a/CaseStudies/Convergence/Algorithm_noB.v +++ b/CaseStudies/Convergence/Algorithm_noB.v @@ -24,6 +24,8 @@ Require Import SetoidDec. Require Import Lia. Require Import SetoidList. Require Import Pactole.Util.Preliminary. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Util.Fin. Require Import Pactole.Setting. Require Import Pactole.Spaces.R2. Require Import Pactole.Observations.SetObservation. @@ -39,13 +41,12 @@ Typeclasses eauto := (bfs). Section ConvergenceAlgo. -(** There are [n] good robots and no byzantine one. *) -Variable n : nat. -Hypothesis n_non_0 : n <> 0%nat. +(** There are [ub] good robots and no byzantine one. *) +Context {k : nat} {ltc_0_k : 0 <c k}. -Instance MyRobots : Names := Robots n 0. +Instance MyRobots : Names := Robots k 0. Instance NoByz : NoByzantine. -Proof. now split. Qed. +Proof using . now split. Qed. (* BUG?: To help finding correct instances, loops otherwise! *) Instance Loc : Location := {| location := R2 |}. @@ -89,11 +90,10 @@ Implicit Type pt : location. (** As there are robots, the observation can never be empty. *) Lemma obs_non_empty : forall config pt, obs_from_config config pt =/= @empty location _ _ _. -Proof using n_non_0. +Proof using ltc_0_k. intros config pt. rewrite obs_from_config_ignore_snd. intro Habs. -assert (Hn : 0%nat < n). { generalize n_non_0. intro. lia. } -pose (g := exist _ 0%nat Hn : G). +pose (g := fin0 : G). specialize (Habs (config (Good g))). rewrite empty_spec in Habs. assert (Hin := pos_in_config config origin (Good g)). @@ -139,16 +139,16 @@ Proof using . intros ? ? Heq ? ? ?. now apply Stream.eventually_compat, imprison no matter what the demon and the starting configuration are. *) Definition solution_SSYNC (r : robogram) : Prop := forall (config : configuration) (d : similarity_demon), Fair d → - forall (ε : R), 0 < ε → exists (pt : R2), attracted pt ε (execute r d config). + exists (pt : R2), forall (ε : R), 0 < ε → attracted pt ε (execute r d config). Definition solution_FSYNC (r : robogram) : Prop := forall (config : configuration) (d : similarity_demon), FSYNC (similarity_demon2demon d) → - forall (ε : R), 0 < ε → exists (pt : R2), attracted pt ε (execute r d config). + exists (pt : R2), forall (ε : R), 0 < ε → attracted pt ε (execute r d config). Lemma synchro : ∀ r, solution_SSYNC r → solution_FSYNC r. Proof using . unfold solution_SSYNC. intros r Hfair config d Hd. -apply Hfair, FSYNC_implies_fair; autoclass. +apply Hfair, FSYNC_implies_Fair; autoclass. Qed. Close Scope R_scope. @@ -170,7 +170,7 @@ Theorem round_simplify : forall da config, SSYNC_da da -> == fun id => if da.(activate) id then isobarycenter (@elements location _ _ _ (!! config)) else config id. -Proof using n_non_0. +Proof using ltc_0_k. intros da config HSSYNC. rewrite SSYNC_round_simplify; trivial; []. intro id. pattern id. apply no_byz. clear id. intro g. unfold round. destruct_match; try reflexivity; []. @@ -196,18 +196,19 @@ Axiom isobarycenter_circle : forall center radius (l : list R2), Lemma contained_isobarycenter : forall c r config, contained c r config -> (dist c (isobarycenter (elements (!! config))) <= r)%R. -Proof using n_non_0. +Proof using . intros c r config Hc. apply isobarycenter_circle. rewrite Forall_forall. intro. rewrite <- InA_Leibniz. change eq with (@equiv location _). rewrite elements_spec, obs_from_config_In. -intros [id Hpt]. rewrite <- Hpt. +intros [id Hpt]. Fail timeout 10 rewrite <- Hpt. (* FIXME: the rewrite should not fail *) +apply (@dist_compat location _ _ _ _ _ _ (reflexivity c)) in Hpt. rewrite <- Hpt. pattern id. apply no_byz. apply Hc. Qed. Lemma contained_next : forall da c r config, SSYNC_da da -> contained c r config -> contained c r (round convergeR2 da config). -Proof using n_non_0. +Proof using ltc_0_k. intros da c r config HSSYNC Hconfig g. rewrite round_simplify; trivial; []. destruct_match. @@ -217,7 +218,7 @@ Qed. Lemma converge_forever : forall d c r config, SSYNC d -> contained c r config -> imprisoned c r (execute convergeR2 d config). -Proof using n_non_0. +Proof using ltc_0_k. cofix Hcorec. intros d c r config [] Hrec. constructor. - apply Hrec. - apply Hcorec; auto using contained_next. @@ -229,13 +230,14 @@ Qed. (************************) Theorem convergence_FSYNC : solution_FSYNC convergeR2. -Proof using n_non_0. -intros config d [Hfair ?] ε Hε. +Proof using ltc_0_k. +intros config d [Hfair ?]. exists (isobarycenter (elements (obs_from_config (Observation := set_observation) config 0))). +intros ε Hε. apply Stream.Later, Stream.Now. rewrite execute_tail. apply converge_forever; auto using FSYNC_SSYNC; []. intro g. rewrite round_simplify; auto using FSYNC_SSYNC_da; []. -rewrite Hfair. changeR2. +hnf in Hfair. rewrite Hfair. changeR2. transitivity 0%R; try (now apply Rlt_le); []. apply Req_le. now apply dist_defined. Qed. diff --git a/CaseStudies/Convergence/Impossibility_2G_1B.v b/CaseStudies/Convergence/Impossibility_2G_1B.v index 70e0fa5f4adf1867d4db3c00a3c0759e3262961c..1c57cf28dfba680daff179ea0f204534b07a26c0 100644 --- a/CaseStudies/Convergence/Impossibility_2G_1B.v +++ b/CaseStudies/Convergence/Impossibility_2G_1B.v @@ -22,10 +22,12 @@ Require Import Utf8. Require Import Reals. Require Import Psatz. Require Import SetoidDec. -Require Import Arith.Div2. Require Import Lia. Require Import SetoidList. Require Import Pactole.Util.Preliminary. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Util.Fin. +Require Import Pactole.Util.Enum. Require Import Pactole.Setting. Require Import Pactole.Spaces.R. Require Import Pactole.Observations.MultisetObservation. @@ -38,9 +40,8 @@ Import Datatypes. Import List. Import SetoidClass. - -Typeclasses eauto := (bfs). - +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Section ConvergenceImpossibility. (** There are [2 * n] good robots and [n] byzantine ones. *) @@ -132,22 +133,26 @@ Proof using n_non_0. assert (Hn0 := n_non_0). rewrite nG_nB, nB_value. destruct n as [| ?]. - lia. -- simpl. rewrite plus_comm. simpl. lia. +- cbn. rewrite Nat.add_comm. cbn. auto with arith. Qed. - +(* TODO: move to a definition/problem file *) (** ** Properties of executions *) Open Scope R_scope. -(** Expressing that all good robots are confined in a small disk. *) +(** All good robots are in a small disk. *) +Definition close (center : R) (radius : R) (config : configuration) := + forall g, dist center (get_location (config (Good g))) <=radius. + +(** All good robots stay confined in a small disk. *) Definition imprisoned (center : R) (radius : R) (e : execution) : Prop := - Stream.forever (Stream.instant - (fun config => forall g, dist center (get_location (config (Good g))) <=radius)) e. + Stream.forever (Stream.instant (close center radius)) e. (** The execution will end in a small disk. *) Definition attracted (c : R) (r : R) (e : execution) : Prop := Stream.eventually (imprisoned c r) e. +(* TODO: split for close *) Instance imprisoned_compat : Proper (equiv ==> Logic.eq ==> equiv ==> iff) imprisoned. Proof using . intros c1 c2 Hc ? r Hr e1 e2 He. subst. split. @@ -171,18 +176,17 @@ Proof using . intros ? ? Heq ? ? ?. now apply Stream.eventually_compat, imprison (** A solution is just convergence property for any demon. *) Definition solution (r : robogram) : Prop := forall (config : configuration) (d : demon), Fair d → - forall (ε : R), 0 < ε → exists (pt : R), attracted pt ε (execute r d config). + exists (pt : R), forall (ε : R), 0 < ε → attracted pt ε (execute r d config). Definition solution_FSYNC (r : robogram) : Prop := forall (config : configuration) (d : demon), FSYNC d → - forall (ε : R), 0 < ε → exists (pt : R), attracted pt ε (execute r d config). + exists (pt : R), forall (ε : R), 0 < ε → attracted pt ε (execute r d config). (** A SSYNC solution is also a FSYNC solution. *) Lemma synchro : ∀ r, solution r → solution_FSYNC r. -Proof using . unfold solution. intros r Hfair config d Hd. apply Hfair, FSYNC_implies_fair; autoclass. Qed. +Proof using . unfold solution. intros r Hfair config d Hd. apply Hfair, FSYNC_implies_Fair; autoclass. Qed. Close Scope R_scope. -Close Scope vector_scope. (** We split good robots into two halves. *) @@ -210,18 +214,17 @@ eapply firstn_skipn_nodup_exclusive; try eassumption. apply Gnames_NoDup. Qed. -Lemma left_spec : forall g, In g left <-> proj1_sig g < Nat.div2 nG. +Lemma left_spec : forall g, In g left <-> fin2nat g < Nat.div2 nG. Proof using . Local Transparent G. -intro g. unfold left, half1. rewrite Gnames_length. unfold Gnames. -apply firstn_enum_spec. +intro g. unfold left, half1. rewrite Gnames_length. apply firstn_enum_spec. Local Opaque G. Qed. -Lemma right_spec : forall g, In g right <-> Nat.div2 nG <= proj1_sig g. +Lemma right_spec : forall g, In g right <-> Nat.div2 nG <= fin2nat g. Proof using . intro g. unfold right, half2. rewrite Gnames_length. unfold Gnames. -rewrite (skipn_enum_spec (Nat.div2 nG) g). intuition. apply proj2_sig. +rewrite (skipn_enum_spec (Nat.div2 nG) g). intuition. apply fin_lt. Qed. (** First and last robots are resp. in the first and in the second half. *) @@ -238,7 +241,7 @@ Lemma glast_right : In glast right. Proof using . rewrite right_spec. simpl. assert (Heven := even_nG). destruct n as [| [| ]]; simpl; auto; []. -apply le_n_S, Nat.div2_decr, le_n_Sn. +apply le_n_S, Nat.div2_decr, Nat.le_succ_diag_r. Qed. Hint Resolve gfirst_left glast_right left_right_exclusive : core. @@ -252,7 +255,6 @@ Hint Resolve gfirst_left glast_right left_right_exclusive : core. - the stack with byzantine is activated, good robots cannot move. *) Open Scope R_scope. -Open Scope vector_scope. (** The reference starting configuration **) Definition config1 : configuration := fun id => @@ -304,7 +306,7 @@ intros pt1 pt2 Hdiff pt k. induction k as [| k]; intros l Hnodup Hlen. cbn [map]. rewrite map_app. cbn [map]. destruct (in_dec Geq_dec a (a :: half1 l)) as [_ | Habs]. destruct (in_dec Geq_dec z (a :: half1 l)) as [Habs | _]. - + inversion_clear Habs; try (now elim Haz); []. + + inversion_clear Habs; try (now contradiction Haz); []. exfalso. now apply Hzl, half1_incl. + rewrite (map_ext_in _ (fun x => if in_dec Geq_dec x (half1 l) then pt1 else pt2)). - cbn [countA_occ]. rewrite countA_occ_app. rewrite IHk; trivial. @@ -316,9 +318,9 @@ intros pt1 pt2 Hdiff pt k. induction k as [| k]; intros l Hnodup Hlen. - intros x Hin. destruct (in_dec Geq_dec x (half1 l)) as [Hx | Hx], (in_dec Geq_dec x (a :: half1 l)) as [Hx' | Hx']; trivial. - -- elim Hx'. intuition. + -- contradiction Hx'. intuition. -- inversion_clear Hx'; subst; contradiction. - + elim Habs. intuition. + + contradiction Habs. intuition. Qed. Theorem obs_config1 : !! config1 == observation1. @@ -376,8 +378,7 @@ intro pt. unfold observation1, observation2, swap. rewrite map_add, map_singleto cbn -[add singleton]. ring_simplify (1 + -1 * (0 + -(1)) + -(1)). ring_simplify (1 + -1 * (1 + -(1)) + -(1)). -destruct (Rdec pt 0); [| destruct (Rdec pt 1)]; subst; -repeat rewrite ?add_same, ?singleton_same, ?singleton_other, ?add_other; auto. +now rewrite <- add_singleton_comm. Qed. (** An execution alternating config1 and config2 *) @@ -516,7 +517,7 @@ CoFixpoint shifting_demon d pt := Lemma Fair_shifting_demon : forall d pt, Fair (shifting_demon d pt). Proof using . -intros d pt. apply FSYNC_implies_fair; autoclass; []. revert pt. +intros d pt. apply FSYNC_implies_Fair; autoclass; []. revert pt. cofix shift_fair. intro pt. constructor. + repeat intro. simpl. reflexivity. + cbn. apply shift_fair. @@ -596,8 +597,8 @@ Proof using . intro. apply keep_moving_by_eq. reflexivity. Qed. Theorem absurd : False. Proof using absurdmove n_non_0 sol. assert (Hthird_move : 0 < Rabs (move / 3)). { apply Rabs_pos_lt. lra. } -specialize (sol (config0 0) (Fair_shifting_demon move 0) Hthird_move). -destruct sol as [pt Hpt]. rewrite keep_moving in Hpt. +destruct (sol (config0 0) (Fair_shifting_demon move 0)) as [pt Hpt]. +specialize (Hpt _ Hthird_move). rewrite keep_moving in Hpt. remember (shifting_execution move 0) as e. remember (Rabs (move / 3)) as ε. revert Heqe. generalize 0. induction Hpt as [e IHpt | e IHpt]; intros start Hstart. @@ -605,10 +606,10 @@ induction Hpt as [e IHpt | e IHpt]; intros start Hstart. clear -absurdmove Hnow1 Hnow2 n_non_0. specialize (Hnow1 gfirst). specialize (Hnow2 gfirst). cut (Rabs move <= Rabs (move / 3) + Rabs (move / 3)). - assert (Hpos : 0 < Rabs move) by now apply Rabs_pos_lt. - unfold Rdiv. rewrite Rabs_mult, Rabs_Rinv; try lra. + unfold Rdiv. rewrite Rabs_mult, Rabs_inv; try lra. assert (Habs3 : Rabs 3 = 3). { apply Rabs_pos_eq. lra. } rewrite Habs3 in *. lra. - - rewrite sqrt_square in Hnow1, Hnow2. + - simpl in *. rewrite sqrt_square in Hnow1, Hnow2. replace move with ((pt - start) - (pt - (start + move))) at 1 by ring. unfold Rminus at 1. eapply Rle_trans; try (now apply Rabs_triang); []. apply Rplus_le_compat; trivial; []. now rewrite Rabs_Ropp. @@ -698,8 +699,8 @@ Theorem noConvergence : forall r, ~solution r. Proof using n_non_0. intros r Hr. assert (Hpos : 0 < 1/4) by lra. -destruct (Hr config1 bad_demon Fair_bad_demon _ Hpos) as [pt Hpt]. -rewrite execute_bad_demon in Hpt; trivial. +destruct (Hr config1 bad_demon Fair_bad_demon) as [pt Hpt]. +specialize (Hpt _ Hpos). rewrite execute_bad_demon in Hpt; trivial. revert Hpt. cut (exec == exec); try reflexivity. generalize exec at -2. intros e He Hpt. induction Hpt using attracted_ind2. + (* First step *) diff --git a/CaseStudies/Exploration/Definitions.v b/CaseStudies/Exploration/Definitions.v deleted file mode 100644 index 31a4febb4cebf6742f4a47df67f5ba4fcdc93ea5..0000000000000000000000000000000000000000 --- a/CaseStudies/Exploration/Definitions.v +++ /dev/null @@ -1,182 +0,0 @@ -(**************************************************************************) -(* Mechanised Framework for Local Interactions & Distributed Algorithms *) -(* C. Auger, P. Courtieu, L. Rieg, X. Urbain , R. Pelle *) -(* PACTOLE project *) -(* *) -(* This file is distributed under the terms of the CeCILL-C licence *) -(* *) -(**************************************************************************) - - -Require Import Reals Lia Arith.Div2 Psatz SetoidDec. -Require Export Pactole.Setting. -Require Export Pactole.Spaces.Ring. -Require Export Pactole.Spaces.Isomorphism. -Require Export Pactole.Observations.MultisetObservation. - - -Close Scope Z_scope. -Set Implicit Arguments. -Typeclasses eauto := (bfs). - - -Section ExplorationDefs. - -(** Setting definitions *) - - -(** Definition of the ring. *) -Context {RR : RingSpec}. -(* We do not care about threshold values, so we just take 1/2 everywhere. *) -Existing Instance localRing. -Notation ring_node := (finite_node ring_size). -(* NB: These instances will be replaced by the glob_* ones so they are local. *) - -(* begin show *) -(** Number of good and Byzantine robots *) -Context {Robots : Names}. - -(** Robots are on nodes *) -Global Instance Loc : Location := make_Location ring_node. - -(** States of robots only contains their location. *) -Global Instance St : State location := OnlyLocation (fun _ => True). -Global Existing Instance proj_graph. - -(** Robots observe the location of others robots with strong multiplicity. *) -Global Existing Instance multiset_observation. - -(** Robots only decide in which direction they want to move. *) -Global Instance RC : robot_choice direction := { robot_choice_Setoid := direction_Setoid }. - -(** Demon's frame choice: we move back the robot to the origin with a translation - and we can choose the orientation of the ring. *) -Global Instance FC : frame_choice (Z * bool) := { - frame_choice_bijection := - fun nb => if snd nb then Ring.sym (fst nb) else Ring.trans (fst nb); - frame_choice_Setoid := eq_setoid _ }. - -Global Existing Instance NoChoice. -Global Existing Instance NoChoiceIna. -Global Existing Instance NoChoiceInaFun. - -Global Instance UpdFun : update_function direction (Z * bool) unit := { - update := fun config g _ dir _ => move_along (config (Good g)) dir; - update_compat := ltac:(repeat intro; subst; now apply move_along_compat) }. -(* end show *) - -(* Global Instance setting : GlobalDefinitions := { - (* Number of good and Byzantine robots *) - glob_Names := Robots; - (* The space in which robots evolve *) - glob_Loc := Loc; - (* The state of robots (must contain at least the current location) *) - glob_info := location; - glob_State := OnlyLocation (fun _ => True); - (* The observation: what robots can see from their surroundings *) - glob_obs := multiset_observation; - (* The output type of robograms: some information that we can use to get a target location *) - glob_Trobot := direction; - glob_robot_choice := RC; - (* The frame decision made by the demon, used to build the frame change *) - glob_Tframe := Z * bool; - glob_frame_choice := FC; - (* The influence of the demon on the state update of robots, when active *) - glob_Tactive := unit; - glob_update_choice := NoChoice; - (* The influence of the demon on the state update of robots, when inactive *) - glob_Tinactive := unit; - glob_inactive_choice := NoChoiceIna; - (* How a robots state is updated: - - if active, using the result of the robogram and the decision from the demon - - if inactive, using only the decision from the demon *) - glob_update_function := UpdFun; - glob_inactive_function := NoChoiceInaFun }. *) - -(** ** Specification of exploration with stop *) - -(** Any node will eventually be visited. *) -Definition is_visited (pt : location) (config : configuration) := - exists g, config (Good g) == pt. - -Definition Will_be_visited (pt : location) (e : execution) : Prop := - Stream.eventually (Stream.instant (is_visited pt)) e. - -(** Eventually, all robots stop moving. *) -Definition Stall (e : execution) := Stream.hd e == (Stream.hd (Stream.tl e)). - -Definition Stopped (e : execution) : Prop := - Stream.forever Stall e. - -Definition Will_stop (e : execution) : Prop := - Stream.eventually Stopped e. - -(** [Exploration_with_stop r d config] means that executing [r] against demon [d] from - configuration [config] indeed solves exploration with stop: after a finite time, every - node of the space has been visited and all robots will stay at the same place forever. *) -Definition ExplorationWithStop (r : robogram) (d : demon) (config : configuration) := - (forall l, Will_be_visited l (execute r d config)) - /\ Will_stop (execute r d config). - -(** [FullSolExplorationWithStop r d] means that the robogram [r] solves exploration with stop - agains demon [d] regardless of the starting configuration. - - This is actually impossible when the number of robots is less than the size of the ring. *) -Definition FullSolExplorationWithStop (r : robogram) (d : demon) := - forall config, ExplorationWithStop r d config. - -(** Acceptable starting configurations contain no tower, - that is, all robots are at different locations. *) -Definition Valid_starting_config config : Prop := - Util.Preliminary.injective (@Logic.eq ident) (@equiv _ state_Setoid) config. -(* forall pt : location, ((obs_from_config config (of_Z 0))[pt] <= 1)%nat. *) - -Definition Explore_and_Stop (r : robogram) := - forall d config, Fair d -> Valid_starting_config config -> - ExplorationWithStop r d config. - -(** Compatibility properties *) -Global Instance is_visited_compat : Proper (equiv ==> equiv ==> iff) is_visited. -Proof using . -intros pt1 pt2 Hpt config1 config2 Hconfig. -split; intros [g Hv]; exists g. -- now rewrite <- Hconfig, Hv, Hpt. -- now rewrite Hconfig, Hv, Hpt. -Qed. - -Global Instance Will_be_visited_compat : Proper (equiv ==> equiv ==> iff) Will_be_visited. -Proof using . -intros ? ? ?. now apply Stream.eventually_compat, Stream.instant_compat, is_visited_compat. -Qed. - -Global Instance Stall_compat : Proper (equiv ==> iff) Stall. -Proof using . -intros e1 e2 He. split; intros Hs; unfold Stall in *; -(now rewrite <- He) || now rewrite He. -Qed. - -Global Instance Stopped_compat : Proper (equiv ==> iff) Stopped. -Proof using . -intros e1 e2 He. split; revert e1 e2 He; coinduction rec. -- destruct H. now rewrite <- He. -- destruct H as [_ H], He as [_ He]. apply (rec _ _ He H). -- destruct H. now rewrite He. -- destruct H as [_ H], He as [_ He]. apply (rec _ _ He H). -Qed. - -Global Instance Will_stop_compat : Proper (equiv ==> iff) Will_stop. -Proof using . apply Stream.eventually_compat, Stopped_compat. Qed. - -Global Instance Valid_starting_config_compat : Proper (equiv ==> iff) Valid_starting_config. -Proof using . -intros ? ? Hconfig. -unfold Valid_starting_config, Util.Preliminary.injective. -now setoid_rewrite Hconfig. -Qed. - -(** We can decide if a given configuration is a valid starting one. *) -Lemma Valid_starting_config_dec : forall config, - {Valid_starting_config config} + {~ Valid_starting_config config}. -Proof using . intro config. unfold Valid_starting_config. apply config_injective_dec. Qed. - -End ExplorationDefs. diff --git a/CaseStudies/Exploration/ExplorationDefs.v b/CaseStudies/Exploration/ExplorationDefs.v new file mode 100644 index 0000000000000000000000000000000000000000..0229171cbc4ad13e37329da963cecaf786eb99cf --- /dev/null +++ b/CaseStudies/Exploration/ExplorationDefs.v @@ -0,0 +1,104 @@ +Require Export Pactole.Setting. +Require Import Pactole.Util.Stream. + +Set Implicit Arguments. + +Section ExplorationDefs. + +Context {Robots : Names} {Loc : Location}. +Context {Tframe Trobot Tinfo Tactive Tinactive: Type}. +Context {RC : robot_choice Trobot} {FC : frame_choice Tframe}. +Context {UC : update_choice Tactive} {IC : inactive_choice Tinactive}. +Context {St : State Tinfo} {Obs : Observation}. +Context {UpdFun : update_function Trobot Tframe Tactive}. +Context {InaFun : inactive_function Tinactive}. + +(** ** Specification of exploration with stop *) + +Definition is_visited (l : location) (c : configuration) : Prop := + exists g, get_location (c (Good g)) == l. + +(** Any node will eventually be visited. *) +Definition Will_be_visited (l : location) (e : execution) : Prop := + eventually (instant (is_visited l)) e. + +Definition simple_exploration (e : execution) := + forall pt, Will_be_visited pt e. + +Definition perpetual_exploration (e : execution) := + Stream.forever simple_exploration e. + +(** Eventually, all robots stop moving. *) +Definition Stall (e : execution) := hd e == (hd (tl e)). + +Definition Stopped (e : execution) : Prop := + forever Stall e. + +Definition Will_stop (e : execution) : Prop := + eventually Stopped e. + +(** [ExplorationStop e] means that after a finite time, every node of the space has been + visited, and after that time, all robots will stay at the same place forever. *) +Definition ExplorationStop (e : execution) := + simple_exploration e /\ Will_stop e. + +Global Instance is_visited_compat : + Proper (equiv ==> equiv ==> iff) is_visited. +Proof using . + intros l1 l2 Hl c1 c2 Hc. unfold is_visited. split; intros [g H]. + - exists g. rewrite <- Hl, <- Hc. exact H. + - exists g. rewrite Hl, Hc. exact H. +Qed. + +Global Instance Will_be_visited_compat : + Proper (equiv ==> equiv ==> iff) Will_be_visited. +Proof using . + intros l1 l2 Hl e1 e2 He. unfold Will_be_visited. split; intros H. + - rewrite <- He, <- Hl. exact H. + - rewrite He, Hl. exact H. +Qed. + +Global Instance simple_exploration_compat : + Proper (equiv ==> iff) simple_exploration. +Proof using . + intros e1 e2 He. unfold simple_exploration. now setoid_rewrite He. +Qed. + +Global Instance perpetual_exploration_compat : + Proper (equiv ==> iff) perpetual_exploration. +Proof using . apply Stream.forever_compat, simple_exploration_compat. Qed. + +Global Instance Stall_compat : Proper (equiv ==> iff) Stall. +Proof using . + intros e1 e2 He. unfold Stall. split; intros H. + - rewrite <- He. exact H. + - rewrite He. exact H. +Qed. + +Global Instance Stopped_compat : Proper (equiv ==> iff) Stopped. +Proof using . + intros e1 e2 He. unfold Stopped. split; intros H. + - rewrite <- He. exact H. + - rewrite He. exact H. +Qed. + +Global Instance Will_stop_compat : Proper (equiv ==> iff) Will_stop. +Proof using . + intros e1 e2 He. unfold Will_stop. split; intros H. + - rewrite <- He. exact H. + - rewrite He. exact H. +Qed. + +Global Instance ExplorationStop_compat : + Proper (equiv ==> iff) ExplorationStop. +Proof using . + intros e1 e2 He. unfold ExplorationStop. split; intros [Hx Hs]. + - split. + + intros l. rewrite <- He. exact (Hx l). + + rewrite <- He. exact Hs. + - split. + + intros l. rewrite He. exact (Hx l). + + rewrite He. exact Hs. +Qed. + +End ExplorationDefs. diff --git a/CaseStudies/Exploration/ImpossibilityKDividesN.v b/CaseStudies/Exploration/ImpossibilityKDividesN.v index 90400cc2457f3fda28841017e9ab1a0600d117e7..0a0269cf42aa284c7d33c79b6444cc8a337a27d7 100644 --- a/CaseStudies/Exploration/ImpossibilityKDividesN.v +++ b/CaseStudies/Exploration/ImpossibilityKDividesN.v @@ -7,50 +7,67 @@ (* *) (**************************************************************************) - -Require Import Psatz Rbase. -Require Import Morphisms. -Require Import Arith.Div2. -Require Import Lia. -Require Import Decidable. -Require Import Equalities. -Require Import List Setoid SetoidList Compare_dec Morphisms. +Require Import Utf8. +Require Import Arith Lia. +Require Import SetoidList. +Require Import Pactole.Util.Stream. Require Import Pactole.Models.NoByzantine. -Require Import Pactole.CaseStudies.Exploration.Definitions. - +Require Import Pactole.Models.RingSSync. +Require Import Pactole.CaseStudies.Exploration.ExplorationDefs. -Open Scope Z_scope. Set Implicit Arguments. -Typeclasses eauto := (bfs). - Section Exploration. (** Given an abitrary ring *) -Context {RR : RingSpec}. -(** There are kG good robots and no byzantine ones. *) -Variable kG : nat. -Instance Robots : Names := Robots kG 0. +Context {n : nat} {ltc_2_n : 2 <c n}. +(** There are k good robots and no byzantine ones. *) +Context {k : nat} {ltc_0_k : 0 <c k}. +Instance Robots : Names := Robots k 0. + +(** Assumptions on the number of robots: it strictly divides the ring size. *) +Hypothesis kdn : (n mod k = 0). +Hypothesis k_inf_n : (k < n). + +Lemma h_not_0: n <> 0. +Proof using ltc_2_n. + unfold ltc in *. + apply neq_lt. + transitivity 2; auto with arith. +Qed. + +Lemma k_not_0: k <> 0. +Proof using ltc_0_k. + unfold ltc in *. + now apply neq_lt. +Qed. + +Local Hint Resolve h_not_0 k_not_0: localDB. -(** Assumptions on the number of robots: it is non zero, less than and divides the ring size. *) -Hypothesis kdn : (ring_size mod kG = 0)%nat. -(* (* FIXME: This version triggers an "out of memory" error in the Program Definition of [da]! *) -Hypothesis k_bounds : (1 < kG < ring_size)%nat. -Definition k_sup_1 : (1 < kG)%nat := proj1 k_bounds. -Definition k_inf_n : (kG < ring_size)%nat := proj2 k_bounds. *) -Hypothesis k_sup_1 : (1 < kG)%nat. -Hypothesis k_inf_n : (kG < ring_size)%nat. +Lemma local_subproof1 : n = k * (n / k). +Proof using kdn ltc_0_k. apply Nat.Div0.div_exact. auto with localDB. Qed. + +Lemma local_subproof2 : n / k ≠0. +Proof using kdn ltc_2_n ltc_0_k. + intros Habs. eapply @neq_u_0. apply ltc_2_n. rewrite local_subproof1, Habs. clear. lia. +Qed. +Local Hint Resolve local_subproof1 local_subproof2 : localDB. + +Lemma local_subproof3 : ∀ m : nat, m < k -> m * (n / k) < n. +Proof using kdn k_inf_n ltc_0_k. +intros * H. pattern n at 2. rewrite (Nat.div_mod_eq n k). (* FIXME: bug in rewrite? *) +rewrite kdn, Nat.add_0_r, <- Nat.mul_lt_mono_pos_r; trivial; []. +apply Nat.div_str_pos. split; auto with arith. +Qed. Instance NoByz : NoByzantine. Proof using . now split. Qed. (** A dummy state used for (inexistant) byzantine robots. *) -Definition origin : location := of_Z 0. -Definition dummy_loc : location := origin. (* could be anything *) - -Notation "!! config" := (obs_from_config config origin) (at level 0). +Definition dummy_loc : location := fin0. (* could be anything *) +Notation "!! config" := (obs_from_config config dummy_loc) (at level 0). (** Let us consider an arbirary robogram. *) Variable r : robogram. @@ -58,13 +75,13 @@ Variable r : robogram. (** The key idea is to prove that we can always make robots think that there are in the same configuration. Thus, is they move at all, then they will never stop. If they do not move, they do not explore the ring. - The configuration to which we will always come back is [ref_config], + The configuration to which we will always come back is [NoByz_periodic_config], in which robots are evenly spaced on the ring. *) -(** *** Definition of the reference configuration and demon used in the proof **) +(** *** Definition of the demon used in the proof **) Definition create_ref_config (g : G) : location := - Ring.of_Z (Z_of_nat (proj1_sig g * (ring_size / kG))). + mod2fin (fin2nat g * (n / k)). (** The starting configuration where robots are evenly spaced: each robot is at a distance of [ring_size / kG] from the previous one, @@ -75,109 +92,75 @@ Definition ref_config : configuration := | Byz b => dummy_loc end. + Lemma ref_config_injective : Util.Preliminary.injective eq equiv (fun id => get_location (ref_config id)). -Proof using k_sup_1 k_inf_n kdn. -intros id1 id2. -assert (ring_size / kG <> 0)%nat by (rewrite Nat.div_small_iff; lia). -apply (no_byz id2), (no_byz id1). clear id1 id2. -intros g1 g2 Heq. f_equal. hnf in Heq. -unfold ref_config, create_ref_config, Ring.of_Z in *. simpl in *. -apply (f_equal to_Z) in Heq. unfold to_Z in Heq. simpl in Heq. -rewrite 2 Z2Nat.id in Heq; try (apply Z.mod_pos_bound; lia); []. -assert (Hlt : forall n, (n < kG)%nat -> Z.of_nat (n * (ring_size / kG)) < Z.of_nat ring_size). -{ intros n Hn. apply Nat2Z.inj_lt. - apply Nat.lt_le_trans with (kG * (ring_size / kG))%nat. - - apply mult_lt_compat_r; lia. - - apply Nat.mul_div_le. lia. } -rewrite 2 Z.mod_small in Heq; try (split; apply Zle_0_nat || apply Hlt, proj2_sig); []. -apply Nat2Z.inj, Nat.mul_cancel_r in Heq; auto. -Local Transparent G. unfold G. now apply eq_proj1. -Qed. - -(** Translating [ref_config] by multiples of [ring_size / kG] does not change its observation. *) -Lemma obs_trans_ref_config : forall g, - !! (map_config (Ring.trans (to_Z (create_ref_config g))) ref_config) == !! ref_config. -Proof using k_sup_1 k_inf_n kdn. -unfold obs_from_config, - MultisetObservation.multiset_observation, MultisetObservation.make_multiset. -intro g. apply MMultisetFacts.from_elements_compat. (* FIXME: [f_equiv] works but is too long *) -rewrite 2 config_list_spec, 4 map_map. -change (finite_node ring_size) with location. -apply NoDupA_equivlistA_PermutationA; autoclass; [| |]. -* apply map_injective_NoDupA with eq; autoclass; [|]. - + intros id1 id2 [Heq _]. apply ref_config_injective. - simpl in Heq. unfold Datatypes.id in *. - apply (f_equal to_Z) in Heq. rewrite 2 Z2Z in Heq. - assert (Heq_mod : (to_Z (ref_config id1)) mod Z.of_nat ring_size - = (to_Z (ref_config id2)) mod Z.of_nat ring_size). - { replace (to_Z (ref_config id1)) - with (to_Z (ref_config id1) - to_Z (create_ref_config g) - + to_Z (create_ref_config g)) by ring. - rewrite Z.add_mod, Heq, <- Z.add_mod; try lia; []. f_equal. ring. } - rewrite 2 Z.mod_small in Heq_mod; auto using to_Z_small; []. - apply to_Z_injective in Heq_mod. now rewrite Heq_mod. - + rewrite NoDupA_Leibniz. apply names_NoDup. -* apply map_injective_NoDupA with eq; autoclass; [|]. - + intros ? ? []. now apply ref_config_injective. - + rewrite NoDupA_Leibniz. apply names_NoDup. -* intro pt. repeat rewrite InA_map_iff; autoclass; []. - assert (HkG : kG <> 0%nat) by lia. - assert (Z.of_nat ring_size <> 0) by lia. - assert (ring_size / kG <> 0)%nat by (rewrite Nat.div_small_iff; lia). - assert (Hg : (proj1_sig g < kG)%nat) by apply proj2_sig. - assert (Hsize : (kG * (ring_size / kG) = ring_size)%nat). - { symmetry. now rewrite Nat.div_exact. } - split; intros [id [Hpt _]]; revert Hpt; apply (no_byz id); clear id; intros g' Hpt. - + assert (Hlt : ((proj1_sig g' + (kG - proj1_sig g)) mod kG < kG)%nat) - by now apply Nat.mod_upper_bound. - pose (id' := exist (fun x => lt x kG) _ Hlt). - change (fin kG) with G in id'. - exists (Good id'). split. - - simpl. rewrite <- Hpt. simpl. split; try reflexivity; []. hnf. simpl. - unfold Datatypes.id, create_ref_config. apply to_Z_injective. rewrite 2 Z2Z. - (* This part is a proof about modular arithmetic; we stay in Z to use the ring structure *) - rewrite 2 Ring.Z2Z, <- Zdiv.Zminus_mod. - unfold id'. simpl. - rewrite <- Nat.mul_mod_distr_r, Hsize, Zdiv.mod_Zmod, Z.mod_mod; try lia; []. - rewrite Nat.mul_add_distr_r, Nat2Z.inj_add, 3 Nat2Z.inj_mul, Nat2Z.inj_sub; try lia; []. - rewrite Z.mul_sub_distr_r, <- (Nat2Z.inj_mul kG), Hsize. - rewrite Z.add_mod, Zdiv.Zminus_mod, Z.mod_same, Z.add_mod_idemp_r; try lia; []. - rewrite Zdiv.Zminus_mod. reflexivity. - - rewrite InA_Leibniz. apply In_names. - + assert (Hlt : ((proj1_sig g' + proj1_sig g) mod kG < kG)%nat) by now apply Nat.mod_upper_bound. - pose (id' := exist (fun x => lt x kG) _ Hlt). - change (fin kG) with G in id'. - exists (Good id'). split. - - simpl. rewrite <- Hpt. simpl. split; try reflexivity; []. hnf. simpl. - unfold Datatypes.id, create_ref_config. apply to_Z_injective. rewrite 2 Z2Z. - (* This part is a proof about modular arithmetic; we stay in Z to use the ring structure *) - rewrite 2 Ring.Z2Z, <- Zdiv.Zminus_mod. - unfold id'. simpl. - rewrite <- Nat.mul_mod_distr_r, Hsize, Zdiv.mod_Zmod; try lia; []. - rewrite Zdiv.Zminus_mod_idemp_l. f_equal. lia. - - rewrite InA_Leibniz. apply In_names. +Proof using kdn ltc_0_k k_inf_n. +intros id1 id2. apply (no_byz id2), (no_byz id1). clear id1 id2. +intros g1 g2 Heq. f_equal. unfold ref_config, create_ref_config in Heq. +eapply mod2fin_betweenI, Nat.mul_cancel_r, fin2natI in Heq; trivial; +try (now split; [apply Nat.le_0_l | apply local_subproof3, fin_lt]). +intro Habs. rewrite Nat.div_small_iff in Habs; try lia; []. now rewrite Nat.neq_0_lt_0. +Qed. + +(** Translating [ref_config] by multiples of [n / k] + does not change its observation. *) +Lemma obs_asbf_ref_config : forall g, + !! (map_config (λ x, subf x (create_ref_config g)) ref_config) == !! ref_config. +Proof using kdn ltc_0_k k_inf_n. + unfold obs_from_config, MultisetObservation.multiset_observation, + MultisetObservation.make_multiset. intro g. f_equiv. + rewrite 2 config_list_spec, 4 map_map. + apply NoDupA_equivlistA_PermutationA. + * autoclass. + * apply map_injective_NoDupA with eq; autoclass; [|]. + + intros id1 id2 [Heq _]. apply ref_config_injective. apply addIm in Heq. + rewrite Heq. reflexivity. + + rewrite NoDupA_Leibniz. apply names_NoDup. + * apply map_injective_NoDupA with eq; autoclass; [|]. + + intros ? ? []. now apply ref_config_injective. + + rewrite NoDupA_Leibniz. apply names_NoDup. + * intro pt. repeat rewrite InA_map_iff; autoclass; []. + split; intros [id [Hpt _]]; revert Hpt; apply (no_byz id); clear id; intros g' Hpt. + + pose (id' := subf g' g). change (fin k) with G in id'. + exists (Good id'). split. 2:{ rewrite InA_Leibniz. apply In_names. } + rewrite <- Hpt. cbn -[create_ref_config BijectionInverse]. split; trivial; []. + change G with (fin k) in *. unfold create_ref_config, Datatypes.id, id'. apply fin2natI. + rewrite 2 subf2nat, 3 mod2fin2nat, Nat.Div0.add_mod_idemp_l, 2 (Nat.mod_small (_ * _));auto with localDB. + - pattern n at 3 5. rewrite local_subproof1. + rewrite <- Nat.mul_sub_distr_r, <- Nat.mul_add_distr_r. + rewrite Nat.Div0.mul_mod_distr_r;try auto with localDB. + - apply local_subproof3, fin_lt. + - pattern n at 2. rewrite local_subproof1, <- Nat.mul_lt_mono_pos_r; try apply mod2fin_lt; []. + apply Nat.neq_0_lt_0, local_subproof2. + + pose (id' := addf g' g). change (fin k) with G in id'. + exists (Good id'). split. 2:{ rewrite InA_Leibniz. apply In_names. } + rewrite <- Hpt. cbn. split. 2: reflexivity. unfold create_ref_config, id', Datatypes.id. apply fin2natI. + rewrite subf2nat, 3 mod2fin2nat, addf2nat, Nat.Div0.add_mod_idemp_l, 2 (Nat.mod_small (_ * _)); + try (pattern n at 2; rewrite local_subproof1, <- Nat.mul_lt_mono_pos_r; try apply fin_lt; + []; + apply Nat.div_str_pos; split; trivial; []; now apply Nat.lt_le_incl); try lia ;[]. + rewrite local_subproof1 at 2 4. + rewrite <- Nat.mul_sub_distr_r, <- Nat.mul_add_distr_r, Nat.Div0.mul_mod_distr_r; auto with localDB. + apply Nat.mul_cancel_r; try apply local_subproof2; []. + rewrite Nat.Div0.add_mod_idemp_l, <- Nat.add_assoc, (Nat.add_comm (fin2nat g)), Nat.sub_add; + try apply Nat.lt_le_incl, fin_lt;auto with localDB. + rewrite <- Nat.Div0.add_mod_idemp_r, Nat.Div0.mod_same, Nat.add_0_r;auto with localDB. apply Nat.mod_small, fin_lt. Qed. (** The demon activate all robots and shifts their view to be on 0. *) Program Definition da : demonic_action := {| activate := fun id => true; relocate_byz := fun _ _ => dummy_loc; - change_frame := fun config g => (to_Z (config (Good g)), false); + change_frame := fun config g => (config (Good g), false); choose_update := fun _ _ _ => tt; choose_inactive := fun _ _ => tt |}. -Next Obligation. (* activate_compat *) -now repeat intro. -Qed. +Next Obligation. (* activate_compat *) now repeat intro. Qed. Next Obligation. (* relocate_byz_compat *) -repeat intro. do 2 f_equal. subst. auto. -Qed. -Next Obligation. (* change_frame_compat *) -now repeat intro. -Qed. -Next Obligation. (* choose_update_compat *) -now repeat intro. + repeat intro. do 2 f_equal. subst. auto. Qed. +Next Obligation. (* change_frame_compat *) now repeat intro. Qed. +Next Obligation. (* choose_update_compat *) now repeat intro. Qed. Definition bad_demon : demon := Stream.constant da. @@ -189,8 +172,16 @@ Proof using . coinduction Hcoind. apply FSYNC_one. Qed. (** As all robots see the same observation, we take for instance the one at location [origin]. *) Definition move := pgm r (!! ref_config). -Definition target := move_along origin move. +Definition target := move_along fin0 move. +(** Acceptable starting configurations contain no tower, + that is, all robots are at different locations. *) +Definition Valid_starting_config config : Prop := + Util.Preliminary.injective (@Logic.eq ident) (@equiv _ state_Setoid) config. + +Definition Explore_and_Stop (r : robogram) := + forall d config, Fair d -> Valid_starting_config config -> + ExplorationStop (execute r d config). (** ** First case: robots do not move **) @@ -201,279 +192,285 @@ Section NoMove. Hypothesis Hmove : move == SelfLoop. Lemma round_id : round r da ref_config == ref_config. -Proof using Hmove k_inf_n k_sup_1 kdn. -rewrite FSYNC_round_simplify; try (now split); []. -apply no_byz_eq. intro g. -cbn -[Ring.trans equiv ring_edge map_config]. -unfold lift. cbn -[map_config Ring.trans equiv]. -rewrite (MultisetObservation.obs_from_config_ignore_snd origin). -rewrite obs_trans_ref_config, Hmove. cbn [move_along map_config]. -apply Bijection.retraction_section. -Qed. - -Lemma NeverVisited_ref_config : forall e, - e == execute r bad_demon ref_config -> - exists pt, ~ Will_be_visited pt e. -Proof using Hmove k_inf_n k_sup_1 kdn. -intros e Heq_e. exists (of_Z 1). -intro Hl. induction Hl as [e [g Hvisited] | e Hlater IHvisited]. -* (* FIXME: why does [rewrite Heq_e in Hvisited] fail? *) - rewrite (Stream.hd_compat Heq_e) in Hvisited. simpl in Hvisited. - apply (f_equal (@proj1_sig _ (fun x => lt x ring_size))) in Hvisited. revert Hvisited. - assert (1 < ring_size / kG)%nat by (apply <- Nat.div_exact in kdn; nia). - unfold Ring.of_Z. simpl. rewrite Z.mod_1_l, Z.mod_small; try lia; [|]. - + change 1 with (Z.of_nat 1). rewrite 2 Nat2Z.id. destruct (proj1_sig g); nia. - + split; try apply Zle_0_nat; []. - apply inj_lt, Nat.lt_le_trans with (kG * (ring_size / kG))%nat. - - apply mult_lt_compat_r; try lia; []. apply proj2_sig. - - rewrite <- Nat.div_exact in kdn; lia. -* apply IHvisited. rewrite Heq_e, execute_tail. - (* FIXME: why does [f_equiv] fail to find [execute_compat]? *) - apply execute_compat; auto using round_id. -Qed. - -Lemma never_visited : ~(forall pt, Will_be_visited pt (execute r bad_demon ref_config)). -Proof using Hmove k_inf_n k_sup_1 kdn. -intros Hw. -destruct (NeverVisited_ref_config (reflexivity _)) as [pt Hpt]. -apply Hpt, Hw. -Qed. - -Theorem no_exploration_idle : ~ Explore_and_Stop r. -Proof using Hmove k_inf_n k_sup_1 kdn. -intros Habs. -destruct (Habs bad_demon ref_config) as [Hexpl _]. -apply FSYNC_implies_fair, FYSNC_setting. -apply ref_config_injective. -now apply never_visited. +Proof using Hmove kdn ltc_0_k k_inf_n. + unfold ltc in *. + rewrite FSYNC_round_simplify. + 2: split. + apply no_byz_eq. intro g. + cbn-[create_ref_config Bijection.BijectionInverse equiv]. + erewrite transvE, asbfVE, obs_from_config_ignore_snd, obs_asbf_ref_config, Hmove, move_along_0. + rewrite Bijection.inv_inv, asbfE. apply subfVKV. +Qed. + +Lemma NeverVisited_ref_config : ∀ e, e == execute r bad_demon ref_config + -> exists pt, ~ Will_be_visited pt e. +Proof using Hmove kdn k_inf_n ltc_0_k. + intros e Heq_e. exists (mod2fin 1). intro Hl. + induction Hl as [e [g Hvisited] | e Hlater IHvisited]. + * rewrite Heq_e in Hvisited. cbn in Hvisited. + apply (f_equal (@fin2nat n)) in Hvisited. revert Hvisited. + cbn-[Nat.modulo Nat.div]. rewrite local_subproof1 at 2. + rewrite Nat.Div0.mul_mod_distr_r, (Nat.mod_small 1), Nat.mul_eq_1. + intros [_ Habs]. + apply Nat.lt_nge in k_inf_n. + contradiction k_inf_n. + rewrite local_subproof1, + Habs, Nat.mul_1_r. reflexivity. + eapply @lt_l_u. apply lt_s_u. auto. + * apply IHvisited. rewrite Heq_e, execute_tail. rewrite round_id. f_equiv. +Qed. + +Lemma never_visited : + ~(∀ pt, Will_be_visited pt (execute r bad_demon ref_config)). +Proof using Hmove kdn k_inf_n ltc_0_k. + intros Hw. destruct (NeverVisited_ref_config (reflexivity _)) as [pt Hpt]. + apply Hpt, Hw. +Qed. + +Theorem no_exploration_idle : ~ Explore_and_Stop r. +Proof using Hmove k_inf_n kdn ltc_0_k. + intros Habs. destruct (Habs bad_demon ref_config) as [Hexpl _]. + apply FSYNC_implies_Fair, FYSNC_setting. apply ref_config_injective. + now apply never_visited. Qed. End NoMove. - (** ** Second case: the robots move **) (** *** Equality of configurations up to translation **) (** Translating a configuration. *) -Definition f_config config k : configuration := map_config (trans (- k)) config. +Definition f_config config m : configuration := map_config (asbm m) config. Instance f_config_compat : Proper (equiv ==> equiv ==> equiv) f_config. Proof using . -unfold f_config. repeat intro. -apply map_config_compat; trivial; []. -intros ? ? Heq. now repeat f_equiv. -Qed. - -Lemma f_config_merge : forall config k1 k2, - f_config (f_config config k1) k2 == f_config config (k1 + k2). -Proof using k_inf_n k_sup_1 kdn. -intros config k1 k2. unfold f_config. rewrite map_config_merge; autoclass; []. -apply no_byz_eq. intro g. -repeat split; simpl. apply to_Z_injective. -repeat rewrite Z2Z, Z.sub_opp_r, ?Zdiv.Zplus_mod_idemp_l. -f_equal. ring. -Qed. - -Lemma f_config_swap : forall config k1 k2, - f_config (f_config config k1) k2 == f_config (f_config config k2) k1. -Proof using k_inf_n k_sup_1 kdn. intros. do 2 rewrite f_config_merge. f_equiv. hnf. ring. Qed. - -Lemma f_config_0 : forall config, f_config config 0 == config. -Proof using . intro. unfold f_config. simpl. intro. now rewrite Z.sub_0_r, V2V. Qed. - -Lemma f_config_injective_local : forall k config1 config2 id, - f_config config1 k id == f_config config2 k id -> config1 id == config2 id. -Proof using k_inf_n k_sup_1 kdn. -intros k config1 config2 id Heq. -setoid_rewrite <- f_config_0. replace 0 with (k + -k) by ring. -setoid_rewrite <- (f_config_merge _ _ _ id). -unfold f_config at 1 3, map_config. now rewrite Heq. -Qed. - -Lemma f_config_injective : forall k config1 config2, - f_config config1 k == f_config config2 k -> config1 == config2. -Proof using k_inf_n k_sup_1 kdn. intros * Heq ?. eapply f_config_injective_local, Heq. Qed. - -Lemma f_config_is_id : forall k config, f_config config k == config <-> of_Z k = origin. -Proof using k_inf_n k_sup_1 kdn. -intros k config. split; intro Heq. -+ assert (g : G). { exists 0%nat. compute. lia. } - specialize (Heq (Good g)). unfold f_config, map_config in Heq. - simpl in Heq. rewrite Z.sub_opp_r in Heq. - apply (f_equal to_Z) in Heq. rewrite Z2Z in Heq. - apply to_Z_injective. rewrite Z2Z. change (to_Z origin) with 0. - replace k with (to_Z (config (Good g)) + k - to_Z (config (Good g))) by ring. - rewrite Zdiv.Zminus_mod, Heq, Zdiv.Zminus_mod_idemp_r, Z.sub_diag, Z.mod_0_l; lia. -+ unfold f_config, map_config. simpl. intro id. rewrite Z.sub_opp_r. - apply to_Z_injective. apply (f_equal to_Z) in Heq. rewrite Z2Z in *. - rewrite Z.add_mod, Heq, Z.add_0_r, Z.mod_mod, Z.mod_small; lia || apply to_Z_small. -Qed. - -Lemma f_config_same_sub : forall k config1 config2, config2 == f_config config1 k -> - forall id id', of_Z (to_Z (config1 id) - to_Z (config1 id')) - == of_Z (to_Z (config2 id) - to_Z (config2 id')). + unfold f_config. intros c1 c2 Hc m1 m2 Hm. rewrite Hc, Hm. reflexivity. +Qed. + +Lemma f_config_merge : ∀ config m1 m2, + f_config (f_config config m1) m2 == f_config config (m1 + m2). +Proof using k_inf_n kdn. + intros. unfold f_config. rewrite map_config_merge; autoclass; []. + intros id. cbn-[equiv]. rewrite 3 asbmE, <- (addm_mod (config id)), <- mod2fin2nat, + <- addmA, addm2nat, mod2fin2nat, Nat.Div0.add_mod_idemp_l;auto with localDB. + apply addm_mod. +Qed. + +Lemma f_config_swap : ∀ config m1 m2, + f_config (f_config config m1) m2 == f_config (f_config config m2) m1. +Proof using k_inf_n kdn. + intros. rewrite 2 f_config_merge, Nat.add_comm. reflexivity. +Qed. + +Lemma f_config_0 : ∀ config, f_config config 0 == config. +Proof using . intros * id. cbn-[equiv]. apply addm0. Qed. + +Lemma f_config_injective_config : ∀ m config1 config2, + f_config config1 m == f_config config2 m -> config1 == config2. +Proof using k_inf_n kdn. + intros * Heq. eapply map_config_inj'. 2: apply Heq. apply Bijection.injective. +Qed. + +Lemma f_config_injective_N : ∀ config m1 m2, + f_config config m1 == f_config config m2 + -> m1 mod n == m2 mod n. +Proof using kdn k_inf_n ltc_0_k. + unfold f_config. intros * Heq. specialize (Heq (Good fin0)). + eapply (@addm_betweenI 2 n ltc_2_n). 3: rewrite 2 addm_mod. + 3: apply Heq. all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. +Qed. + +Lemma f_config_modulo : ∀ config m, + f_config config (m mod n) == f_config config m. +Proof using . intros * id. apply addm_mod. Qed. + +Lemma asbf_f_config : ∀ (config : configuration) (id1 id2 : ident) (m : nat), + asbf (config id1)â»Â¹ (config id2) + == asbf (f_config config m id1)â»Â¹ (f_config config m id2). Proof using . -intros k config1 config2 Heq id id'. -rewrite Heq. unfold f_config. simpl. apply to_Z_injective. -rewrite 2 Z.sub_opp_r, 4 Z2Z, <- Zdiv.Zminus_mod. f_equal. ring. + intros. unfold f_config. cbn-[equiv]. symmetry. apply subf_addm_addm. Qed. (** Two configurations are equivalent if they are equal up to a global translation. *) -Definition equiv_config_k k config1 config2 : Prop := config2 == f_config config1 k. -Definition equiv_config config1 config2 : Prop := exists k, equiv_config_k k config1 config2. +Definition equiv_config_m m config1 config2 : Prop + := config2 == f_config config1 m. +Definition equiv_config config1 config2 : Prop + := ∃ m, equiv_config_m m config1 config2. -Lemma equiv_config_k_sym : forall k config1 config2, - equiv_config_k k config1 config2 -> equiv_config_k (- k) config2 config1. -Proof using k_inf_n k_sup_1 kdn. -unfold equiv_config_k. intros k config1 config2 Hequiv. -rewrite Hequiv, f_config_merge, <- f_config_0 at 1. -f_equiv. hnf. ring. +Lemma equiv_config_m_sym : ∀ m config1 config2, + equiv_config_m m config1 config2 + -> equiv_config_m (@oppm 2 n ltc_2_n m) config2 config1. +Proof using k_inf_n kdn. + unfold equiv_config_m. intros * Hequiv. unshelve erewrite Hequiv, + f_config_merge, <- f_config_modulo, (proj2 (Nat.Lcm0.mod_divide _ _));auto with localDB. + symmetry. apply f_config_0. apply divide_addn_oppm. Qed. -Lemma equiv_config_k_trans : forall k1 k2 config1 config2 config3, - equiv_config_k k1 config1 config2 -> equiv_config_k k2 config2 config3 -> - equiv_config_k (k1 + k2) config1 config3. -Proof using k_inf_n k_sup_1 kdn. -unfold equiv_config_k. intros * Hequiv12 Hequiv23. -now rewrite Hequiv23, Hequiv12, f_config_merge. +Lemma equiv_config_m_trans : ∀ m1 m2 config1 config2 config3, + equiv_config_m m1 config1 config2 -> equiv_config_m m2 config2 config3 -> + equiv_config_m (m1 + m2) config1 config3. +Proof using k_inf_n kdn. + unfold equiv_config_m. intros * Hequiv12 Hequiv23. + now rewrite Hequiv23, Hequiv12, f_config_merge. Qed. - Instance equiv_config_equiv : Equivalence equiv_config. -Proof using k_inf_n k_sup_1 kdn. split. -+ intro config. exists 0. unfold equiv_config_k. now rewrite f_config_0. -+ intros config1 config2 [k Hk]. exists (- k). now apply equiv_config_k_sym. -+ intros ? ? ? [k1 Hk1] [k2 Hk2]. exists (k1 + k2). - revert Hk1 Hk2. apply equiv_config_k_trans. +Proof using k_inf_n kdn. split. + + intro config. exists 0. unfold equiv_config_m. now rewrite f_config_0. + + intros config1 config2 [m Hm]. exists (@oppm 2 n ltc_2_n m). + apply (equiv_config_m_sym Hm). + + intros ? ? ? [m1 Hm1] [m2 Hm2]. exists (m1 + m2). + revert Hm1 Hm2. apply equiv_config_m_trans. +Qed. + +Lemma equiv_config_mod : ∀ (m : nat) (config1 config2 : configuration), + equiv_config_m (m mod n) config1 config2 + <-> equiv_config_m m config1 config2. +Proof using . + intros. split; intros H id. + - rewrite <- f_config_modulo. apply H. + - rewrite f_config_modulo. apply H. Qed. (* It is actually an equivalence. *) Instance eq_equiv_subrelation : subrelation equiv equiv_config. -Proof using . intros ? ? ?. exists 0. unfold equiv_config_k. now rewrite f_config_0. Qed. +Proof using . + intros. exists 0. unfold equiv_config_m. now rewrite f_config_0. +Qed. (** Equivalent configurations produce the same observation hence the same answer from the robogram. *) -Lemma config1_obs_equiv : forall config1 config2, - equiv_config config1 config2 -> - forall g, !! (map_config (trans (to_Z (config1 (Good g)))) config1) - == !! (map_config (trans (to_Z (config2 (Good g)))) config2). -Proof using k_inf_n k_sup_1 kdn. -intros config1 config2 [offset Hequiv] g. -f_equiv. apply no_byz_eq. intro g'. simpl. -apply to_Z_injective. rewrite 2 Z2Z. -unfold equiv_config_k in Hequiv. rewrite Hequiv. -unfold f_config. simpl. rewrite 2 Z2Z, <- Zdiv.Zminus_mod. -f_equal. ring. -Qed. - -Theorem equiv_config_k_round : forall k config1 config2, - equiv_config_k k config1 config2 -> equiv_config_k k (round r da config1) (round r da config2). -Proof using k_inf_n k_sup_1 kdn. -unfold equiv_config_k. intros k config1 config2 Hequiv id. -apply (no_byz id). clear id. intro g. -rewrite (FSYNC_round_simplify r config2 FSYNC_one). -rewrite (f_config_compat (FSYNC_round_simplify r config1 FSYNC_one) (reflexivity k)). -simpl. unfold f_config. simpl. apply to_Z_injective. repeat rewrite Z2Z. -rewrite 2 Z.sub_diag, Z.sub_opp_r, Z.add_mod_idemp_l; try lia; []. -unfold Datatypes.id. rewrite <- Z.add_assoc. setoid_rewrite Z.add_mod; try lia; []. -do 2 f_equal. -+ do 3 f_equiv. apply (pgm_compat r), obs_from_config_compat; try reflexivity; []. - intro. symmetry. apply (f_config_same_sub Hequiv). -+ rewrite Hequiv. unfold f_config. simpl. rewrite Z2Z, Z.sub_opp_r, Z.mod_mod; lia. -Qed. - -Corollary equiv_config_round : forall config1 config2, equiv_config config1 config2 -> - equiv_config (round r da config1) (round r da config2). -Proof using k_inf_n k_sup_1 kdn. intros config1 config2 [k Hequiv]. exists k. now apply equiv_config_k_round. Qed. +Lemma config1_obs_equiv : ∀ config1 config2, equiv_config config1 config2 + -> ∀ g, !! (map_config (asbf (config1 (Good g))â»Â¹) config1) + == !! (map_config (asbf (config2 (Good g))â»Â¹) config2). +Proof using k_inf_n kdn. + intros config1 config2 [offset Hequiv] g. f_equiv. + unfold equiv_config_m in Hequiv. rewrite Hequiv. rewrite Hequiv. + intros id. cbn[map_config]. apply asbf_f_config. +Qed. + +Theorem equiv_config_m_round : ∀ m config1 config2, + equiv_config_m m config1 config2 + -> equiv_config_m m (round r da config1) (round r da config2). +Proof using k_inf_n kdn. + unfold equiv_config_m. intros * Hequiv. unfold f_config. + rewrite 2 (FSYNC_round_simplify r _ FSYNC_one). intros id. apply (no_byz id). + clear id. intro g. cbn-[equiv Bijection.BijectionInverse]. + setoid_rewrite <- config1_obs_equiv. 2,3: reflexivity. + repeat rewrite Hequiv. + setoid_rewrite transvVE. rewrite 2 transvE, 2 asbfE, 2 asbfVE, asbmE. cbn-[f_config equiv]. + rewrite <- 2 subf_move_along'. unfold f_config, map_config. cbn. rewrite 2 subfVKV, addm_move_along. + apply move_along_compat. reflexivity. apply (pgm_compat r), + obs_from_config_compat. 2: reflexivity. intros id. apply subf_addm_addm. +Qed. +Corollary equiv_config_round : ∀ config1 config2, equiv_config config1 config2 + -> equiv_config (round r da config1) (round r da config2). +Proof using k_inf_n kdn. + intros config1 config2 [m Hequiv]. exists m. now apply equiv_config_m_round. +Qed. (** *** Equality of executions up to translation **) -Definition AlwaysEquiv k (e1 e2 : execution) : Prop := - Stream.forever2 (Stream.instant2 (equiv_config_k k)) e1 e2. +Definition AlwaysEquiv m (e1 e2 : execution) : Prop := + Stream.forever2 (Stream.instant2 (equiv_config_m m)) e1 e2. -Lemma AlwaysEquiv_refl : forall e, AlwaysEquiv 0 e e. +Lemma AlwaysEquiv_refl : ∀ e, AlwaysEquiv 0 e e. Proof using . -coinduction Hcoind. -unfold Stream.instant2, equiv_config_k. -now rewrite f_config_0. + coinduction Hcoind. unfold Stream.instant2, equiv_config_m. + now rewrite f_config_0. Qed. -Lemma AlwaysEquiv_sym : forall k (e1 e2 : execution), - AlwaysEquiv k e1 e2 -> AlwaysEquiv (- k) e2 e1. -Proof using k_inf_n k_sup_1 kdn. -cofix Hcoind. -intros k1 e1 e2 [Hnow Hlater]. -constructor. -+ now apply equiv_config_k_sym. -+ apply Hcoind; auto. +Lemma AlwaysEquiv_sym : ∀ m (e1 e2 : execution), + AlwaysEquiv m e1 e2 -> AlwaysEquiv (@oppm 2 n ltc_2_n m) e2 e1. +Proof using k_inf_n kdn. + cofix Hcoind. intros m1 e1 e2 [Hnow Hlater]. constructor. + + now apply equiv_config_m_sym. + + apply Hcoind; auto. Qed. -Lemma AlwaysEquiv_trans : forall k1 k2 (e1 e2 e3 : execution), - AlwaysEquiv k1 e1 e2 -> AlwaysEquiv k2 e2 e3 -> AlwaysEquiv (k1 + k2) e1 e3. -Proof using k_inf_n k_sup_1 kdn. -cofix Hrec. -intros k1 k2 e1 e2 e3 [Hnow12 Hlater12] [Hnow23 Hnlater23]. -constructor. -+ eapply equiv_config_k_trans; eauto. -+ apply Hrec with (Stream.tl e2); auto. +Lemma AlwaysEquiv_trans : ∀ m1 m2 (e1 e2 e3 : execution), + AlwaysEquiv m1 e1 e2 -> AlwaysEquiv m2 e2 e3 -> AlwaysEquiv (m1 + m2) e1 e3. +Proof using k_inf_n kdn. + cofix Hrec. intros * [Hnow12 Hlater12] [Hnow23 Hnlater23]. constructor. + + eapply equiv_config_m_trans; eauto. + + apply Hrec with (tl e2); auto. Qed. -Instance execute_equiv_compat : forall k, - Proper (equiv_config_k k ==> AlwaysEquiv k) (execute r bad_demon). -Proof using k_inf_n k_sup_1 kdn. intro k. coinduction Hrec; trivial; []. simpl. now apply equiv_config_k_round. Qed. +Instance execute_equiv_compat : ∀ m, + Proper (equiv_config_m m ==> AlwaysEquiv m) (execute r bad_demon). +Proof using k_inf_n kdn. + intros. coinduction Hrec; trivial; []. simpl. + now apply equiv_config_m_round. +Qed. -(** Stopping is invariant by this notion of equivalence. *) -Instance Stall_equiv_compat : forall k, Proper (AlwaysEquiv k ==> iff) Stall. -Proof using k_inf_n k_sup_1 kdn. -intros k s1 s2 Hequiv. unfold Stall. destruct Hequiv as [Hequiv [Hequiv' Hlater]]. -unfold Stream.instant2, equiv_config_k in *. -rewrite Hequiv, Hequiv'. split. -- intro Heq. now rewrite Heq. -- apply f_config_injective. +Lemma AlwaysEquiv_mod : ∀ (m : nat) (e1 e2 : execution), + AlwaysEquiv (m mod n) e1 e2 <-> AlwaysEquiv m e1 e2. +Proof using . + intros. split. + - revert m e1 e2. cofix Hrec. intros * H. constructor. + setoid_rewrite <- equiv_config_mod. apply H. apply Hrec, H. + - revert m e1 e2. cofix Hrec. intros * H. constructor. + setoid_rewrite equiv_config_mod. apply H. apply Hrec, H. Qed. -Lemma Stopped_equiv_compat_aux : forall k e1 e2, - AlwaysEquiv k e1 e2 -> Stopped e1 -> Stopped e2. -Proof using k_inf_n k_sup_1 kdn. -cofix Hcoind. intros k e1 e2 Hequiv Hstop. -constructor. -+ rewrite <- (Stall_equiv_compat Hequiv). apply Hstop. -+ destruct Hequiv. apply (Hcoind _ _ _ Hequiv), Hstop. +(** Stopping is invariant by this notion of equivalence. *) +Instance Stall_equiv_compat : ∀ m, Proper (AlwaysEquiv m ==> iff) Stall. +Proof using k_inf_n kdn. + intros m s1 s2 Hequiv. unfold Stall. + destruct Hequiv as [Hequiv [Hequiv' Hlater]]. + unfold instant2, equiv_config_m in *. rewrite Hequiv, Hequiv'. split. + - intro Heq. now rewrite Heq. + - apply f_config_injective_config. Qed. -Instance Stopped_equiv_compat : forall k, Proper (AlwaysEquiv k ==> iff) Stopped. -Proof using k_inf_n k_sup_1 kdn. intros ? ? ? ?. split; eapply Stopped_equiv_compat_aux; eauto using AlwaysEquiv_sym. Qed. +Lemma Stopped_equiv_compat_aux : ∀ m e1 e2, + AlwaysEquiv m e1 e2 -> Stopped e1 -> Stopped e2. +Proof using k_inf_n kdn. + cofix Hcoind. intros m e1 e2 Hequiv Hstop. constructor. + + rewrite <- (Stall_equiv_compat Hequiv). apply Hstop. + + destruct Hequiv. apply (Hcoind _ _ _ Hequiv), Hstop. +Qed. -Instance NoStopped_equiv_compat : forall k, Proper (AlwaysEquiv k ==> iff) (fun e => ~Stopped e). -Proof using k_inf_n k_sup_1 kdn. intros ? ? ? Hequiv. now rewrite (Stopped_equiv_compat Hequiv). Qed. +Instance Stopped_equiv_compat : ∀ m, Proper (AlwaysEquiv m ==> iff) Stopped. +Proof using k_inf_n kdn. + intros ? ? ? ?. + split; eapply Stopped_equiv_compat_aux; eauto using AlwaysEquiv_sym. +Qed. +Instance NoStopped_equiv_compat : ∀ m, + Proper (AlwaysEquiv m ==> iff) (fun e => ~Stopped e). +Proof using k_inf_n kdn. + intros ? ? ? Hequiv. now rewrite (Stopped_equiv_compat Hequiv). +Qed. (** An execution that never stops is always moving. *) Definition AlwaysMoving (e : execution) : Prop := - Stream.forever (fun e1 => ~Stopped e1) e. + forever (fun e1 => ~Stopped e1) e. -Lemma AlwaysMoving_equiv_compat_aux : forall k e1 e2, - AlwaysEquiv k e1 e2 -> AlwaysMoving e1 -> AlwaysMoving e2. -Proof using k_inf_n k_sup_1 kdn. -cofix Hcoind. intros k e1 e2 Hequiv He. -constructor. -+ rewrite <- (NoStopped_equiv_compat Hequiv). apply He. -+ destruct Hequiv. apply (Hcoind _ _ _ Hequiv), He. +Lemma AlwaysMoving_equiv_compat_aux : ∀ m e1 e2, + AlwaysEquiv m e1 e2 -> AlwaysMoving e1 -> AlwaysMoving e2. +Proof using k_inf_n kdn. + cofix Hcoind. intros m e1 e2 Hequiv He. constructor. + + rewrite <- (NoStopped_equiv_compat Hequiv). apply He. + + destruct Hequiv. apply (Hcoind _ _ _ Hequiv), He. Qed. -Instance AlwaysMoving_equiv_compat : forall k, Proper (AlwaysEquiv k ==> iff) AlwaysMoving. -Proof using k_inf_n k_sup_1 kdn. -intros ? ? ? ?. -split; eapply AlwaysMoving_equiv_compat_aux; eauto using AlwaysEquiv_sym. +Instance AlwaysMoving_equiv_compat : ∀ m, + Proper (AlwaysEquiv m ==> iff) AlwaysMoving. +Proof using k_inf_n kdn. + intros ? ? ? ?. + split; eapply AlwaysMoving_equiv_compat_aux; eauto using AlwaysEquiv_sym. Qed. -Instance AlwaysMoving_execute_compat : forall k, - Proper (equiv_config_k k ==> iff) (fun config => AlwaysMoving (execute r bad_demon config)). -Proof using k_inf_n k_sup_1 kdn. intros k ? ? Hequiv. apply (AlwaysMoving_equiv_compat (execute_equiv_compat Hequiv)). Qed. - +Instance AlwaysMoving_execute_compat : ∀ m, Proper (equiv_config_m m ==> iff) + (λ config, AlwaysMoving (execute r bad_demon config)). +Proof using k_inf_n kdn. + intros m ? ? Hequiv. + apply (AlwaysMoving_equiv_compat (execute_equiv_compat Hequiv)). +Qed. (** *** Proof when robots move **) @@ -484,82 +481,72 @@ Section DoesMove. Hypothesis Hmove : move =/= SelfLoop. (** After a round, the configuration obtained from ref_config is equivalent to ref_config. *) -Lemma round_simplify : round r da ref_config == f_config ref_config (to_Z target). -Proof using k_inf_n k_sup_1 kdn. -apply no_byz_eq. intro g. -rewrite (FSYNC_round_simplify r ref_config FSYNC_one). -cbn -[equiv map_config trans]. -rewrite MultisetObservation.obs_from_config_ignore_snd. -rewrite obs_trans_ref_config. -cbn -[trans equiv]. rewrite trans_same. fold origin. -unfold f_config, map_config. simpl. now rewrite Z.add_comm, Z.sub_opp_r. +Lemma round_simplify : round r da ref_config == f_config ref_config target. +Proof using k_inf_n kdn ltc_0_k. + rewrite (FSYNC_round_simplify _ _ FSYNC_one). apply no_byz_eq. intro g. + cbn-[Bijection.BijectionInverse mod2fin equiv]. + setoid_rewrite transvVE. rewrite transvE, asbfE, asbfVE. + rewrite obs_from_config_ignore_snd, obs_asbf_ref_config. + unfold target. rewrite subff, 2 move_along0_, addfC. unfold f_config. + rewrite asbmE. cbn. rewrite addm_addf, <- dir2nodE. reflexivity. Qed. -Corollary round_ref_config : equiv_config_k (to_Z target) ref_config (round r da ref_config). -Proof using k_inf_n k_sup_1 kdn. apply round_simplify. Qed. +Corollary round_ref_config : + equiv_config_m target ref_config (round r da ref_config). +Proof using k_inf_n kdn ltc_0_k. apply round_simplify. Qed. Corollary AlwaysEquiv_ref_config : - AlwaysEquiv (to_Z target) (execute r bad_demon ref_config) - (Stream.tl (execute r bad_demon ref_config)). -Proof using k_inf_n k_sup_1 kdn. apply execute_equiv_compat, round_simplify. Qed. + AlwaysEquiv target (execute r bad_demon ref_config) + (tl (execute r bad_demon ref_config)). +Proof using k_inf_n kdn ltc_0_k. + apply execute_equiv_compat, round_simplify. +Qed. (** An execution that is always moving cannot stop. *) -Lemma AlwaysMoving_not_WillStop : forall e, AlwaysMoving e -> ~Will_stop e. +Lemma AlwaysMoving_not_WillStop : ∀ e, AlwaysMoving e -> ~Will_stop e. Proof using . -intros e [Hnow Hmoving] Hstop. -induction Hstop as [e Hstop | e Hstop IHstop]. -+ contradiction. -+ inv Hmoving. now apply IHstop. + intros e [Hnow Hmoving] Hstop. induction Hstop as [e Hstop | e Hstop IHstop]. + contradiction. inv Hmoving. now apply IHstop. Qed. (** The starting configuration is always moving. *) Lemma ref_config_AlwaysMoving : AlwaysMoving (execute r bad_demon ref_config). -Proof using Hmove k_inf_n k_sup_1 kdn. -generalize (AlwaysEquiv_refl (execute r bad_demon ref_config)). -generalize 0, (execute r bad_demon ref_config) at 1 3. -cofix Hcoind. intros k e Hequiv. constructor. -+ clear Hcoind. rewrite Hequiv. intros [Hstop _]. - unfold Stall in Hstop. - rewrite execute_tail, round_simplify in Hstop. simpl Stream.hd in Hstop. - symmetry in Hstop. rewrite f_config_is_id in Hstop. - apply (f_equal to_Z) in Hstop. revert Hstop. - unfold of_Z, to_Z, target, move_along. simpl. - destruct move; simpl; repeat rewrite Z2Nat.id; try (apply Z.mod_pos_bound; lia); [| |]. - - rewrite 2 Z.mod_1_l; lia || discriminate. - - rewrite Z.mod_mod, <- (Z.mod_add _ 1); try lia; []. - replace (-1 + 1 * Z.of_nat ring_size) with (Z.of_nat ring_size - 1) by ring. - rewrite Z.mod_small; lia. - - now elim Hmove. -+ apply (Hcoind (k - to_Z target)). clear Hcoind. - apply AlwaysEquiv_trans with (Stream.tl (execute r bad_demon ref_config)). - - now inv Hequiv. - - apply AlwaysEquiv_sym, AlwaysEquiv_ref_config. +Proof using Hmove k_inf_n kdn ltc_0_k. + generalize (AlwaysEquiv_refl (execute r bad_demon ref_config)). + generalize 0 at 1. generalize (execute r bad_demon ref_config) at 1 3. + cofix Hcoind. intros e m Hequiv. constructor. + + clear Hcoind. rewrite Hequiv. intros [Hstop _]. contradiction Hmove. + unfold Stall in Hstop. rewrite execute_tail, round_simplify in Hstop. + simpl hd in Hstop. apply (move_along_I fin0). rewrite move_along_0. + apply fin2natI. setoid_rewrite <- Nat.mod_small. 2,3: apply fin_lt. + eapply f_config_injective_N. rewrite fin02nat, f_config_0. + symmetry. apply Hstop. + + eapply (Hcoind _ (addm (oppf target) m)). clear Hcoind. rewrite addm2nat, + Nat.add_comm. apply AlwaysEquiv_mod. + apply AlwaysEquiv_trans with (tl (execute r bad_demon ref_config)). + apply Hequiv. rewrite oppf_oppm. + apply AlwaysEquiv_sym, AlwaysEquiv_ref_config. Qed. (** Final theorem when robots move. *) Theorem no_exploration_moving : ~ Explore_and_Stop r. -Proof using Hmove k_inf_n k_sup_1 kdn. - intros Habs. - unfold Explore_and_Stop in *. - destruct (Habs bad_demon ref_config) as [_ Hstop]. - apply FSYNC_implies_fair. apply FYSNC_setting. - unfold Valid_starting_config. - apply ref_config_injective. - revert Hstop. -now apply AlwaysMoving_not_WillStop, ref_config_AlwaysMoving. +Proof using Hmove k_inf_n kdn ltc_0_k. + intros Habs. unfold Explore_and_Stop in *. + destruct (Habs bad_demon ref_config) as [_ Hstop]. apply FSYNC_implies_Fair. + apply FYSNC_setting. unfold Valid_starting_config. apply ref_config_injective. + revert Hstop. now apply AlwaysMoving_not_WillStop, ref_config_AlwaysMoving. Qed. End DoesMove. - (** Final theorem combining both cases: In the asynchronous model, if the number of robots [kG] divides the size [n] of the ring, then the exploration with stop of a n-node ring is not possible. *) Theorem no_exploration : ~ Explore_and_Stop r. -Proof using k_inf_n k_sup_1 kdn. -destruct (move =?= SelfLoop) as [Hmove | Hmove]. -+ now apply no_exploration_idle. -+ now apply no_exploration_moving. +Proof using k_inf_n kdn ltc_0_k. + destruct (move =?= SelfLoop) as [Hmove | Hmove]. + + now apply no_exploration_idle. + + now apply no_exploration_moving. Qed. End Exploration. diff --git a/CaseStudies/Exploration/Tower.v b/CaseStudies/Exploration/Tower.v index fc378f5f1880858a17b50cfd4ecefbc3fba85a15..39c0e8ed69f271fc9fc9d022482ba5a24cbb959f 100644 --- a/CaseStudies/Exploration/Tower.v +++ b/CaseStudies/Exploration/Tower.v @@ -16,188 +16,171 @@ *) (**************************************************************************) - Require Import Utf8. Require Import List SetoidList. -Require Import Decidable. -Require Import Setoid Equalities Morphisms. -Require Import Compare_dec FinFun. -Require Import ZArith Arith_base Arith.Div2 Lia Psatz. +Require Import Arith_base Lia. +Require Import Pactole.Util.Stream. +Require Import Pactole.Util.Enum. +Require Import Pactole.Util.Fin. Require Import Pactole.Models.NoByzantine. -Require Import Pactole.CaseStudies.Exploration.Definitions. - +Require Import Pactole.Models.RingSSync. +Require Import Pactole.CaseStudies.Exploration.ExplorationDefs. Open Scope list_scope. Set Implicit Arguments. -Typeclasses eauto := (bfs). +(* Typeclasses eauto := (bfs). *) Section Tower. (** Given an abitrary ring *) -Context {RR : RingSpec}. -(** There are kG good robots and no byzantine ones. *) -Variable kG : nat. -Instance Robots : Names := Robots kG 0. +Context {n : nat} {ltc_2_n : 2 <c n}. +(** There are k good robots and no byzantine ones. *) +Context {k : nat} {ltc_0_k : 0 <c k}. +Instance Robots : Names := Robots k 0. (** Assumptions on the number of robots: it is non zero and strictly divides the ring size. *) -Hypothesis kdn : (ring_size mod kG = 0)%nat. -Hypothesis k_inf_n : (kG < ring_size)%nat. +Hypothesis kdn : (n mod k = 0). +Hypothesis k_inf_n : (k < n). (** There is no byzantine robot. *) Instance NoByz : NoByzantine. Proof using . now split. Qed. -Definition origin : location := of_Z 0. - -Notation "!! config" := (@obs_from_config _ _ _ _ multiset_observation config origin) (at level 0). +Notation "!! config" := + (@obs_from_config _ _ _ _ multiset_observation config fin0) (at level 0). Notation execute := (execute (UpdFun := UpdFun)). (** In order to prove that at least one position is occupied, we define the list of positions. *) -Definition Vlist := Identifiers.enum ring_size. +Definition Vlist := enum (n). Lemma Vlist_NoDup : NoDupA equiv Vlist. Proof using . rewrite NoDupA_Leibniz. apply enum_NoDup. Qed. -Lemma Vlist_length : length Vlist = ring_size. +Lemma Vlist_length : length Vlist = n. Proof using . apply enum_length. Qed. (** As there is strictly less robots than location, there is an empty location. *) -Lemma ConfigExistsEmpty : forall config, ¬ (∀ pt, In pt (!! config)). +Lemma ConfigExistsEmpty : ∀ config, ¬ (∀ pt, In pt (!! config)). Proof using k_inf_n kdn. -generalize k_inf_n; intros Hkin config Hall. -assert (Hsize : size (!! config) < ring_size). -{ apply le_lt_trans with (cardinal (!! config)). - - apply size_cardinal. - - cut (cardinal (!! config) = kG); try lia; []. - change (cardinal (make_multiset (List.map get_location (config_list config))) = kG). - rewrite cardinal_make_multiset, config_list_spec, map_map, map_length. - rewrite names_length. simpl. lia. } -assert (Hle : ring_size <= size (!! config)). -{ rewrite size_spec. - assert (Hobs : forall pt, InA equiv pt (support (!! config))). - { intro pt. specialize (Hall pt). now rewrite support_spec. } - rewrite <- Vlist_length. - apply (Preliminary.inclA_length setoid_equiv). - - apply Vlist_NoDup. - - repeat intro. apply Hobs. } -lia. + generalize k_inf_n; intros Hkin config Hall. + assert (Hsize : size (!! config) < n). + { apply Nat.le_lt_trans with (cardinal (!! config)). apply size_cardinal. + cut (cardinal (!! config) = k). intros H. rewrite H. apply k_inf_n. + change (cardinal (make_multiset (List.map get_location + (config_list config))) = k). rewrite cardinal_make_multiset, + config_list_spec, map_map, map_length, names_length. simpl. + rewrite Nat.add_0_r. reflexivity. } + assert (Hle : n <= size (!! config)). + { rewrite size_spec. assert (Hobs : ∀ pt, InA equiv pt (support (!! config))). + { intro pt. specialize (Hall pt). now rewrite support_spec. } + rewrite <- Vlist_length at 1. apply (Preliminary.inclA_length setoid_equiv). + apply Vlist_NoDup. repeat intro. apply Hobs. } lia. Qed. -Lemma Stopped_same : forall e, Stopped e -> e == Stream.tl e. +Lemma Stopped_same : ∀ e, Stopped e -> e == Stream.tl e. Proof using . -cofix Hcoind. intros e Hstop. constructor. -+ clear Hcoind. apply Hstop. -+ apply Hcoind. apply Hstop. + cofix Hcoind. intros e Hstop. constructor. + + clear Hcoind. apply Hstop. + + apply Hcoind. apply Hstop. Qed. -Lemma Will_stop_tl : forall e, Will_stop e -> Will_stop (Stream.tl e). +Lemma Will_stop_tl : ∀ e, Will_stop e -> Will_stop (Stream.tl e). Proof using . -intros e He. induction He. -+ left. match goal with H : Stopped _ |- _ => apply H end. -+ right. apply IHHe. + intros e He. induction He. + + left. match goal with H : Stopped _ |- _ => apply H end. + + right. apply IHHe. Qed. +(** Acceptable starting configurations contain no tower, + that is, all robots are at different locations. *) +Definition Valid_starting_config config : Prop := + Util.Preliminary.injective (@Logic.eq ident) (@equiv _ state_Setoid) config. + +Definition Explore_and_Stop (r : robogram) := + forall d config, Fair d -> Valid_starting_config config -> + ExplorationStop (execute r d config). + (** No algorithm can stop on a starting configuration. *) -Theorem no_stop_on_starting_config : forall r d config, - Fair d -> - Explore_and_Stop r -> - Valid_starting_config config -> - ~Stopped (execute r d config). +Theorem no_stop_on_starting_config : ∀ r d config, + Fair d -> Explore_and_Stop r -> Valid_starting_config config -> + ~ Stopped (execute r d config). Proof using k_inf_n kdn. -intros r d config. -generalize (@reflexivity execution equiv _ (execute r d config)). -generalize (execute r d config) at -2. -intros e Heqe Hfair Hsol Hvalid Hsto. -destruct (Hsol d config Hfair Hvalid) as [Hvisit Hstop]. -assert (Hfalse := ConfigExistsEmpty config). -(* TODO: remove the use of classical logic: everything is decidable here *) -apply Logic.Classical_Pred_Type.not_all_ex_not in Hfalse. -destruct Hfalse as [loc Hfalse]. -specialize (Hvisit loc). -rewrite <- Heqe in *. -induction Hvisit. -+ rewrite Heqe in *. - match goal with H : Stream.instant _ _ |- _ => destruct H as [g Hg] end. - rewrite (obs_from_config_In config origin) in Hfalse; - destruct Hfalse. - exists (Good g). - apply Hg. -+ apply IHHvisit. - - rewrite <- Heqe. symmetry. now apply Stopped_same. - - apply Hsto. - - now apply Will_stop_tl. + intros r d config. + generalize (@reflexivity execution equiv _ (execute r d config)). + generalize (execute r d config) at -2. + intros e Heqe Hfair Hsol Hvalid Hsto. + destruct (Hsol d config Hfair Hvalid) as [Hvisit Hstop]. + assert (Hfalse := ConfigExistsEmpty config). + (* TODO: remove the use of classical logic: everything is decidable here *) + apply Logic.Classical_Pred_Type.not_all_ex_not in Hfalse. + destruct Hfalse as [loc Hfalse]. specialize (Hvisit loc). + rewrite <- Heqe in *. induction Hvisit. + + rewrite Heqe in *. + match goal with H : Stream.instant _ _ |- _ => destruct H as [g Hg] end. + rewrite (obs_from_config_In config fin0) in Hfalse; destruct Hfalse. + exists (Good g). apply Hg. + + apply IHHvisit. + - rewrite <- Heqe. symmetry. now apply Stopped_same. + - apply Hsto. + - now apply Will_stop_tl. Qed. (** In particular, there is a tower on any final configuration. *) -Lemma tower_on_final_config : forall r d config, - Fair d -> - Explore_and_Stop r -> - Stopped (execute r d config) -> - exists loc, ((!! config)[loc] > 1)%nat. +Lemma tower_on_final_config : ∀ r d config, Fair d -> Explore_and_Stop r + -> Stopped (execute r d config) -> ∃ loc, (!! config)[loc] > 1. Proof using k_inf_n kdn. -intros r d config Hfair Hsol Hstop. -assert (Hequiv := @no_stop_on_starting_config r d config Hfair Hsol). -assert (Hvalid : ~Valid_starting_config config) by tauto. -apply config_not_injective in Hvalid. -destruct Hvalid as [id [id' [Hid Heq]]]. -exists (config id). -assert (Hobs := obs_from_config_spec config origin (config id)). -assert (Hperm : exists l, PermutationA equiv (config_list config) (config id :: config id' :: l)). -{ assert (Hin : List.In id names) by apply In_names. - assert (Hin' : List.In id' names) by apply In_names. - assert (Hperm : exists l, PermutationA eq names (id :: id' :: l)). - { rewrite <- InA_Leibniz in Hin, Hin'. - apply PermutationA_split in Hin; autoclass; []. - destruct Hin as [l' Hperm']. rewrite Hperm', InA_cons in Hin'. - destruct Hin' as [| Hin']; try congruence; []. - apply PermutationA_split in Hin'; autoclass; []. - destruct Hin' as [l Hperm]. exists l. now rewrite Hperm', Hperm. } - destruct Hperm as [l Hperm]. - exists (List.map config l). - now rewrite config_list_spec, Hperm. } -destruct Hperm as [l Hperm]. -rewrite Hobs. -(* FIXME: why does [rewrite Hperm] fail here? *) -rewrite (countA_occ_compat _ equiv_dec _ _ (reflexivity (config id)) - (PermutationA_map _ _ Hperm)). -simpl List.map. rewrite List.map_id. unfold Datatypes.id. simpl. -repeat destruct_match; solve [lia | exfalso; auto]. + intros r d config Hfair Hsol Hstop. + assert (Hequiv := @no_stop_on_starting_config r d config Hfair Hsol). + assert (Hvalid : ~Valid_starting_config config) by tauto. + apply config_not_injective in Hvalid. destruct Hvalid as [id [id' [Hid Heq]]]. + exists (config id). + assert (Hobs := obs_from_config_spec config fin0 (config id)). + assert (Hperm : exists l, PermutationA equiv (config_list config) + (config id :: config id' :: l)). + { assert (Hin : List.In id names) by apply In_names. + assert (Hin' : List.In id' names) by apply In_names. + assert (Hperm : exists l, PermutationA eq names (id :: id' :: l)). + { rewrite <- InA_Leibniz in Hin, Hin'. + apply PermutationA_split in Hin; autoclass; []. + destruct Hin as [l' Hperm']. rewrite Hperm', InA_cons in Hin'. + destruct Hin' as [| Hin']; try congruence; []. + apply PermutationA_split in Hin'; autoclass; []. + destruct Hin' as [l Hperm]. exists l. now rewrite Hperm', Hperm. } + destruct Hperm as [l Hperm]. exists (List.map config l). + now rewrite config_list_spec, Hperm. } + destruct Hperm as [l Hperm]. rewrite Hobs. + (* FIXME: why does [rewrite Hperm] fail here? *) + rewrite (countA_occ_compat _ equiv_dec _ _ (reflexivity (config id)) + (PermutationA_map _ _ Hperm)). + simpl List.map. rewrite List.map_id. unfold Datatypes.id. simpl. + repeat destruct_match; solve [lia | exfalso; auto]. Qed. -Lemma exec_stopped r : forall d c, Fair d -> Will_stop (execute r d c) -> - exists d' c', Fair d'/\ Stopped (execute r d' c'). -(*exists e, exec_r_comp e r /\ Stopped e.*) +Lemma exec_stopped (r : robogram) : + ∀ (d : demon) (c : configuration), Fair d -> Will_stop (execute r d c) + -> exists d' c', Fair d'/\ Stopped (execute r d' c'). Proof using . -intros d' config' Hfair Hstop. -remember (execute r d' config') as e'. -revert Heqe'. -revert d' config' Hfair. -induction Hstop as [e' Hstop | e' Hstop IHstop]. -+ intros d' config' Hfair Heq. - exists d', config'; now rewrite Heq in *. -+ intros d' config' Hfair Heq. - apply (IHstop (Stream.tl d') (Stream.hd (Stream.tl e'))). - - destruct Hfair as [_ Hfair]. constructor; apply Hfair. - - now rewrite Heq, execute_tail. + intros d' config' Hfair Hstop. remember (execute r d' config') as e'. + revert Heqe'. revert d' config' Hfair. + induction Hstop as [e' Hstop | e' Hstop IHstop]. + + intros d' config' Hfair Heq. exists d', config'; now rewrite Heq in *. + + intros d' config' Hfair Heq. + apply (IHstop (tl d') (hd (tl e'))). + - destruct Hfair as [_ Hfair]. constructor; apply Hfair. + - now rewrite Heq, execute_tail. Qed. -Lemma no_exploration_k_inf_2 : forall r d config, - Fair d -> - Explore_and_Stop r -> - Valid_starting_config config -> - (kG > 1)%nat. +Lemma no_exploration_k_inf_2 : ∀ r d config, + Fair d -> Explore_and_Stop r -> Valid_starting_config config -> k > 1. Proof using k_inf_n kdn. -intros r d config Hfair Hsol Hvalid. -assert (Hexr := exec_stopped r). -assert (Htower := tower_on_final_config). -destruct (Hsol d config Hfair Hvalid) as [Hvisit Hstop]. -destruct (Hexr d config Hfair Hstop) as [d' [config' [Hfair' Hstop']]]. -specialize (Htower r d' config' Hfair' Hsol Hstop'). -destruct Htower. -assert (Hcard := cardinal_lower x (!! config')). -rewrite cardinal_obs_from_config in Hcard. -unfold nG, nB in *. -unfold Robots in *. simpl in *. lia. + intros r d config Hfair Hsol Hvalid. assert (Hexr := exec_stopped r). + assert (Htower := tower_on_final_config). + destruct (Hsol d config Hfair Hvalid) as [Hvisit Hstop]. + destruct (Hexr d config Hfair Hstop) as [d' [config' [Hfair' Hstop']]]. + specialize (Htower r d' config' Hfair' Hsol Hstop'). destruct Htower. + assert (Hcard := cardinal_lower x (!! config')). + rewrite cardinal_obs_from_config in Hcard. unfold nG, nB in *. + unfold Robots in *. simpl in *. lia. Qed. End Tower. diff --git a/CaseStudies/Gathering/Definitions.v b/CaseStudies/Gathering/Definitions.v index 871e214b0f1f0abd7558bb66ec1a043dfc2d8165..c144450b85e8f43e5dcaa5421c887b3c94df78f0 100644 --- a/CaseStudies/Gathering/Definitions.v +++ b/CaseStudies/Gathering/Definitions.v @@ -19,7 +19,6 @@ (**************************************************************************) -Require Import Arith.Div2. Require Import Lia. Require Export SetoidDec. Require Export Pactole.Util.Preliminary. @@ -48,11 +47,13 @@ Notation "!!" := (fun config => obs_from_config config origin). (** [gathered_at conf pt] means that in configuration [conf] all good robots are at the same location [pt] (exactly). *) -Definition gathered_at (pt : location) (config : configuration) := forall g, get_location (config (Good g)) == pt. +Definition gathered_at (pt : location) (config : configuration) := + forall g, get_location (config (Good g)) == pt. (** [Gather pt e] means that at all rounds of (infinite) execution [e], robots are gathered at the same position [pt]. *) -Definition Gather (pt: location) (e : execution) : Prop := Stream.forever (Stream.instant (gathered_at pt)) e. +Definition Gather (pt: location) (e : execution) : Prop := + Stream.forever (Stream.instant (gathered_at pt)) e. (** [WillGather e] means that the (infinite) execution [e] is *eventually* [Gather]ed. *) Definition WillGather (e : execution) : Prop := diff --git a/CaseStudies/Gathering/Impossibility.v b/CaseStudies/Gathering/Impossibility.v index d1a8d363ff8be4984effa07b47f4ffde19504a6d..ac16573c37feb01dc71961afa10365f782bd526d 100644 --- a/CaseStudies/Gathering/Impossibility.v +++ b/CaseStudies/Gathering/Impossibility.v @@ -17,17 +17,13 @@ This file is distributed under the terms of the CeCILL-C licence. *) (**************************************************************************) -Require Import Reals. -Require Import Psatz. -Require Import Morphisms. -Require Import Arith.Div2. -Require Import Lia. -Require Import List SetoidList. -Require Import Pactole.Util.Preliminary. +Require Import Reals Psatz Morphisms Lia List SetoidList. +From Pactole Require Import Util.Preliminary Util.Fin. +Require Import Pactole.Util.Bijection. Require Import Pactole.Setting. Require Import Pactole.Spaces.EuclideanSpace. Require Import Pactole.CaseStudies.Gathering.Definitions. -Require Pactole.CaseStudies.Gathering.WithMultiplicity. +Require Import Pactole.CaseStudies.Gathering.WithMultiplicity. Import Pactole.Observations.MultisetObservation. Require Import Pactole.Models.Rigid. Require Import Pactole.Models.NoByzantine. @@ -35,11 +31,12 @@ Set Implicit Arguments. Close Scope R_scope. Close Scope VectorSpace_scope. Import Datatypes. (* To recover Datatypes.id *) -Typeclasses eauto := (bfs). Remove Hints eq_setoid : typeclass_instances. (* TODO: make equiv not unfolded everywhere. *) Arguments equiv : simpl never. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Section ImpossibilityProof. @@ -56,7 +53,7 @@ Hypothesis nG_non_0 : n <> 0. Local Transparent G B. (** The setting is an arbitrary metric space over R. *) -Context `{Location}. +Context `{Loc:Location}. (* Instance St : State location := OnlyLocation (fun _ => True). *) Context {VS : RealVectorSpace location}. Context {ES : EuclideanSpace location}. @@ -121,7 +118,7 @@ assert (Heven := even_nG). assert (HnG0 := nG_non_0). simpl. destruct n as [| [| ?]]. - lia. - destruct Heven. lia. -- simpl. lia. +- simpl. solve [ auto with arith | lia]. Qed. (* We need to unfold [obs_is_ok] for rewriting *) @@ -130,7 +127,7 @@ Definition obs_from_config_spec : forall config (pt : location), := WithMultiplicity.obs_from_config_spec. Definition mk_info : location -> location := id. -Arguments mk_info _%VectorSpace_scope. +Arguments mk_info _%_VectorSpace_scope. Lemma mk_info_get_location : forall pt, get_location (mk_info pt) == pt. Proof using . reflexivity. Qed. (* @@ -144,6 +141,9 @@ Proof. simpl. repeat intro. now subst. Qed. (* To avoid passing the [nB = 0] argument each time. *) Definition invalid_dec := WithMultiplicity.invalid_dec (reflexivity nB). +Existing Instance app_eqlistA_compat. + + (** From [elements], we can rebuild the [config_list]. *) Lemma obs_makes_config_list : forall config : configuration, PermutationA equiv (List.map get_location (config_list config)) @@ -160,7 +160,10 @@ assert (Hcompat : Proper (PermutationA equiv ==> PermutationA equiv ==> Permutat revert l1 l1' Hl1. pattern l2, l2'. apply PermutationA_ind_bis with equiv; autoclass. + intros [] [] ? ? [Hl Hn] Hperm Hrec ? ? ?. simpl in *. - rewrite Hn, Hl at 1. unfold equiv. now rewrite Hrec. + hnf in Hn. + rewrite Hl at 1. + rewrite Hn. + unfold equiv. now rewrite Hrec. + intros [] [] ? ? Hperm Hrec ? ? ?. simpl. rewrite Hrec, 2 app_assoc; try eassumption; []. f_equiv. apply PermutationA_app_comm; autoclass. + intros ? ? ? Hperm1 Hperm2 Hperm Hrec ? ? ?. @@ -204,7 +207,7 @@ destruct (Nat.eq_dec (from_elements (List.map (fun x : location => (x, 1)) l))[e change (InA eq_pair (e, (from_elements (List.map (fun y => (y, 1)) l))[e]) (elements (from_elements (List.map (fun y => (y, 1)) l)))). rewrite elements_spec; simpl. - split; trivial; []. apply neq_0_lt. auto. + split; trivial; []. apply Nat.neq_0_lt_0. auto. -- revert_all. rewrite removeA_InA; autoclass. tauto. Qed. @@ -244,7 +247,7 @@ Hint Resolve half_size_config : core. Definition lift_config {A} (config : G -> A) : ident -> A := fun id => match id with | Good g => config g - | Byz b => ltac:(exfalso; now apply (Nat.nlt_0_r (proj1_sig b)), proj2_sig) + | Byz b => ltac:(exfalso; now apply (Nat.nlt_0_r (fin2nat b)), fin_lt) end. Local Opaque G B. @@ -312,8 +315,8 @@ assert (Hperm : PermutationA equiv (pt1 :: pt2 :: nil) (pt3 :: pt4 :: nil)). revert Hobs. rewrite 2 support_add; auto; []. destruct (In_dec pt3 (singleton pt4 (Nat.div2 nG))) as [Habs |]; [| destruct (In_dec pt1 (singleton pt2 (Nat.div2 nG))) as [Habs |]]. - - rewrite In_singleton in Habs. now elim Hdiff'. - - rewrite In_singleton in Habs. now elim Hdiff. + - rewrite In_singleton in Habs. now contradiction Hdiff'. + - rewrite In_singleton in Habs. now contradiction Hdiff. - rewrite 2 support_singleton; auto; []. now symmetry. } repeat destruct_match; simpl; rewrite Hgh in *; intro. + assert (Heq1 : pt1 == pt3) by now transitivity (get_location (config2 (Good h))). @@ -364,7 +367,7 @@ assert (Heven : Nat.Even nG). repeat split; trivial; [|]. + rewrite <- Nat.even_spec in Heven. assert (HnG := nG_non_0). simpl nG in *. - destruct n as [| [| ?]]; simpl; discriminate || lia || now elim HnG. + destruct n as [| [| ?]]; simpl; discriminate || lia || now contradiction HnG. + exists (sim origin), (sim one). repeat split. - intro Heq. apply Similarity.injective in Heq. symmetry in Heq. revert Heq. apply non_trivial. @@ -465,20 +468,20 @@ destruct (invalid_dec config) as [Hvalid | Hvalid]. rewrite In_singleton in Hin, Hin'; try solve [ simpl in *; tauto ]; [|]. - destruct Hin' as [Hin' _]. apply Similarity.injective in Hin'. - symmetry in Hin'. elim (non_trivial Hin'). + symmetry in Hin'. contradiction (non_trivial Hin'). - rewrite 2 support_singleton; auto. } rewrite PermutationA_2 in Hperm; autoclass; []. destruct_match. + assert (Hpt1 : sim origin == pt1) by (etransitivity; eauto). assert (Hpt2 : sim one == pt2). - { decompose [and or] Hperm; auto; []. rewrite Hpt1 in *. now elim Hdiff. } + { decompose [and or] Hperm; auto; []. rewrite Hpt1 in *. now contradiction Hdiff. } now rewrite <- Hpt2. + assert (Hpt2 : sim origin == pt2). { destruct (Hcase (Good g)); try contradiction; []. etransitivity; eauto. } assert (Hpt1 : sim one == pt1). - { decompose [and or] Hperm; auto; []. rewrite Hpt2 in *. now elim Hdiff. } + { decompose [and or] Hperm; auto; []. rewrite Hpt2 in *. now contradiction Hdiff. } now rewrite <- Hpt1. -* elim Hvalid. +* contradiction Hvalid. apply (invalid_reverse (build_similarity non_trivial Hdiff)). rewrite Hobs. unfold observation0. rewrite map_add, map_singleton, build_similarity_eq1, build_similarity_eq2; autoclass; []. @@ -504,7 +507,7 @@ assert (Hconfig : round r da1 config == map_config (lift (existT precondition si destruct (Hcase (Good g)) as [Hg' | Hg']; rewrite Hg' in *; solve [ symmetry; apply build_similarity_eq1 | symmetry; apply build_similarity_eq2 - | auto; now elim Hg ]. } + | auto; now contradiction Hg ]. } (* Let us pick an arbitrary robot (here [g0]) and consider a similarity [sim1] that maps [!! config] to [observation0] and such that [sim1 g0 = origin]. *) destruct (invalid_obs Hvalid g0) as [sim1 Hsim1 ?]. @@ -654,13 +657,13 @@ destruct (invalid_dec config) as [Hvalid | Hvalid]. assert (Hin := pos_in_config config origin id). rewrite Hobs', add_In,In_singleton in Hin0, Hin. destruct Hin as [[] | []], Hin0 as [[] | []]; - tauto || elim Hdiff; etransitivity; eauto. } + tauto || contradiction Hdiff; etransitivity; eauto. } exists (symmetry Hdiff). repeat destruct_match; simpl in *; destruct Hcase as [[] | []]; unfold Datatypes.id in *; try solve [ congruence | now apply Hb2; auto - | elim Hdiff'; transitivity (config (Good g0)); auto ]. + | contradiction Hdiff'; transitivity (config (Good g0)); auto ]. + contradiction. Qed. @@ -814,7 +817,10 @@ intros config g Hinvalid. pose (sim := change_frame2 config g). fold sim. unfold observation0. rewrite map_add, map_singleton; autoclass; []. assert (Hdiff_sim : simâ»Â¹ origin =/= simâ»Â¹ one). -{ intro Habs. apply Similarity.injective in Habs. now apply non_trivial. } +{ intro Habs. apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } destruct (WithMultiplicity.invalid_strengthen (reflexivity _) Hinvalid) as [pt1 [pt2 Hdiff Hobs]]. assert (Hin0 : In (simâ»Â¹ origin) (!! config)). { change (In (Similarity.center sim) (!! config)). @@ -836,9 +842,9 @@ assert (Hin1 : In (simâ»Â¹ one) (!! config)). intro pt. rewrite Hobs, 2 add_spec, 2 singleton_spec. repeat destruct_match; solve [ lia - | elim Hdiff; transitivity pt; eauto - | elim Hdiff_sim; transitivity pt; eauto - | elim Hdiff_sim; + | contradiction Hdiff; transitivity pt; eauto + | contradiction Hdiff_sim; transitivity pt; eauto + | contradiction Hdiff_sim; apply (WithMultiplicity.invalid_same_location (reflexivity _) Hinvalid (pt3 := pt)); auto; try (now symmetry); []; rewrite Hobs, add_In, In_singleton; auto | rewrite Hobs, add_In, In_singleton in Hin0, Hin1; @@ -912,7 +918,7 @@ destruct_match_eq Hcase. * (* The robot is on the second tower so it does not move. *) rewrite activate2_spec2 in Hcase; trivial; []. fold sim. - destruct_match; reflexivity || now elim Hcase; etransitivity; eauto. + destruct_match; reflexivity || now contradiction Hcase; etransitivity; eauto. Qed. Lemma invalid_da2_left_next : forall config, @@ -942,7 +948,10 @@ assert (Hconfig : round r (da2_left config) config - rewrite Hobs. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - rewrite Hobs. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - assumption. - - intro Habs. apply Similarity.injective in Habs. now apply non_trivial. } + - intro Habs. apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } rewrite Heq'. cbn -[equiv sim']. rewrite Bijection.section_retraction. unfold sim'. now rewrite build_similarity_eq2. } rewrite Hconfig. @@ -967,7 +976,10 @@ do 2 destruct_match. + split; intro Heq. - etransitivity; eauto. - rewrite change_frame2_eq in Heq; trivial; []. now rewrite Heq. -+ revert_one equiv. intro Heq1. rewrite Heq1. ++ Typeclasses eauto := (bfs). + revert_one equiv. + Typeclasses eauto := (dfs). + intro Heq1. rewrite Heq1. change ((sim â»Â¹) 0%VS) with (Similarity.center sim) in Heq1. unfold sim in Heq1. rewrite center_change_frame2, change_frame2_eq in Heq1; trivial; []. assert (Heq2 : get_location (config (Good g2)) == (sim â»Â¹) one). @@ -976,8 +988,12 @@ do 2 destruct_match. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. now destruct Hin as [[] | []]. } rewrite Heq2, Heq1. fold sim. - split; intro Heq; apply Similarity.injective in Heq; contradiction || now elim non_trivial. + split; intro Heq; apply Similarity.injective in Heq. + * contradiction. + * Typeclasses eauto := (bfs). + now elim non_trivial. + revert_one equiv. intro Heq2. rewrite Heq2. + Typeclasses eauto := (dfs). change ((sim â»Â¹) 0%VS) with (Similarity.center sim) in Heq2. unfold sim in Heq2. rewrite center_change_frame2, change_frame2_eq in Heq2; trivial; []. assert (Heq1 : get_location (config (Good g1)) == (sim â»Â¹) one). @@ -986,7 +1002,11 @@ do 2 destruct_match. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. now destruct Hin as [[] | []]. } rewrite Heq2, Heq1. fold sim. - split; intro Heq; apply Similarity.injective in Heq; (now symmetry in Heq) || now elim non_trivial. + split; intro Heq; apply Similarity.injective in Heq. + * now symmetry in Heq. + * Typeclasses eauto := (bfs). + now elim non_trivial. + Typeclasses eauto := (dfs). + split; intro; solve [ etransitivity; eauto ]. Qed. @@ -1013,18 +1033,25 @@ destruct_match_eq Hcase. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. destruct Hin as [[] | []]; trivial; []. rewrite activate2_spec2 in Hcase; trivial; []. - elim Hcase. etransitivity; eauto; []. now apply center_change_frame2. } + contradiction Hcase. etransitivity; eauto; []. now apply center_change_frame2. } rewrite (obs_from_config_ignore_snd origin). assert (Hobs' : observation0 == map (change_frame2 config g) (!! config)). - { rewrite <- map_id. Time change id with (Bijection.section Similarity.id). + { rewrite <- map_id. + Typeclasses eauto := (bfs). + Time change id with (Bijection.section Similarity.id). + Typeclasses eauto := (dfs). rewrite <- (map_extensionality_compat _ _ (Similarity.compose_inverse_r (change_frame2 config g))). rewrite (change_frame2_obs g Hinvalid), map_merge; autoclass. } rewrite obs_from_config_ignore_snd, <- obs_from_config_map, <- Hobs'; autoclass; []. destruct_match; try reflexivity; []. - fold move. rewrite Hsim0 in *. elim non_trivial. eapply Similarity.injective; eauto. + fold move. rewrite Hsim0 in *. + Typeclasses eauto := (bfs). + contradiction non_trivial. + eapply Similarity.injective; eauto. * (* The robot is on the first tower so it does not move. *) destruct_match; try reflexivity; []. exfalso. revert_one equiv. intro Heq. + Typeclasses eauto := (dfs). rewrite activate2_spec1 in Hcase; trivial; []. contradict Hcase. now rewrite <- (center_change_frame2 g0). Qed. @@ -1044,7 +1071,10 @@ pose (sim1 := change_frame2 config g1). assert (Hdiff : get_location (config (Good g0)) =/= get_location (config (Good g1))). { rewrite Hg1, <- center_change_frame2; trivial; []. unfold Similarity.center. fold sim0. intro Habs. - apply Similarity.injective in Habs. now apply non_trivial. } + apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } assert (Hobs1 := change_frame2_obs g1 Hinvalid). fold sim1 in Hobs1. assert (Hg0 : get_location (config (Good g0)) == (sim1 â»Â¹) one). { apply (WithMultiplicity.invalid_same_location (reflexivity _) Hinvalid (pt3 := (sim1 â»Â¹) 0%VS)). @@ -1052,7 +1082,10 @@ assert (Hg0 : get_location (config (Good g0)) == (sim1 â»Â¹) one). - rewrite Hobs1. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - rewrite Hobs1. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - change ((sim1 â»Â¹) 0%VS) with (Similarity.center sim1). unfold sim1. now rewrite center_change_frame2. - - intro Habs. apply Similarity.injective in Habs. now apply non_trivial. } + - intro Habs. apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } assert (Hdiff_move0 : sim0â»Â¹ move =/= sim0â»Â¹ one). { intro Heq. now apply Similarity.injective in Heq. } assert (Hdiff_move1 : sim1â»Â¹ move =/= sim1â»Â¹ one). @@ -1074,7 +1107,10 @@ assert (Hconfig : round r (da2_right config) config - rewrite Hobs0. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - rewrite Hobs0. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. - assumption. - - intro Habs. apply Similarity.injective in Habs. now apply non_trivial. } + - intro Habs. apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } rewrite Heq'. cbn -[equiv sim']. rewrite Bijection.section_retraction. unfold sim'. rewrite build_similarity_eq1. rewrite <- Hg1, change_frame2_eq in Heq'; trivial; []. now rewrite Heq'. } @@ -1105,10 +1141,15 @@ pose (sim1 := change_frame2 config g3). assert (Hdiff : get_location (config (Good g0)) =/= get_location (config (Good g3))). { rewrite Hg3, <- center_change_frame2; trivial; []. unfold Similarity.center. fold sim0. intro Habs. - apply Similarity.injective in Habs. now apply non_trivial. } + apply Similarity.injective in Habs. + Typeclasses eauto := (bfs). + now apply non_trivial. + Typeclasses eauto := (dfs). } do 2 destruct_match. + reflexivity. -+ revert_one equiv. intro Heq1. rewrite Heq1. ++ Typeclasses eauto := (bfs). + revert_one equiv. intro Heq1. rewrite Heq1. + Typeclasses eauto := (dfs). split; intro Heq. - destruct (change_frame2_case g3 Hinvalid (symmetry Hdiff) g2) as [Hcase | Hcase]. * rewrite <- Hcase. now apply center_change_frame2. @@ -1119,11 +1160,13 @@ do 2 destruct_match. { rewrite Hobs0. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. } rewrite (change_frame2_obs g3 Hinvalid) in Hin. fold sim1 in Hin. unfold observation0 in Hin. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass. - decompose [and or] Hin; trivial; elim Hdiff; []. now rewrite <- 2 center_change_frame2. } + decompose [and or] Hin; trivial; contradiction Hdiff; []. now rewrite <- 2 center_change_frame2. } rewrite Hcase in Heq. fold sim0 sim1 in Heq. rewrite Heq in Habs. apply Similarity.injective in Habs. contradiction. - symmetry in Heq. contradiction. -+ revert_one equiv. intro Heq2. rewrite Heq2. ++ Typeclasses eauto := (bfs). + revert_one equiv. intro Heq2. rewrite Heq2. + Typeclasses eauto := (dfs). change ((change_frame2 config g0 â»Â¹) 0%VS) with (Similarity.center (change_frame2 config g0)) in Heq2. rewrite center_change_frame2, change_frame2_eq in Heq2; trivial; []. assert (Heq1 : get_location (config (Good g1)) == (sim0 â»Â¹) one). @@ -1132,7 +1175,10 @@ do 2 destruct_match. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. now destruct Hin as [[] | []]. } rewrite Heq1. - split; intro Heq; try (apply Similarity.injective in Heq; now elim non_trivial); []. + + Typeclasses eauto := (bfs). + split; intro Heq; try (apply Similarity.injective in Heq; now contradiction non_trivial); []. + Typeclasses eauto := (dfs). destruct (change_frame2_case g3 Hinvalid (symmetry Hdiff) g1) as [Hcase | Hcase]. * rewrite <- Hcase, <- Heq1. symmetry. now apply center_change_frame2. * change ((change_frame2 config g0 â»Â¹) 0%VS) with (Similarity.center (change_frame2 config g0)). @@ -1142,7 +1188,7 @@ do 2 destruct_match. { rewrite Hobs0. unfold observation0. rewrite map_add, map_singleton, add_In, In_singleton; autoclass. } rewrite (change_frame2_obs g3 Hinvalid) in Hin. unfold observation0 in Hin. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. - decompose [and or] Hin; trivial; elim Hdiff; []. now rewrite <- 2 center_change_frame2. } + decompose [and or] Hin; trivial; contradiction Hdiff; []. now rewrite <- 2 center_change_frame2. } rewrite Hcase in Heq. fold sim0 sim1 in Heq. rewrite <- Heq in Habs. apply Similarity.injective in Habs. contradiction. + assert (Heq1 : get_location (config (Good g1)) == (sim0 â»Â¹) one). diff --git a/CaseStudies/Gathering/InR/Algorithm.v b/CaseStudies/Gathering/InR/Algorithm.v index 96b5f0f185f2f323af87b9ef79ae90c2e548693f..09fe0415aace3c749d129f095d960111fca270b2 100644 --- a/CaseStudies/Gathering/InR/Algorithm.v +++ b/CaseStudies/Gathering/InR/Algorithm.v @@ -18,7 +18,6 @@ (**************************************************************************) Require Import Bool. -Require Import Arith.Div2. Require Import Lia. Require Import Rbase Rbasic_fun. Require Import List SetoidList. @@ -26,17 +25,17 @@ Require Import RelationPairs. Require Import Morphisms. Require Import Psatz. Require Import Inverse_Image. -(* Pactole basic definitions *) -Require Export Pactole.Setting. Require Import FMapFacts. (* Specific to R topology *) Require Import Pactole.Spaces.R. (* Specific to gathering *) -Require Pactole.CaseStudies.Gathering.WithMultiplicity. -(* I don't like this Import, but gathered_at is too frequent *) -Require Import Pactole.CaseStudies.Gathering.Definitions. +Require Import Pactole.CaseStudies.Gathering.WithMultiplicity. +Require Import Pactole.Core.State. (* Specific to multiplicity *) Require Import Pactole.Observations.MultisetObservation. +Require Import Pactole.Util.NumberComplements. +(* I don't like this Import, but gathered_at is too frequent *) +Require Import Pactole.CaseStudies.Gathering.Definitions. Require Import Pactole.Models.Similarity. (* Specific to rigidity *) Require Export Pactole.Models.Rigid. @@ -49,9 +48,11 @@ Import Datatypes. Close Scope R_scope. (* rule of thumb *) -Typeclasses eauto := 10. Set Implicit Arguments. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. + (* Now we declare the context of our proof: things left abstract are variables and hypothesis. Other things are defined by declaring type class instances. *) @@ -89,6 +90,8 @@ Instance InaFun : inactive_function unit := { Instance Rigid : RigidSetting. Proof using . split. reflexivity. Qed. +Existing Instance multiset_observation. + (* Trying to avoid notation problem with implicit arguments *) Notation "s [ x ]" := (multiplicity x s) (at level 2, no associativity, format "s [ x ]"). Notation obs_from_config := (@obs_from_config _ _ _ _ multiset_observation). @@ -102,7 +105,8 @@ Arguments origin : simpl never. Ltac changeR := change R with location in *; change R_Setoid with location_Setoid in *; - change R_EqDec with location_EqDec in *. + change R_EqDec with location_EqDec in *; + change SetoidDefs.R_EqDec with (@location_EqDec Loc) in *. Lemma similarity_middle : forall (sim : similarity R) x y, (sim ((x + y) / 2) = (sim x + sim y) / 2)%R. Proof using . @@ -128,7 +132,7 @@ apply Permutation_nil. setoid_rewrite Permuted_sort at 2. rewrite Habs. reflexiv Qed. Lemma half_size_config : Nat.div2 nG > 0. -Proof using size_G. assert (Heven := size_G). simpl. destruct n as [| [| ?]]; simpl; lia. Qed. +Proof using size_G. assert (Heven := size_G). simpl. destruct n as [| [| ?]]; simpl; solve [ auto with arith | lia]. Qed. (* We need to unfold [obs_is_ok] for rewriting *) Definition obs_from_config_spec : forall config l, @@ -179,9 +183,9 @@ intros config pt. split; intro Hmaj. + intro y. rewrite InA_singleton. rewrite support_spec, max_spec1_iff; try (now apply obs_non_nil); []. split; intro Hpt. - - rewrite Hpt. intro x. destruct (Rdec x pt). + - rewrite Hpt. intro x. destruct (Req_dec_T x pt). -- subst x. reflexivity. - -- apply lt_le_weak. now apply (Hmaj x). + -- apply Nat.lt_le_incl. now apply (Hmaj x). - destruct (Rdec y pt) as [? | Hy]; trivial. exfalso. apply (Hmaj y) in Hy. specialize (Hpt pt). simpl in *. lia. * intros x Hx. apply max_spec_lub. @@ -191,32 +195,32 @@ Qed. (** ** Some properties of [invalid] **) -Lemma invalid_even : forall conf, WithMultiplicity.invalid conf -> Nat.Even nG. +Lemma invalid_even : forall conf, (@WithMultiplicity.invalid Loc _ _ conf) -> Nat.Even nG. Proof using . now intros conf [? _]. Qed. Lemma invalid_support_length : forall config, WithMultiplicity.invalid config -> size (!! config) = 2. Proof using size_G. intros config [Heven [_ [pt1 [pt2 [Hdiff [Hpt1 Hpt2]]]]]]. -rewrite <- (@cardinal_total_sub_eq _ _ _ _ _ (add pt2 (Nat.div2 nG) (singleton pt1 (Nat.div2 nG)))). +rewrite <- (@cardinal_total_sub_eq _ _ _ _ _ (MMultisetInterface.add pt2 (Nat.div2 nG) (singleton pt1 (Nat.div2 nG)))). * rewrite size_add; try (now apply half_size_config); []. destruct (In_dec pt2 (singleton pt1 (Nat.div2 nG))) as [Hin | Hin]. + exfalso. rewrite In_singleton in Hin. - destruct Hin. now elim Hdiff. + destruct Hin. now contradiction Hdiff. + rewrite size_singleton; trivial. apply half_size_config. -* intro pt. destruct (Rdec pt pt2), (Rdec pt pt1); subst. - + now elim Hdiff. +* intro pt. destruct (Req_dec_T pt pt2), (Req_dec_T pt pt1); subst. + + now contradiction Hdiff. + rewrite add_spec, singleton_spec. setoid_rewrite Hpt2. repeat destruct_match; simpl in *; contradiction || (try lia). + rewrite add_other, singleton_spec. - - setoid_rewrite Hpt1. repeat destruct_match; simpl in *; contradiction || lia. + - changeR. setoid_rewrite Hpt1. repeat destruct_match; simpl in *; contradiction || lia. - assumption. + rewrite add_other, singleton_spec. - repeat destruct_match; simpl in *; contradiction || lia. - assumption. * rewrite cardinal_add, cardinal_singleton, cardinal_obs_from_config. - simpl. rewrite plus_0_r. now apply even_div2. + simpl. rewrite Nat.add_0_r. now apply even_div2. Qed. Lemma support_max_1_not_invalid : forall config pt, @@ -236,17 +240,17 @@ assert (Hsup : Permutation (support (!! config)) (pt1 :: pt2 :: nil)). assert (Hin2 : InA equiv pt2 (support (!! config))). { rewrite support_spec. unfold In. setoid_rewrite Hpt2. apply half_size_config. } apply (PermutationA_split _) in Hin1. destruct Hin1 as [l Hl]. rewrite Hl in Hin2. - inversion_clear Hin2; try now subst; elim Hdiff. + inversion_clear Hin2; try now subst; contradiction Hdiff. rewrite size_spec, Hl in Hsuplen. destruct l as [| x [| ? ?]]; simpl in Hsuplen; try lia. inversion_clear H; (now inversion H0) || (cbn in H0; subst). now rewrite <- PermutationA_Leibniz. } assert (Hpt : pt = pt1 \/ pt = pt2). { assert (Hin : List.In pt (pt1 :: pt2 :: nil)). { rewrite <- Hsup, <- InA_Leibniz. change eq with (@equiv location _). + changeR. rewrite support_spec, <- max_subset, <- support_spec, Hmaj. now left. } inversion_clear Hin; auto. inversion_clear H; auto. inversion H0. } -apply (lt_irrefl (Nat.div2 nG)). destruct Hpt; subst pt. -- rewrite <- Hpt1 at 2. rewrite <- Hpt2. apply max_spec_lub; try (now rewrite Hmax); []. - rewrite Hmax. auto. +apply (Nat.lt_irrefl (Nat.div2 nG)). destruct Hpt; subst pt. +- rewrite <- Hpt1 at 2. rewrite <- Hpt2. apply max_spec_lub; try (now rewrite Hmax). - rewrite <- Hpt1 at 1. rewrite <- Hpt2. apply max_spec_lub; now rewrite Hmax. Qed. @@ -289,7 +293,7 @@ Proof using size_G. intro abs. subst. inversion hnodup;subst. - elim H1. + contradiction H1. constructor. reflexivity. * assert (h : inclA equiv (support (max (!! config))) (support (!! config))). @@ -314,13 +318,13 @@ Proof using size_G. rewrite Hsupp in hnodup. inversion hnodup;subst. match goal with - | H: ~ InA equiv pt2 (pt2 :: nil) |- _ => elim H + | H: ~ InA equiv pt2 (pt2 :: nil) |- _ => contradiction H end. constructor 1. reflexivity. } assert (heq_config: equiv (!!config) - (add pt1 ((!! config)[pt1]) - (add pt2 ((!! config)[pt2]) empty))). + (MMultisetInterface.add pt1 ((!! config)[pt1]) + (MMultisetInterface.add pt2 ((!! config)[pt2]) empty))). { intros x. destruct (equiv_dec x pt1) as [heqxpt1 | hneqxpt1]. - rewrite heqxpt1. @@ -355,7 +359,7 @@ Proof using size_G. - intros ? ? ? ? ? Heq; subst; now rewrite Heq. - intros. lia. - symmetry. - transitivity ((pt1, (!! config)[pt1]) :: (elements (add pt2 (!! config)[pt2] empty))). + transitivity ((pt1, (!! config)[pt1]) :: (elements (MMultisetInterface.add pt2 (!! config)[pt2] empty))). eapply elements_add_out;auto. + rewrite heq_config, add_same. cut ((!! config)[pt1] > 0). lia. change (In pt1 (!! config)). rewrite <- support_spec, Hsupp. now left. @@ -425,7 +429,7 @@ Lemma not_invalid_not_majority_length : forall config, Proof using size_G. intros config H1 H2. assert (size (!! config) > 1)%nat. -{ unfold gt. eapply lt_le_trans; try eassumption. +{ unfold gt. eapply Nat.lt_le_trans; try eassumption. f_equiv. apply max_subset. } destruct (size (!! config)) as [| [| [| ?]]] eqn:Hlen; try lia. exfalso. apply H2. now rewrite invalid_equiv. @@ -481,10 +485,10 @@ destruct Hfmon as [Hfinc | Hfdec]. - rewrite map_last in Heq. apply Hfinj in Heq. contradiction. - intro Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). + rewrite (hd_indep _ (f 0)) in Hneq. - - elim Hneq. rewrite map_hd. now f_equal. + - contradiction Hneq. rewrite map_hd. hnf. now f_equal. - intro Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). + rewrite (last_indep _ (f 0)) in Hneq0. - - elim Hneq0. rewrite map_last. now f_equal. + - contradiction Hneq0. rewrite map_last. hnf. now f_equal. - intro Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). * repeat Rdec_full; trivial; rewrite map_injective_support, (sort_map_decreasing Hfdec) in Heq @@ -499,10 +503,10 @@ destruct Hfmon as [Hfinc | Hfdec]. - rewrite last_rev_hd, map_hd in Heq. apply Hfinj in Heq. contradiction. - intro Habs. rewrite rev_nil in Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). + rewrite (last_indep _ (f 0)) in Hneq0. - - elim Hneq0. rewrite last_rev_hd, map_hd. now f_equal. + - contradiction Hneq0. rewrite last_rev_hd, map_hd. hnf. now f_equal. - intro Habs. rewrite rev_nil in Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). + rewrite (hd_indep _ (f 0)) in Hneq. - - elim Hneq. rewrite hd_rev_last, map_last. now f_equal. + - contradiction Hneq. rewrite hd_rev_last, map_last. hnf. now f_equal. - intro Habs. rewrite rev_nil in Habs. apply map_eq_nil in Habs. now apply (sort_non_nil config). Qed. @@ -520,7 +524,7 @@ Definition extreme_center (s : observation) := (mini s + maxi s) / 2. Instance extreme_center_compat : Proper (equiv ==> eq) extreme_center. Proof using . intros s s' Hs. unfold extreme_center, mini, maxi. now rewrite Hs. Qed. -Lemma extreme_center_similarity : forall (sim : similarity location) s, s =/= empty -> +Lemma extreme_center_similarity : forall (sim : (@similarity location (@state_Setoid Loc (@location Loc) Info) (@state_EqDec Loc (@location Loc) Info) _ _)) (s:multiset location), s =/= empty -> extreme_center (map sim s) = sim (extreme_center s). Proof using . intros sim s Hs. @@ -531,10 +535,10 @@ destruct (similarity_monotonic sim) as [Hinc | Hdec]. * rewrite map_injective_support, (sort_map_increasing Hinc); trivial; []. assert (Hperm := Permuted_sort (support s)). changeR. destruct (sort (support s)) as [| x l']. - + symmetry in Hperm. apply Permutation_nil in Hperm. elim Hs. now rewrite <- support_nil. + + symmetry in Hperm. apply Permutation_nil in Hperm. contradiction Hs. now rewrite <- support_nil. + clear s Hs Hperm. simpl hd. cut (x :: l' <> nil). generalize (x :: l'). intro l. generalize 0. induction l; intros r Hl. - - now elim Hl. + - now contradiction Hl. - simpl. destruct l. simpl. symmetry. now apply similarity_middle. rewrite <- IHl. reflexivity. discriminate. @@ -543,10 +547,10 @@ destruct (similarity_monotonic sim) as [Hinc | Hdec]. rewrite last_rev_hd, hd_rev_last. assert (Hperm := Permuted_sort (support s)). changeR. destruct (sort (support s)) as [| x l']. - + symmetry in Hperm. apply Permutation_nil in Hperm. elim Hs. now rewrite <- support_nil. + + symmetry in Hperm. apply Permutation_nil in Hperm. contradiction Hs. now rewrite <- support_nil. + clear s Hs Hperm. simpl hd. cut (x :: l' <> nil). generalize (x :: l'). intro l. generalize 0. induction l; intros r Hl. - - now elim Hl. + - now contradiction Hl. - simpl. destruct l. -- simpl. rewrite similarity_middle. now rewrite Rplus_comm. -- rewrite <- IHl. reflexivity. discriminate. @@ -597,7 +601,7 @@ Definition gatherR_pgm (s : observation) : location := | nil => 0 (* only happen with no robots *) | pt :: nil => pt (* case 1: one majority stack *) | _ => (* several majority stacks *) - if beq_nat (size s) 3 + if Nat.eqb (size s) 3 then middle s else if is_extremal 0 s then 0 else extreme_center s end. @@ -634,7 +638,7 @@ Lemma round_simplify : forall config, | nil => config id (* only happen with no robots *) | pt :: nil => pt (* case 1: one majority stack *) | _ => (* several majority stacks *) - if beq_nat (size s) 3 + if Nat.eqb (size s) 3 then middle s else if is_extremal (config id) s then config id else extreme_center s end @@ -794,9 +798,10 @@ Theorem MajTower_at_forever : forall pt config, MajTower_at pt config -> MajTower_at pt (round gatherR da config). Proof using Hda size_G. intros pt config Hmaj x Hx. assert (Hs := Hmaj x Hx). -apply le_lt_trans with ((!! config)[x]); try eapply lt_le_trans; try eassumption; [|]. +apply Nat.le_lt_trans with ((!! config)[x]). - eapply Majority_wither; eauto. -- eapply Majority_grow; eauto. +- apply Nat.lt_le_trans with (!! config)[pt]; try eassumption; []. + eapply Majority_grow; eauto. Qed. Theorem Majority_not_invalid : forall pt config, MajTower_at pt config -> ~WithMultiplicity.invalid config. @@ -826,7 +831,7 @@ Lemma Generic_min_max_lt : forall config, no_Majority config -> mini (!! config) < maxi (!! config). Proof using size_G. intros config Hmaj. apply Generic_min_max_lt_aux. -+ apply lt_le_trans with (size (max (!! config))); trivial. ++ apply Nat.lt_le_trans with (size (max (!! config))); trivial. rewrite <- size_spec. f_equiv. apply max_subset. + apply support_NoDupA. Qed. @@ -855,7 +860,7 @@ assert (Heq : config id == round gatherR da config id). { rewrite (round_simplify_Generic Hmaj Hlen id); trivial; []. destruct (da.(activate) id); try reflexivity; []. unfold is_extremal. Rdec_full; try reflexivity; []. - elim Hneq. rewrite Hid. apply hd_indep. apply sort_non_nil. } + contradiction Hneq. } (** Main proof *) apply Rle_antisym. * apply sort_min. @@ -898,7 +903,7 @@ assert (Heq : config id == round gatherR da config id). { rewrite (round_simplify_Generic Hmaj Hlen id). destruct (da.(activate) id); try reflexivity; []. unfold is_extremal. repeat Rdec_full; try reflexivity; []. - elim Hneq0. rewrite Hid. apply last_indep. apply sort_non_nil. } + contradiction Hneq0. } (** Main proof *) apply Rle_antisym. * apply sort_max. @@ -945,16 +950,22 @@ do 2 rewrite obs_from_config_spec, config_list_spec. induction names as [| id l]; simpl in *; unfold Datatypes.id in *. * reflexivity. * changeR. destruct (activate da id). - + destruct_match_eq Hext; repeat destruct_match. + + changeR. destruct_match_eq Hext;repeat destruct_match. - f_equal. apply IHl. - apply IHl. - - elim (Rlt_irrefl (mini (!! config))). + - contradiction (Rlt_irrefl (mini (!! config))). + changeR. match goal with H : extreme_center _ == _ |- _ => rewrite <- H at 2 end. now apply Generic_min_mid_lt. - - elim (Rlt_irrefl (mini (!! config))). + - contradiction (Rlt_irrefl (mini (!! config))). match goal with H : extreme_center _ == _ |- _ => rewrite <- H at 2 end. now apply Generic_min_mid_lt. - - exfalso. revert Hext. unfold is_extremal. repeat destruct_match; tauto || discriminate. + - exfalso. revert Hext. unfold is_extremal. + rewrite e. + repeat destruct_match. + -- discriminate. + -- discriminate. + -- intros _. now apply c0. - apply IHl. + destruct_match. - f_equal. apply IHl. @@ -976,13 +987,18 @@ induction names as [| id l]; simpl. + destruct_match_eq Hext; repeat destruct_match. - f_equal. apply IHl. - apply IHl. - - elim (Rlt_irrefl (maxi (!! config))). + - contradiction (Rlt_irrefl (maxi (!! config))). match goal with H : extreme_center _ == _ |- _ => rewrite <- H at 1 end. now apply Generic_mid_max_lt. - - elim (Rlt_irrefl (maxi (!! config))). + - contradiction (Rlt_irrefl (maxi (!! config))). match goal with H : extreme_center _ == _ |- _ => rewrite <- H at 1 end. now apply Generic_mid_max_lt. - - exfalso. revert Hext. unfold is_extremal. repeat destruct_match; tauto || discriminate. + - exfalso. revert Hext. unfold is_extremal. + rewrite e. + repeat destruct_match;intros ?. + -- discriminate. + -- discriminate. + -- now apply c1. - apply IHl. + destruct_match. - f_equal. apply IHl. @@ -1018,7 +1034,7 @@ destruct (da.(activate) id1) eqn:Hmove1; [destruct (da.(activate) id2) eqn:Hmove cbv zeta. destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. + (* no robots *) - rewrite support_nil, max_is_empty in Hmaj. elim (obs_non_nil _ Hmaj). + rewrite support_nil, max_is_empty in Hmaj. contradiction (obs_non_nil _ Hmaj). + (* a majority tower *) reflexivity. + destruct (size (!! config) =? 3) eqn:Hlen. @@ -1053,7 +1069,7 @@ Proof using size_G. intros config pt1 pt2 Hlen Hpt1 Hpt2 Hdiff12. rewrite <- support_spec in Hpt1, Hpt2. rewrite size_spec in Hlen. apply (PermutationA_split _) in Hpt1. destruct Hpt1 as [supp1 Hperm]. -rewrite Hperm in Hpt2. inversion_clear Hpt2; try (now elim Hdiff12); []. rename H into Hpt2. +rewrite Hperm in Hpt2. inversion_clear Hpt2; try (now contradiction Hdiff12); []. rename H into Hpt2. apply (PermutationA_split _) in Hpt2. destruct Hpt2 as [supp2 Hperm2]. rewrite Hperm2 in Hperm. rewrite Hperm in Hlen. destruct supp2 as [| pt3 supp]; try (now simpl in Hlen; lia); []. @@ -1085,7 +1101,7 @@ Proof using . assert (Hg : forall id, get_location (round r da config id) <> pt \/ get_location (config id) = pt). { intro id. specialize (Hex id (In_names _)). revert Hex. repeat destruct_match; auto. } (** We prove a contradiction by showing that the opposite inequality of Hlt holds. *) - clear Hex. revert Hlt. apply le_not_lt. + clear Hex. revert Hlt. apply Nat.le_ngt. do 2 rewrite obs_from_config_spec, config_list_spec. induction names as [| id l]; simpl; trivial; []. unfold Datatypes.id in *. repeat destruct_match; auto using le_n_S; []. @@ -1108,7 +1124,7 @@ intros config pt. split. { intros id' Hid'. destruct (get_location (round gatherR da config id') =?= pt) as [Heq | Heq]; trivial; []. assert (Habs := Heq). rewrite <- Hid' in Habs. assert (Habs' : round gatherR da config id' =/= config id'). - { intro Habs'. rewrite Habs' in Habs. now elim Habs. } + { intro Habs'. rewrite Habs' in Habs. now contradiction Habs. } rewrite <- (moving_spec gatherR) in Habs'. apply Hdest in Habs'. contradiction. } do 2 rewrite obs_from_config_spec, config_list_spec. assert (Hin : List.In id names) by apply In_names. @@ -1116,13 +1132,13 @@ intros config pt. split. inversion_clear Hin. + subst id'. clear IHl. simpl. hnf in Hroundid. unfold Datatypes.id. destruct_match. revert_one @equiv. intro Heq. - - rewrite <- Hid in Heq. elim Hroundid. now apply WithMultiplicity.no_info. + - rewrite <- Hid in Heq. contradiction Hroundid. changeR. now apply WithMultiplicity.no_info. - destruct_match; try contradiction; []. apply le_n_S. induction l; simpl. -- reflexivity. -- repeat destruct_match; try now idtac + apply le_n_S + apply le_S; apply IHl. - revert_one @complement. intro Hneq. elim Hneq. now apply Hstay. + revert_one @complement. intro Hneq. contradiction Hneq. now apply Hstay. + apply IHl in H. simpl in *. repeat destruct_match; try lia; []. - revert_one @complement. intro Hneq. elim Hneq. + revert_one @complement. intro Hneq. contradiction Hneq. apply Hdest. rewrite moving_spec. intro Habs. apply Hneq. now rewrite Habs. Qed. @@ -1161,7 +1177,7 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. { assert (Hperm : Permutation (support (!! (round gatherR da config))) (pt1 :: pt2 :: nil)). { symmetry. apply NoDup_Permutation_bis. + repeat constructor. - - intro Habs. inversion Habs. now elim Hdiff. now inversion H. + - intro Habs. inversion Habs. now contradiction Hdiff. now inversion H. - intro Habs. now inversion Habs. (* NoDupA_Leibniz had a useless hyp in coq stdlib until april 2020. *) (* + rewrite <- NoDupA_Leibniz. change eq with (@equiv location _). apply support_NoDupA. *) @@ -1184,7 +1200,7 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. { intros id Hid. rewrite <- Hrmove_pt. apply same_destination; auto. rewrite moving_spec. congruence. } - assert ((div2 nG) <= (!! config)[pt']). + assert ((Nat.div2 nG) <= (!! config)[pt']). { transitivity ((!! (round gatherR da config))[pt']). - decompose [and or] Hpt; subst. + setoid_rewrite Hpt2. reflexivity. @@ -1196,18 +1212,18 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. rewrite <- Hid1. symmetry. apply Hdest. rewrite moving_spec. assumption. } - assert (Hlt : forall p, p <> pt' -> (!! config)[p] < div2 nG). + assert (Hlt : forall p, p <> pt' -> (!! config)[p] < Nat.div2 nG). { assert (Hpt'_in : In pt' (!! config)). - { unfold In. eapply lt_le_trans; try eassumption. apply half_size_config. } + { unfold In. eapply Nat.lt_le_trans; try eassumption. apply half_size_config. } assert (Hle := not_invalid_not_majority_length Hmaj Hok). - intros p Hp. apply Nat.nle_gt. intro Habs. apply (lt_irrefl nG). + intros p Hp. apply Nat.nle_gt. intro Habs. apply (Nat.lt_irrefl nG). destruct (@towers_elements_3 config pt' p Hle Hpt'_in) as [pt3' [Hdiff13 [Hdiff23 Hpt3_in]]]; trivial. - + apply lt_le_trans with (div2 nG); try assumption. apply half_size_config. + + apply Nat.lt_le_trans with (Nat.div2 nG); try assumption. apply half_size_config. + auto. - + eapply lt_le_trans; try apply (sum3_le_total config Hp Hdiff13 Hdiff23); []. + + eapply Nat.lt_le_trans; try apply (sum3_le_total config Hp Hdiff13 Hdiff23); []. unfold In in Hpt3_in. rewrite <- (even_div2 _ HnG). lia. } assert (Hmaj' : MajTower_at pt' config). - { intros x Hx. apply lt_le_trans with (div2 nG); trivial. now apply Hlt. } + { intros x Hx. apply Nat.lt_le_trans with (Nat.div2 nG); trivial. now apply Hlt. } apply MajTower_at_forever, Majority_not_invalid in Hmaj'. contradiction. } Qed. @@ -1220,7 +1236,7 @@ Definition config_to_NxN config := | nil => (0, 0) | pt :: nil => (1, nG - s[pt]) | _ :: _ :: _ => - if beq_nat (size s) 3 + if Nat.eqb (size s) 3 then (2, nG - s[nth 1 (sort (support s)) 0%R]) else (3, nG - (s[extreme_center s] + s[hd 0%R (sort (support s))] @@ -1338,7 +1354,7 @@ assert (Hstep : da.(activate) gmove = true). rewrite moving_spec in Hmove. destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. * (* No robots *) - elim (support_max_non_nil _ Hmaj). + contradiction (support_max_non_nil _ Hmaj). * (* A majority tower *) rewrite <- MajTower_at_equiv in Hmaj. assert (Hmaj' : MajTower_at pt (round gatherR da config)) by now apply MajTower_at_forever. @@ -1349,7 +1365,7 @@ destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. cut ((!! config)[pt] < (!! (round gatherR da config))[pt]). lia. assert (Hdestg : get_location (round gatherR da config gmove) = pt). { rewrite (round_simplify_Majority Hmaj gmove). - destruct (da.(activate) gmove); trivial; now elim Hstep. } + destruct (da.(activate) gmove); reflexivity || contradiction diff_false_true. } rewrite increase_move_iff. eauto. * rename Hmaj into Hmaj'. assert (Hmaj : no_Majority config). @@ -1365,7 +1381,7 @@ destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. assert (Hmaj' : no_Majority (round gatherR da config)). { unfold no_Majority. rewrite size_spec, Hmaj''. simpl. lia. } clear Hmaj''. assert (Hlen' : size (!! (round gatherR da config)) = 3). - { apply le_antisym. + { apply Nat.le_antisymm. + apply (support_decrease Hmaj Hlen). + apply (not_invalid_not_majority_length Hmaj'). now apply never_invalid. } rewrite (config_to_NxN_Three_spec Hmaj' Hlen'). apply Lexprod.right_lex. @@ -1385,7 +1401,7 @@ destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. unfold gt. rewrite increase_move_iff. exists gmove. split; trivial; []. rewrite (round_simplify_Three Hmaj Hlen gmove). - destruct (da.(activate) gmove); reflexivity || now elim Hstep. + destruct (da.(activate) gmove); reflexivity || contradiction diff_false_true. + (* Generic case *) red. rewrite (config_to_NxN_Generic_spec Hmaj Hlen). destruct (support (max (!! (round gatherR da config)))) as [| pt' [| ? ?]] eqn:Hmaj'. @@ -1414,9 +1430,9 @@ destruct (support (max (!! config))) as [| pt [| ? ?]] eqn:Hmaj. lia. rewrite increase_move_iff. exists gmove. split; trivial. rewrite (round_simplify_Generic Hmaj Hlen gmove) in Hmove |- *; trivial; []. - destruct (da.(activate) gmove); try (now elim Hstep); []. + destruct (da.(activate) gmove); [| contradiction diff_false_true]. destruct (is_extremal (config gmove) (!! config)). - -- now elim Hmove. + -- now contradiction Hmove. -- reflexivity. Qed. @@ -1453,8 +1469,10 @@ Lemma not_gathered_exists : forall config pt, Proof using . intros config pt Hgather. destruct (forallb (fun x => if get_location (config x) =?= pt then true else false) names) eqn:Hall. -- elim Hgather. rewrite forallb_forall in Hall. - intro id'. setoid_rewrite Rdec_bool_true_iff in Hall. hnf. repeat rewrite Hall; trivial; apply In_names. +- contradiction Hgather. rewrite forallb_forall in Hall. + intro id'. changeR. + Print Instances Proper. + setoid_rewrite Rdec_bool_true_iff in Hall. hnf. repeat rewrite Hall; trivial; apply In_names. - rewrite <- negb_true_iff, existsb_forallb, existsb_exists in Hall. destruct Hall as [id' [_ Hid']]. revert Hid'. destruct_match; discriminate || intro. now exists id'. Qed. @@ -1465,7 +1483,7 @@ Lemma not_invalid_gathered_Majority_size : forall config id, Proof using size_G. intros config id Hinvalid Hgather Hmaj. assert (size (!! config) > 1). -{ unfold no_Majority in Hmaj. eapply lt_le_trans; try eassumption; []. now rewrite max_subset. } +{ unfold no_Majority in Hmaj. eapply Nat.lt_le_trans; try eassumption; []. now rewrite max_subset. } rewrite invalid_equiv in Hinvalid. destruct (size (!! config)) as [| [| [| ?]]]; lia || tauto. Qed. @@ -1476,7 +1494,7 @@ Theorem OneMustMove : forall config id, ~ WithMultiplicity.invalid config -> ~ga Proof using size_G. intros config id Hinvalid Hgather. destruct (support (max (!! config))) as [| pt [| pt' l]] eqn:Hmaj. -* elim (support_max_non_nil _ Hmaj). +* contradiction (support_max_non_nil _ Hmaj). * rewrite <- MajTower_at_equiv in Hmaj. apply not_gathered_generalize with _ _ pt in Hgather. apply not_gathered_exists in Hgather. destruct Hgather as [gmove Hmove]. @@ -1508,7 +1526,7 @@ destruct (support (max (!! config))) as [| pt [| pt' l]] eqn:Hmaj. assert (Hnodup := support_NoDupA (!! config)). rewrite NoDupA_Leibniz, Permuted_sort, Hsup in Hnodup. inversion_clear Hnodup. inversion_clear H0. inversion_clear H2. - destruct (Rdec pt2 (extreme_center (!! config))) as [Heq | Heq]; subst. + destruct (Rdec pt2 (extreme_center (!! config))) as [Heq | Heq]; simpl in Heq;subst. - exists pt3. repeat split; try now intro; subst; intuition. rewrite <- support_spec, InA_Leibniz, Permuted_sort, Hsup. intuition. - exists pt2. repeat split; try now intro; subst; intuition. @@ -1558,7 +1576,7 @@ intros da Hda config pt Hgather. rewrite (round_simplify_Majority). induction names as [| id l]. + reflexivity. + simpl. destruct_match. - - elim Hdiff. simpl in *. subst. apply (no_byz id), Hgather. + - contradiction Hdiff. simpl in *. subst. apply (no_byz id), Hgather. - apply IHl. } rewrite H0. specialize (Hgather g1). rewrite <- Hgather. apply pos_in_config. Qed. diff --git a/CaseStudies/Gathering/InR/Impossibility.v b/CaseStudies/Gathering/InR/Impossibility.v index 751914ecf63c85e67c42efff85a14e79d7f05682..a26c70346ae574dcd60a774d8e953f37df75024b 100644 --- a/CaseStudies/Gathering/InR/Impossibility.v +++ b/CaseStudies/Gathering/InR/Impossibility.v @@ -20,10 +20,10 @@ Require Import Reals. Require Import Psatz. Require Import Morphisms. -Require Import Arith.Div2. Require Import Lia. Require Import List SetoidList. -Require Import Pactole.Util.Preliminary. +Require Import Pactole.Util.Preliminary Pactole.Util.Fin. +Require Import Pactole.Util.Bijection. Require Import Pactole.Setting. Require Import FMapFacts. Require Import Pactole.Spaces.R. @@ -95,7 +95,7 @@ assert (Heven := even_nG). assert (H0 := nG_non_0). simpl. destruct n as [| [| ?]]. - lia. - destruct Heven. lia. -- simpl. lia. +- simpl. solve [ auto with arith | lia]. Qed. (* We need to unfold [obs_is_ok] for rewriting *) @@ -169,7 +169,7 @@ destruct (Nat.eq_dec (from_elements (List.map (fun x : location => (x, 1)) l))[e change (InA eq_pair (e, (from_elements (List.map (fun x0 : R => (x0, 1)) l))[e]) (elements (from_elements (List.map (fun x : location => (x, 1)) l)))). rewrite elements_spec; simpl. - split; trivial; []. apply neq_0_lt. auto. + split; trivial; []. apply Nat.neq_0_lt_0. auto. -- revert_all. rewrite removeA_InA; autoclass. tauto. Qed. @@ -209,7 +209,7 @@ Hint Resolve half_size_config : core. Definition lift_config {A} (config : G -> A) : ident -> A := fun id => match id with | Good g => config g - | Byz b => ltac:(exfalso; now apply (Nat.nlt_0_r (proj1_sig b)), proj2_sig) + | Byz b => ltac:(exfalso; now apply (Nat.nlt_0_r (fin2nat b)), fin_lt) end. Local Opaque G B. @@ -308,7 +308,7 @@ assert (Heven : Nat.Even nG). repeat split; trivial; [|]. + rewrite <- Nat.even_spec in Heven. assert (HnG := nG_non_0). simpl nG in *. - destruct n as [| [| ?]]; simpl; discriminate || lia || now elim HnG. + destruct n as [| [| ?]]; simpl; discriminate || lia || now apply HnG. + exists (sim origin), (sim one). repeat split. - intro Heq. apply Similarity.injective in Heq. symmetry in Heq. revert Heq. apply non_trivial. @@ -429,7 +429,7 @@ destruct (invalid_dec config) as [Hvalid | Hvalid]. { decompose [and or] Hperm; unfold origin in *; simpl in *; congruence. } simpl get_location. unfold id. rewrite <- Hpt1. f_equiv. apply Hmove. -* elim Hvalid. +* contradiction Hvalid. apply (invalid_reverse (build_similarity neq_0_1 Hdiff)). rewrite Hobs. unfold observation0. rewrite map_add, map_singleton, build_similarity_eq1, build_similarity_eq2; autoclass. @@ -577,7 +577,7 @@ destruct (invalid_dec config) as [Hvalid | Hvalid]. repeat destruct_match; try contradiction; [|]. - simpl in *. destruct Hperm as [[] | []]; subst; auto. - simpl in *. destruct Hperm as [[] | []]; subst; auto; congruence. -+ elim Hvalid. ++ contradiction Hvalid. apply (invalid_reverse (build_similarity neq_0_1 Hdiff)). rewrite Hobs. unfold observation0. rewrite map_add, map_singleton; autoclass; []. @@ -719,7 +719,7 @@ destruct (get_location (config (Good g)) =?= get_location (config (Good g0))) as - now rewrite build_similarity_eq1, Hsim0. - rewrite build_similarity_eq2. rewrite Hobs in Hpt'. unfold observation0 in Hpt'. rewrite map_add, map_singleton, add_In, In_singleton in Hpt'; autoclass; []. - decompose [and or] Hpt'; auto; []. elim Hdiff'. etransitivity; eauto. } + decompose [and or] Hpt'; auto; []. contradiction Hdiff'. etransitivity; eauto. } rewrite obs_from_config_ignore_snd, <- obs_from_config_map; autoclass; []. rewrite Hobs, map_merge; autoclass; []. rewrite <- (map_extensionality_compat Similarity.id), map_id; autoclass; [|]. @@ -735,11 +735,11 @@ destruct (get_location (config (Good g)) =?= get_location (config (Good g0))) as { assert (Hin := pos_in_config config origin (Good g)). rewrite Hobs in Hin. unfold observation0 in Hin. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. - destruct Hin as [[] | []]; trivial. elim Hcase. etransitivity; eauto. } + destruct Hin as [[] | []]; trivial. contradiction Hcase. etransitivity; eauto. } destruct (select_tower_case_2 (fun pt1 pt2 (_ : pt1 =/= pt2) => true) (fun pt1 pt2 (_ : pt1 =/= pt2) => false) true (Good g) Hvalid Hcase) as [Hdiff Hactivate]. rewrite Hactivate. - destruct_match; trivial; []. elim Hcase. etransitivity; eauto. + destruct_match; trivial; []. contradiction Hcase. etransitivity; eauto. Qed. Lemma invalid_da2_left_next : forall config, @@ -816,7 +816,7 @@ destruct (get_location (config (Good g)) =?= get_location (config (Good g0))) as { assert (Hin := pos_in_config config origin (Good g)). rewrite Hobs in Hin. unfold observation0 in Hin. rewrite map_add, map_singleton, add_In, In_singleton in Hin; autoclass; []. - destruct Hin as [[] | []]; trivial; []. elim Hcase. etransitivity; eauto. } + destruct Hin as [[] | []]; trivial; []. contradiction Hcase. etransitivity; eauto. } destruct (select_tower_case_2 (fun pt1 pt2 (_ : pt1 =/= pt2) => false) (fun pt1 pt2 (_ : pt1 =/= pt2) => true) true (Good g) Hvalid Hcase) as [Hdiff Hactivate]. destruct (select_tower_case_2 (fun pt1 pt2 (Hdiff0 : pt1 =/= pt2) => build_similarity Hdiff0 neq_0_1) @@ -831,7 +831,7 @@ destruct (get_location (config (Good g)) =?= get_location (config (Good g0))) as rewrite Hobs, map_merge; autoclass; []. rewrite <- (map_extensionality_compat Similarity.id), map_id; autoclass; [|]. + destruct_match. - - elim neq_0_1. apply (Similarity.injective sim). + - contradiction neq_0_1. apply (Similarity.injective sim). now transitivity (get_location (config (Good g))). - transitivity ((build_similarity (symmetry Hcase) neq_1_0)â»Â¹ move); try reflexivity; []. rewrite Hsim, build_similarity_inverse. @@ -870,7 +870,7 @@ assert (Hcase' : forall id, get_location (config id) = sim 0 \/ get_location (co { intro id. assert (Hin := pos_in_config config origin id). unfold observation0 in *. rewrite Hobs', map_add, map_singleton, add_In, In_singleton in Hin; autoclass; simpl in *; tauto. } assert (Hsim1 : get_location (config (Good g0)) == sim 1). -{ destruct (Hcase' (Good g0)); trivial; []. elim Hg3. now rewrite <- Hsim0. } +{ destruct (Hcase' (Good g0)); trivial; []. contradiction Hg3. now rewrite <- Hsim0. } clear pt1 pt2 g1 g2 Hg1 Hg2 Hdiff Hobs Hcase. assert (Hdiff_move : sim move =/= sim 1). { intro Heq. now apply Similarity.injective in Heq. } @@ -928,7 +928,7 @@ assert (Hcase' : forall id, get_location (config id) = sim 0 \/ get_location (co { intro id. assert (Hin := pos_in_config config origin id). unfold observation0 in *. rewrite Hobs', map_add, map_singleton, add_In, In_singleton in Hin; autoclass; simpl in *; tauto. } assert (Hsim1 : get_location (config (Good g0)) == sim 1). -{ destruct (Hcase' (Good g0)); trivial; []. elim Hg3. now rewrite <- Hsim0. } +{ destruct (Hcase' (Good g0)); trivial; []. contradiction Hg3. now rewrite <- Hsim0. } clear pt1 pt2 g1' g2' Hg1 Hg2 Hdiff Hobs Hcase. erewrite 2 round_simplify2_right; auto; []. rewrite 2 mk_info_get_location. diff --git a/CaseStudies/Gathering/InR2/Algorithm.v b/CaseStudies/Gathering/InR2/Algorithm.v index 6923ff9d78c69741961e7349b5c477dfe43f4bcb..c879256a4d315b9cfcce93194c8ea5a0b5856fe6 100644 --- a/CaseStudies/Gathering/InR2/Algorithm.v +++ b/CaseStudies/Gathering/InR2/Algorithm.v @@ -19,8 +19,7 @@ Require Import Bool. -Require Import Arith.Div2. -Require Import Lia Field. +Require Import Lia Field Lra. Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def. Require Import List. Require Import SetoidList. @@ -37,7 +36,7 @@ Require Export Pactole.Setting. (* Specific to R^2 topology *) Require Import Pactole.Spaces.R2. (* Specific to gathering *) -Require Pactole.CaseStudies.Gathering.WithMultiplicity. +Require Import Pactole.CaseStudies.Gathering.WithMultiplicity. Require Import Pactole.CaseStudies.Gathering.Definitions. (* Specific to multiplicity *) Require Import Pactole.Observations.MultisetObservation. @@ -54,6 +53,8 @@ Set Implicit Arguments. Close Scope R_scope. Close Scope VectorSpace_scope. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. (** * The Gathering Problem **) @@ -127,7 +128,7 @@ Ltac changeR2 := Lemma config_list_alls : forall pt, config_list (fun _ => pt) = alls pt nG. Proof using . intro. rewrite config_list_spec, map_cst. -rewrite names_length. simpl. now rewrite plus_0_r. +rewrite names_length. simpl. now rewrite Nat.add_0_r. Qed. Lemma no_byz_eq : forall config1 config2 : configuration, @@ -244,9 +245,9 @@ intros ? ? Heq. unfold is_clean. destruct (inclA_bool _ equiv_dec (support x) (SECT x)) eqn:Hx, (inclA_bool _ equiv_dec (support y) (SECT y)) eqn:Hy; trivial; rewrite ?inclA_bool_true_iff, ?inclA_bool_false_iff in *; [|]. -+ elim Hy. intros e Hin. rewrite <- Heq in Hin. ++ contradiction Hy. intros e Hin. rewrite <- Heq in Hin. apply SECT_compat in Heq. rewrite <- Heq. now apply Hx. -+ elim Hx. intros e Hin. rewrite Heq in Hin. ++ contradiction Hx. intros e Hin. rewrite Heq in Hin. apply SECT_compat in Heq. rewrite Heq. now apply Hy. Qed. @@ -360,10 +361,10 @@ intros config pt. split; intro Hmaj. simpl equiv. split; intro Hpt. - subst y. intro x. destruct (equiv_dec x pt). -- rewrite e. reflexivity. - -- apply lt_le_weak. now apply (Hmaj x). + -- apply Nat.lt_le_incl. now apply (Hmaj x). - destruct (equiv_dec y pt) as [? | Hy]; trivial. - exfalso. apply (Hmaj y) in Hy. elim (lt_irrefl (!! config)[pt]). - eapply le_lt_trans; try eassumption; []. + exfalso. apply (Hmaj y) in Hy. contradiction (Nat.lt_irrefl (!! config)[pt]). + eapply Nat.le_lt_trans; try eassumption; []. apply Hpt. * intros x Hx. apply max_spec_lub. - rewrite <- support_spec, Hmaj. now left. @@ -453,7 +454,7 @@ assert (Hsup : Permutation (support (!! config)) (pt1 :: pt2 :: nil)). assert (Hin2 : InA equiv pt2 (support (!! config))). { rewrite support_spec. unfold In. changeR2. setoid_rewrite Hpt2. now apply Exp_prop.div2_not_R0. } apply (PermutationA_split _) in Hin1. destruct Hin1 as [l Hl]. rewrite Hl in Hin2. - inversion_clear Hin2; try now subst; elim Hdiff. + inversion_clear Hin2; try now subst; contradiction Hdiff. rewrite size_spec, Hl in Hsuplen. destruct l as [| x [| ? ?]]; simpl in Hsuplen; try lia. inversion_clear H. - inversion H0; simpl in H1; subst. @@ -471,9 +472,8 @@ assert (Hpt : pt = pt1 \/ pt = pt2). setoid_rewrite Hmaj. now left. } inversion_clear Hin; auto. inversion_clear H0; auto. inversion H1. } -apply (lt_irrefl (Nat.div2 nG)). destruct Hpt; subst pt. +apply (Nat.lt_irrefl (Nat.div2 nG)). destruct Hpt; subst pt. - rewrite <- Hpt1 at 2. rewrite <- Hpt2. apply max_spec_lub; try now rewrite Hmax. - rewrite Hmax. auto. - rewrite <- Hpt1 at 1. rewrite <- Hpt2. apply max_spec_lub; now rewrite Hmax. Qed. @@ -514,7 +514,7 @@ intro config. unfold no_Majority. split. intro abs. subst. inversion hnodup; subst. - elim H1. + contradiction H1. constructor. reflexivity. * assert (h : inclA equiv (support (max (!! config))) (support (!! config))). @@ -537,7 +537,7 @@ intro config. unfold no_Majority. split. assert (hnodup := support_NoDupA (!! config)). rewrite Hsupp in hnodup. inversion hnodup; subst. - match goal with H : ~ InA equiv pt2 (pt2 :: nil) |- _ => elim H end. + match goal with H : ~ InA equiv pt2 (pt2 :: nil) |- _ => contradiction H end. constructor 1. reflexivity. } assert (heq_config: !! config == Madd pt1 ((!! config)[pt1]) (Madd pt2 ((!! config)[pt2]) empty)). @@ -645,7 +645,7 @@ Lemma not_invalid_no_majority_size : forall config, Proof using size_G. intros config H1 H2. assert (size (!! config) > 1)%nat. -{ unfold gt. eapply lt_le_trans; try eassumption; []. f_equiv. apply max_subset. } +{ unfold gt. eapply Nat.lt_le_trans; try eassumption; []. f_equiv. apply max_subset. } destruct (size (!! config)) as [| [| [| ?]]] eqn:Hlen; try lia. exfalso. apply H2. now rewrite invalid_equiv. Qed. @@ -684,7 +684,7 @@ Qed. Lemma SECT_cardinal_le_nG : forall config, SECT_cardinal (!! config) <= nG. Proof using . intro config. unfold SECT_cardinal. -replace nG with (nG + nB) by (simpl; apply plus_0_r). +replace nG with (nG + nB) by (simpl; apply Nat.add_0_r). rewrite <- (cardinal_obs_from_config config origin). apply cardinal_sub_compat, filter_subset. intros ? ? H. now rewrite H. @@ -824,10 +824,10 @@ intros sim s s_nonempty. unfold is_clean. changeR2. destruct (inclA_bool _ equiv_dec (support (map sim s)) (SECT (map sim s))) eqn:Hx, (inclA_bool _ equiv_dec (support s) (SECT s)) eqn:Hy; trivial; rewrite ?inclA_bool_true_iff, ?inclA_bool_false_iff, ?inclA_Leibniz in *; [|]. -- elim Hy. intros x Hin. apply (in_map sim) in Hin. rewrite <- map_sim_support in Hin. +- contradiction Hy. intros x Hin. apply (in_map sim) in Hin. rewrite <- map_sim_support in Hin. apply Hx in Hin. rewrite SECT_morph, in_map_iff in Hin;auto. destruct Hin as [x' [Heq ?]]. apply (Similarity.injective sim) in Heq. now rewrite <- Heq. -- elim Hx. intros x Hin. rewrite SECT_morph; auto. rewrite map_sim_support in Hin. +- contradiction Hx. intros x Hin. rewrite SECT_morph; auto. rewrite map_sim_support in Hin. rewrite in_map_iff in *. destruct Hin as [x' [? Hin]]. subst. exists x'. repeat split. now apply Hy. Qed. @@ -1109,7 +1109,7 @@ destruct (support (max (!! config))) as [| pt1 [| pt2 l]] eqn:Hmax, (support (max (!! (map_config sim config)))) as [| pt1' [| pt2' l']]; simpl in Hlen; discriminate || clear Hlen; [| |]. * (* No maximal tower *) - rewrite support_nil, max_is_empty in Hmax. elim (obs_non_nil _ Hmax). + rewrite support_nil, max_is_empty in Hmax. contradiction (obs_non_nil _ Hmax). * (* One maximal tower *) simpl in Hperm. rewrite <- PermutationA_Leibniz, (PermutationA_1 _) in Hperm. subst pt1'. apply Bijection.retraction_section. @@ -1202,7 +1202,7 @@ destruct (da.(activate) id) eqn:Hactive. destruct (is_clean (!! config)) eqn:Hclean. + reflexivity. + destruct (mem equiv_dec (get_location (config id)) (SECT (!! config))) eqn:Hmem. - - now elim Hmove. + - now contradiction Hmove. - reflexivity. * apply moving_active in Hmove; trivial; []. rewrite active_spec in Hmove. congruence. Qed. @@ -1217,14 +1217,17 @@ destruct (le_lt_dec 2 (length (support (max (!! config))))) as [Hle |Hlt]. now repeat rewrite destination_is_target. + rewrite moving_spec in Hmove1, Hmove2. rewrite (round_simplify _ id1) in Hmove1 |- *. rewrite (round_simplify _ id2) in Hmove2 |- *. - destruct (da.(activate) id1), (da.(activate) id2); try (now elim Hmove1 + elim Hmove2); []. + destruct (da.(activate) id1), (da.(activate) id2); try (now contradiction Hmove1 + contradiction Hmove2); []. cbn zeta in *. destruct (support (max (!! config))) as [| ? [| ? ?]] eqn:Hsupp. - - now elim Hmove1. + - now contradiction Hmove1. - reflexivity. - simpl in Hlt. lia. Qed. + + + (** Generic result of robograms using multiset obsra. *) Lemma increase_move : forall r config pt, @@ -1244,7 +1247,7 @@ destruct (existsb (fun x => if get_location (round r da config x) =?= pt then assert (Hg : forall id, get_location (round r da config id) <> pt \/ get_location (config id) = pt). { intro id. specialize (Hex id (In_names _)). revert Hex. repeat destruct_match; try discriminate; auto. } (** We prove a contradiction by showing that the opposite inequality of Hlt holds. *) - clear Hex. revert Hlt. apply le_not_lt. + clear Hex. revert Hlt. apply Nat.le_ngt. setoid_rewrite WithMultiplicity.obs_from_config_spec. do 2 rewrite config_list_spec. induction names as [| id l]; trivial; []. @@ -1278,14 +1281,14 @@ intros config pt. split. induction names as [| id' l]; try (now inversion Hin); []. inversion_clear Hin. + subst id'. clear IHl. simpl. R2dec_full. - - rewrite <- Hid in Heq. now elim Hroundid. + - rewrite <- Hid in Heq. now contradiction Hroundid. - R2dec_full; try contradiction; []. apply le_n_S. induction l as [| id' ?]; simpl. -- reflexivity. -- repeat R2dec_full; try now idtac + apply le_n_S + apply le_S; apply IHl. exfalso. now generalize (Hstay id' ltac:(assumption)). + apply IHl in H. simpl. repeat R2dec_full; try (simpl in *; lia); []. - elim Hneq. apply Hdest. rewrite moving_spec. intro Habs. rewrite Habs in Hneq. contradiction. + contradiction Hneq. apply Hdest. rewrite moving_spec. intro Habs. rewrite Habs in Hneq. contradiction. Qed. (** *** Generic results about the [SEC] after one round **) @@ -1311,7 +1314,7 @@ destruct (@increase_move gatherR2 config x) as [r_moving [Hdest_rmoving Hrmoving * simpl in *. lia. * destruct (le_lt_dec 2 (length (support (max (!! config))))) as [Hle | Hlt]. + rewrite destination_is_target in Hdest_rmoving. - - now elim Heq. + - now contradiction Heq. - unfold no_Majority. now rewrite size_spec. - rewrite moving_spec. intro Habs. apply Hrmoving. now rewrite Habs. + assert ((support (max (!! config))) = x :: nil). @@ -1383,7 +1386,7 @@ intros config id Hmaj Hclean Hcircle. rewrite (round_simplify_dirty Hmaj Hclean id). destruct (da.(activate) id); try reflexivity; []. destruct (mem equiv_dec (get_location (config id)) (SECT (!! config))) eqn:Hmem; try reflexivity; []. -rewrite mem_false_iff in Hmem. elim Hmem. +rewrite mem_false_iff in Hmem. contradiction Hmem. unfold SECT. right. unfold on_SEC. rewrite filter_InA; autoclass; []. split; trivial; []. rewrite support_spec. apply pos_in_config. @@ -1481,7 +1484,7 @@ Theorem MajTower_at_forever : forall pt config, MajTower_at pt config -> MajTower_at pt (round gatherR2 da config). Proof using Hssync. intros pt config Hmaj x Hx. assert (Hs := Hmaj x Hx). -apply le_lt_trans with ((!! config)[x]); try eapply lt_le_trans; try eassumption; [|]. +apply Nat.le_lt_trans with ((!! config)[x]); try eapply Nat.lt_le_trans; try eassumption; [|]. - eapply Majority_wither; eauto. - eapply Majority_grow; eauto. Qed. @@ -1538,8 +1541,8 @@ assert (Hext : forall x, f (!! (round gatherR2 da config)) x = f (!! config) x). { intro pt. unfold f. destruct (InA_dec equiv_dec pt (SECT (!! config))) as [Htest1 | Htest1], (InA_dec equiv_dec pt (SECT (!! (round gatherR2 da config)))) as [Htest2 | Htest2]; trivial. - - elim Htest2. now rewrite HsameSECT. - - elim Htest1. now rewrite <- HsameSECT. } + - contradiction Htest2. now rewrite HsameSECT. + - contradiction Htest1. now rewrite <- HsameSECT. } unfold f in Hext. rewrite (filter_extensionality_compat _ _ Hext). clear Hext. pose (f_target := fun x => if equiv_dec x (target (!! config)) then true else false). @@ -1586,7 +1589,7 @@ assert (Heq : @equiv observation observation_Setoid (filter f_out_target (!! (ro unfold f_target in Htest. revert Htest. destruct_match; try discriminate; auto. -- apply IHl. - - destruct_match; (now elim Hneq) || apply IHl. } + - destruct_match; (now contradiction Hneq) || apply IHl. } rewrite Heq. lia. Qed. @@ -1602,7 +1605,7 @@ Proof using size_G. intros config pt1 pt2 Hlen Hpt1 Hpt2 Hdiff12. rewrite <- support_spec in Hpt1, Hpt2. rewrite size_spec in Hlen. apply (PermutationA_split _) in Hpt1. destruct Hpt1 as [supp1 Hperm]. -rewrite Hperm in Hpt2. inversion_clear Hpt2; try (now elim Hdiff12); []. rename H into Hpt2. +rewrite Hperm in Hpt2. inversion_clear Hpt2; try (now contradiction Hdiff12); []. rename H into Hpt2. apply (PermutationA_split _) in Hpt2. destruct Hpt2 as [supp2 Hperm2]. rewrite Hperm2 in Hperm. rewrite Hperm in Hlen. destruct supp2 as [| pt3 supp]; try (now simpl in Hlen; lia); []. @@ -1668,7 +1671,7 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. { assert (Hperm : Permutation (support (!! (round gatherR2 da config))) (pt1 :: pt2 :: nil)). { symmetry. apply NoDup_Permutation_bis. + repeat constructor. - - intro Habs. inversion Habs. now elim Hdiff. now inversion H. + - intro Habs. inversion Habs. now contradiction Hdiff. now inversion H. - intro Habs. now inversion Habs. (* NoDupA_Leibniz had a useless hyp in coq stdlib until april 2020. *) (* + rewrite <- NoDupA_Leibniz. change eq with (@equiv location _). apply support_NoDupA. *) @@ -1690,7 +1693,7 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. assert (Hdest : forall g, List.In g (moving gatherR2 da config) -> get_location (round gatherR2 da config g) == pt). { intros id Hid. rewrite <- Hrmove_pt. apply same_destination; auto. rewrite moving_spec. congruence. } - assert ((div2 nG) <= (!! config)[pt']). + assert ((Nat.div2 nG) <= (!! config)[pt']). { transitivity ((!! (round gatherR2 da config))[pt']). - decompose [and or] Hpt; clear Hpt; subst. + setoid_rewrite Hpt2. simpl. reflexivity. @@ -1702,19 +1705,21 @@ destruct (support (max (!! config))) as [| pt [| pt' l']] eqn:Hmaj. rewrite <- Hid1. symmetry. apply Hdest. rewrite moving_spec. intro Habs. apply Hid2. now rewrite Habs. } - assert (Hlt : forall p, p <> pt' -> (!! config)[p] < div2 nG). + assert (Hlt : forall p, p <> pt' -> (!! config)[p] < Nat.div2 nG). { assert (Hpt'_in : In pt' (!! config)). - { unfold In. eapply lt_le_trans; try eassumption. apply Exp_prop.div2_not_R0. apply HsizeG. } + { unfold In. eapply Nat.lt_le_trans; try eassumption. apply Exp_prop.div2_not_R0. apply HsizeG. } assert (Hle := not_invalid_no_majority_size Hmaj Hok). - intros p Hp. apply Nat.nle_gt. intro Habs. apply (lt_irrefl nG). + intros p Hp. apply Nat.nle_gt. intro Habs. apply (Nat.lt_irrefl nG). destruct (@towers_elements_3 config pt' p Hle Hpt'_in) as [pt3' [Hdiff13 [Hdiff23 Hpt3_in]]]; trivial. - + apply lt_le_trans with (div2 nG); try assumption. apply Exp_prop.div2_not_R0. apply HsizeG. + + apply Nat.lt_le_trans with (Nat.div2 nG); try assumption. apply Exp_prop.div2_not_R0. apply HsizeG. + auto. - + eapply lt_le_trans; try apply (sum3_le_total config Hp Hdiff13 Hdiff23); []. - unfold In in Hpt3_in. rewrite <- ?Even.even_equiv in *. - rewrite (even_double nG); auto. unfold Nat.double. lia. } + + eapply Nat.lt_le_trans; try apply (sum3_le_total config Hp Hdiff13 Hdiff23); []. + unfold In in Hpt3_in. rewrite <- ?Nat.Even_alt_Even in *. + rewrite (Nat.Even_double nG). + 2:{ now apply Nat.Even_alt_Even. } + unfold Nat.double. lia. } assert (Hmaj' : MajTower_at pt' config). - { intros x Hx. apply lt_le_trans with (div2 nG); trivial. now apply Hlt. } + { intros x Hx. apply Nat.lt_le_trans with (Nat.div2 nG); trivial. now apply Hlt. } apply MajTower_at_equiv in Hmaj'. red in Hmaj. rewrite size_spec in Hmaj. @@ -1844,7 +1849,7 @@ intros config ptx pty Hinvalid Hmaj Hclean Hsec. assert (Hperm := diameter_clean_support Hinvalid Hmaj Hclean Hsec). destruct (support (max (!! (round gatherR2 da config)))) as [| pt [| ? ?]] eqn:Hmax'. - rewrite support_nil, max_is_empty, <- support_nil in Hmax'. - now elim (support_non_nil _ Hmax'). + now contradiction (support_non_nil _ Hmax'). - left. exists pt. rewrite MajTower_at_equiv. now rewrite Hmax'. - right. @@ -1988,7 +1993,7 @@ Qed. Ltac inv_notin H := match type of H with | ~ List.In ?x nil => clear H - | ~ InA (@equiv R2 _) ?x ?l => + | ~ InA (@equiv _ _) ?x ?l => let h := fresh H in assert (h:~ List.In x l); [ rewrite <- InA_Leibniz; assumption | inv_notin h ] @@ -2002,13 +2007,15 @@ Ltac inv_notin H := Ltac inv_nodup H := lazymatch type of H with - | NoDupA (@equiv R2 _) nil => clear H - | NoDupA (@equiv R2 _) (?x::nil) => clear H - | NoDupA (@equiv R2 _) (?x::?y::?l) => + | NoDupA (@equiv _ _) nil => clear H + | NoDupA (@equiv _ _) (?x::nil) => clear H + | NoDupA (@equiv _ _) (?x::?y::?l) => let x := fresh "x" in let l := fresh "l" in let C := fresh "h_notin" in let D := fresh "h_nodup" in + let E := fresh "E" in + let F := fresh "F" in inversion H as [|x l C D [E F]]; match type of E with | ?x = _ => subst x @@ -2019,7 +2026,7 @@ Ltac inv_nodup H := inv_notin C; inv_nodup D (* try clear C; try clear D *) - | NoDupA (@equiv R2 _) (?x :: ?l) => idtac (* nothing to do here *) + | NoDupA (@equiv _ _) (?x :: ?l) => idtac (* nothing to do here *) end. (** **** Merging results about the different kinds of triangles **) @@ -2039,10 +2046,10 @@ Lemma triangle_next_maj_or_diameter_or_triangle : forall config, \/ no_Majority (round gatherR2 da config) /\ PermutationA equiv (on_SEC (support (!! (round gatherR2 da config)))) (on_SEC (support (!! config))). -Proof. +Proof using Hssync da n size_G. intros config Hinvalid [Hmaj [ptx [pty [ptz Hsec]]]]. destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eqn:Hmax'. -- rewrite support_nil, max_is_empty in Hmax'. elim (obs_non_nil _ Hmax'). +- rewrite support_nil, max_is_empty in Hmax'. contradiction (obs_non_nil _ Hmax'). - now left. - right. get_case (round gatherR2 da config). rename Hmaj0 into Hmaj'. @@ -2063,7 +2070,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq assert (Hcenter := on_SEC_pair_is_diameter _ Hsec'). assert (Hperm : PermutationA equiv (support (!! (round gatherR2 da config))) (middle pt1 pt2 :: pt1 :: pt2 :: nil)). - apply diameter_clean_support;auto. + { apply diameter_clean_support;auto. } destruct (is_clean (!! config)) eqn:Hclean. ** assert (Hincl : inclA equiv (support (!! (round gatherR2 da config))) (target (!! config) :: ptx :: pty :: ptz :: nil)). @@ -2132,119 +2139,118 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq { rewrite Hsec'. now right; left. } rewrite InA_Leibniz in Hin |- *. now apply on_SEC_In. } -(* FIXME: Auto-names changed... - inversion H1; trivial; []. + inversion Hin; trivial; []. exfalso. - rewrite <- H2 in Htarget. - rewrite Htarget in H3. + rewrite <- H0 in Htarget. + rewrite Htarget in H. subst pt2; inv hNoDup; intuition. } - unfold inclA in H0. + unfold inclA in Hincl. assert (hmid:InA equiv (middle pt1 pt2) (middle pt1 pt2 :: pt1 :: pt2 :: nil)). { left. reflexivity. } - specialize (H0 (middle pt1 pt2) hmid). - rewrite InA_Leibniz in H0. - simpl in H0. - decompose [or False] H0;clear H0. - + rewrite Htarget in H1. - rewrite <- H2 in H1. - elim (@middle_diff pt1 pt2). + specialize (Hincl (middle pt1 pt2) hmid). + clear hmid. + rewrite InA_Leibniz in Hincl. + lazy beta iota delta [List.In] in Hincl. + decompose [or False] Hincl;clear Hincl. + + + rewrite Htarget in H. + rewrite <- H0 in H. + apply (@middle_diff pt1 pt2). * intro abs. rewrite abs in hNoDup. inversion hNoDup. - apply H4. - left; reflexivity. - * rewrite <- H1. - left;reflexivity. - + assert(ptx = pt2). - { rewrite middle_comm in H3. + apply H3. + left; reflexivity. + * changeR2. + rewrite <- H. + simpl. auto. + + assert(ptx == pt2). + { rewrite middle_comm in H1. eapply equilateral_isobarycenter_degenerated_gen with (ptopp:=pt2) (mid:=ptx) (white:=pt1);eauto. left. reflexivity. } subst ptx. - rewrite middle_comm in H0. - rewrite middle_eq in H0. - rewrite H0 in hNoDup. + rewrite middle_comm in H. + apply (middle_eq pt2 pt1) in H. + rewrite H in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. + assert(pty = pt2). - { rewrite middle_comm in H1. + { rewrite middle_comm in H. eapply equilateral_isobarycenter_degenerated_gen with (ptopp:=pt2) (mid:=pty) (white:=pt1);eauto. right;left. reflexivity. } subst pty. - rewrite middle_comm in H0. - rewrite middle_eq in H0. - rewrite H0 in hNoDup. + rewrite middle_comm in H1. + rewrite (middle_eq pt2 pt1) in H1. + rewrite H1 in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. + assert(ptz = pt2). - { rewrite middle_comm in H3. + { rewrite middle_comm in H1. eapply equilateral_isobarycenter_degenerated_gen with (ptopp:=pt2) (mid:=ptz) (white:=pt1);eauto. right;right;left. reflexivity. } subst ptz. - rewrite middle_comm in H0. - rewrite middle_eq in H0. - rewrite H0 in hNoDup. + rewrite middle_comm in H. + rewrite (middle_eq pt2 pt1) in H. + rewrite H in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. - (* if (target (config)) is in (SEC (round config)) then two previously SEC-towers have moved to (target (config)). therefore there are two towers => majority (or contradicting invalid). *) - assert (hIn:In pt1 (ptx :: pty :: ptz :: nil)). - { assert (In pt1 (target (!! config) :: ptx :: pty :: ptz :: nil)). + assert (hIn:List.In pt1 (ptx :: pty :: ptz :: nil)). + { assert (Hin:List.In pt1 (target (!! config) :: ptx :: pty :: ptz :: nil)). { rewrite <- Hsec. apply InA_Leibniz. - eapply incl_clean_next with da;auto. - assert (InA equiv pt1 (on_SEC (support (!! (round gatherR2 da config))))). + eapply incl_clean_next ;auto;[]. + assert (Hin:InA equiv pt1 (on_SEC (support (!! (round gatherR2 da config))))). { rewrite Hsec'. left;reflexivity. } - rewrite InA_Leibniz in H2 |-*. + rewrite InA_Leibniz in Hin |-*. apply on_SEC_In. assumption. } - inversion H2. - - exfalso. - rewrite H3 in Htarget. - rewrite <- Htarget in H1. - subst pt2. - inversion hNoDup. - apply H5. - left;reflexivity. - - assumption. } - unfold inclA in H0. + inversion Hin;trivial;[]. + exfalso. + rewrite <- H in Htarget. + rewrite Htarget in H0. + subst pt2; inv hNoDup; intuition. } + lazy beta iota delta [inclA] in Hincl. assert (hmid:InA equiv (middle pt1 pt2) (middle pt1 pt2 :: pt1 :: pt2 :: nil)). { left. reflexivity. } - specialize (H0 (middle pt1 pt2) hmid). - rewrite InA_Leibniz in H0. - simpl in H0. - decompose [or False] H0;clear H0. - + rewrite Htarget in H2. - rewrite <- H1 in H2. - elim (@middle_diff pt1 pt2). + specialize (Hincl (middle pt1 pt2) hmid). + rewrite InA_Leibniz in Hincl. + lazy beta iota delta [List.In] in Hincl. + decompose [or False] Hincl;clear Hincl. + + rewrite Htarget in H0. + rewrite <- H in H0. + apply (@middle_diff pt1 pt2). * intro abs. rewrite abs in hNoDup. inversion hNoDup. - apply H4. - left; reflexivity. - * rewrite <- H2. - right;left;reflexivity. + apply H3. + left; reflexivity. + * changeR2. + rewrite <- H0. + simpl. auto. + assert(ptx = pt1). { eapply equilateral_isobarycenter_degenerated_gen with (ptopp:=pt1) (mid:=ptx) (white:=pt2);eauto. left. reflexivity. } subst ptx. - rewrite middle_eq in H0. + rewrite (middle_eq pt1 pt2) in H0. rewrite H0 in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. + assert(pty = pt1). @@ -2253,10 +2259,10 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq right;left. reflexivity. } subst pty. - rewrite middle_eq in H0. - rewrite H0 in hNoDup. + rewrite (middle_eq pt1 pt2) in H1. + rewrite H1 in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. + assert(ptz = pt1). @@ -2265,48 +2271,48 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq right;right;left. reflexivity. } subst ptz. - rewrite middle_eq in H0. + rewrite (middle_eq pt1 pt2) in H0. rewrite H0 in hNoDup. inversion hNoDup. - apply H4. + apply H3. left. reflexivity. } --- (* (ptx :: pty :: ptz :: nil) = (middle pt1 pt2 :: pt1 :: pt2 :: nil) contradiction with calssify_triangle = equilateral *) assert (PermutationA equiv (ptx :: pty :: ptz :: nil) (middle pt1 pt2 :: pt1 :: pt2 :: nil)). - { apply inclA_skip in H0;autoclass. + { apply inclA_skip in Hincl;autoclass. - symmetry. - apply NoDupA_inclA_length_PermutationA with (1:=equiv_equiv);auto. - + rewrite <- H. + Set Printing Depth 100. + apply NoDupA_inclA_length_PermutationA with (1:=setoid_equiv);auto. + + rewrite <- Hperm. apply support_NoDupA;auto. + rewrite <- Hsec. apply on_SEC_NoDupA;auto. - apply support_NoDupA;auto. - - rewrite InA_Leibniz. - assumption. } + apply support_NoDupA;auto. } assert (classify_triangle (middle pt1 pt2) pt1 pt2 = Equilateral). - { rewrite PermutationA_Leibniz in H1. now rewrite (classify_triangle_compat H1) in Htriangle. } - functional inversion H2. clear H2. + { rewrite PermutationA_Leibniz in H. now rewrite (classify_triangle_compat H) in Htriangle. } + functional inversion H0. (*clear H0.*) rewrite -> ?Rdec_bool_true_iff in *. - rewrite R2.dist_sym in H3. - rewrite R2dist_middle in H3. - assert (R2.dist pt1 pt2 = 0%R). - { lra. } - apply R2.dist_defined in H2. + rewrite dist_sym in H1. + rewrite R2dist_middle in H1. + assert (dist pt1 pt2 = 0%R). + { changeR2. + lra. } + apply dist_defined in H3. assert (hNoDup:NoDupA equiv (pt1 :: pt2 :: nil)). { rewrite <- Hsec'. apply on_SEC_NoDupA. apply support_NoDupA. } - rewrite H2 in hNoDup. + rewrite H3 in hNoDup. inversion hNoDup. - apply H7. left;reflexivity. + apply H6. left;reflexivity. ** rewrite <- dirty_next_on_SEC_same in Hsec;auto. rewrite Hsec' in Hsec. assert (length (pt1 :: pt2 :: nil) = length (ptx :: pty :: ptz :: nil)). { rewrite Hsec. reflexivity. } - simpl in H0;lia. + simpl in H;lia. -- (* Valid case: the center of the SEC is not on a diameter *) left. repeat split; trivial; eauto. @@ -2315,8 +2321,8 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq - assumption. - exfalso. apply not_true_is_false in heq_clean_false. - assert (hdirty:=@dirty_next_SEC_same da config Hmaj heq_clean_false). - setoid_rewrite <- (@dirty_next_on_SEC_same da config Hmaj heq_clean_false) in Hsec. + assert (hdirty:=@dirty_next_SEC_same config Hmaj heq_clean_false). + setoid_rewrite <- (@dirty_next_on_SEC_same config Hmaj heq_clean_false) in Hsec. rewrite Hsec' in Hsec. apply PermutationA_length in Hsec. simpl in Hsec. @@ -2351,6 +2357,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq apply support_NoDupA. } inv_nodup hnodupxyz. inv_nodup hnodup. + (* simpl in H;lia. *) destruct (equiv_dec pt1 (isobarycenter_3_pts ptx pty ptz)) as [heq_pt1_bary | hneq_pt1_bary]. ++ { exfalso. assert(hpermut_config: PermutationA equiv (support (!! (round gatherR2 da config))) @@ -2364,31 +2371,29 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq clear h_pt2. inversion h_incl_pt1_pt2 as [pt lpt heq_pt2_ptx [__h heq_lpt]| pt lpt h_in_pt2_lpt [__h heq_lpt]]. (* pt2 = ptx *) - * unfold equiv, R2def.eq in heq_pt2_ptx. + * unfold equiv, R2_Setoid in heq_pt2_ptx. subst. assert (hpermut:PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil) (pty :: ptz :: isobarycenter_3_pts ptx pty ptz :: ptx :: nil)) by permut_3_4. rewrite hpermut in hincl_round;clear hpermut. assert (h_ynotin:~ InA equiv pty (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty) (ptz:=ptz) ;eauto. - rewrite Hsec'. reflexivity. } assert (h_znotin:~ InA equiv ptz (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty) (ptz:=ptz);eauto. - rewrite Hsec'. reflexivity. } do 2 (apply inclA_skip in hincl_round; autoclass). apply NoDupA_inclA_length_PermutationA; autoclass. -- apply support_NoDupA. -- now rewrite heq_pt1_bary. - -- simpl. - transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). + -- transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). ++ now rewrite Hsec'. - ++ unfold on_SEC. + ++ lazy beta iota delta [on_SEC]. rewrite filter_length. + changeR2. lia. * { (* InA equiv pt2 (pt2 :: nil) *) @@ -2397,20 +2402,18 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq inversion h_in_pt2_lpt as [pt lpt heq_pt2_pty [__h heq_lpt] | pt lpt h_in_pt2_lpt' [__h heq_lpt]]. (* pt2 = pty *) - * unfold equiv, R2def.eq in heq_pt2_pty. + * unfold equiv, R2_Setoid in heq_pt2_pty. subst. assert (Hperm:PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil) (ptx :: ptz :: isobarycenter_3_pts ptx pty ptz :: pty :: nil)) by permut_3_4. rewrite Hperm in hincl_round;clear Hperm. assert (h_ynotin:~ InA equiv ptx (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. reflexivity. } assert (h_znotin:~ InA equiv ptz (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. reflexivity. } apply inclA_skip in hincl_round;autoclass. @@ -2419,19 +2422,21 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq -- apply support_NoDupA. -- rewrite heq_pt1_bary. assumption. - -- simpl. + -- changeR2. transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). - ++ rewrite Hsec'. + ++ changeR2. + rewrite Hsec'. reflexivity. ++ unfold on_SEC. rewrite filter_length. + changeR2. lia. * subst pt. subst lpt. { inversion h_in_pt2_lpt' as [pt lpt heq_pt2_pty [__h heq_lpt] | pt lpt h_in_pt2_lpt'' [__h heq_lpt]]. (* pt2 = pty *) - * unfold equiv, R2def.eq in heq_pt2_pty. + * unfold equiv, R2_Setoid in heq_pt2_pty. subst. assert (Hperm : PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil) @@ -2439,26 +2444,25 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq by permut_3_4. rewrite Hperm in hincl_round;clear Hperm. assert (h_ynotin:~ InA equiv ptx (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. - - rewrite Hsec'. - reflexivity. } + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. + - rewrite Hsec'. + reflexivity. } assert (h_znotin:~ InA equiv pty (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. - - rewrite Hsec'. - reflexivity. } + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. + - rewrite Hsec'. + reflexivity. } apply inclA_skip in hincl_round;autoclass. apply inclA_skip in hincl_round;autoclass. apply NoDupA_inclA_length_PermutationA;autoclass. -- apply support_NoDupA. -- now rewrite heq_pt1_bary. - -- simpl. + -- transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). ++ rewrite Hsec'. reflexivity. ++ unfold on_SEC. rewrite filter_length. + changeR2. lia. * inversion h_in_pt2_lpt''. } } + intro Hin. apply heq2. now inversion Hin. } @@ -2482,7 +2486,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq inversion h_incl_pt1_pt2 as [pt lpt heq_pt1_ptx [__h heq_lpt] | pt lpt h_in_pt1_lpt [__h heq_lpt]]. (* pt1 = ptx *) - * unfold equiv, R2def.eq in heq_pt1_ptx. + * unfold equiv, R2_Setoid in heq_pt1_ptx. subst ptx. subst pt. assert (Hperm:PermutationA equiv (isobarycenter_3_pts pt1 pty ptz :: pt1 :: pty :: ptz :: nil) @@ -2490,13 +2494,11 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq by permut_3_4. rewrite Hperm in hincl_round;clear Hperm. assert (h_ynotin:~ InA equiv pty (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=pt1)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. permut_3_4. } assert (h_znotin:~ InA equiv ptz (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=pt1)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. permut_3_4. } apply inclA_skip in hincl_round;autoclass. @@ -2510,11 +2512,11 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq ** now rewrite InA_nil in *. ++ now rewrite InA_nil. -- now rewrite heq_pt2_bary. - -- simpl. - transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). + -- transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). ++ now rewrite Hsec'. ++ unfold on_SEC. rewrite filter_length. + changeR2. lia. * { (* InA equiv pt1 (pt1 :: nil) *) @@ -2523,7 +2525,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq inversion h_in_pt1_lpt as [pt lpt heq_pt1_pty [__h heq_lpt] | pt lpt h_in_pt1_lpt' [__h heq_lpt]]. (* pt1 = pty *) - * unfold equiv, R2def.eq in heq_pt1_pty. + * unfold equiv, R2_Setoid in heq_pt1_pty. subst. assert (Hperm : PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil) @@ -2531,13 +2533,11 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq by permut_3_4. rewrite Hperm in hincl_round;clear Hperm. assert (h_xnotin:~ InA equiv ptx (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. permut_3_4. } assert (h_znotin:~ InA equiv ptz (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. permut_3_4. } apply inclA_skip in hincl_round;autoclass. @@ -2552,19 +2552,19 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq ++ now rewrite InA_nil. -- rewrite heq_pt2_bary. assumption. - -- simpl. - transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). + -- transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). ++ rewrite Hsec'. reflexivity. ++ unfold on_SEC. rewrite filter_length. + changeR2. lia. * subst pt. subst lpt. { inversion h_in_pt1_lpt' as [pt lpt heq_pt1_ptz [__h heq_lpt] | pt lpt h_in_pt1_lpt'' [__h heq_lpt]]. (* pt1 = pty *) - * unfold equiv, R2def.eq in heq_pt1_ptz. + * unfold equiv, R2_Setoid in heq_pt1_ptz. subst. assert (hpermut : PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil) @@ -2572,13 +2572,11 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq by permut_3_4. rewrite hpermut in hincl_round;clear hpermut. assert (h_xnotin:~ InA equiv ptx (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2;eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz);eauto. - rewrite Hsec'. permut_3_4. } assert (h_ynotin:~ InA equiv pty (support (!! (round gatherR2 da config)))). - { eapply SEC_3_to_2; eauto. - - repeat econstructor; reflexivity. + { eapply SEC_3_to_2 with (ptx:=ptx)(pty:=pty)(ptz:=ptz); eauto. - rewrite Hsec'. permut_3_4. } @@ -2592,12 +2590,12 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq ** now rewrite InA_nil in *. ++ now rewrite InA_nil. -- now rewrite heq_pt2_bary. - -- simpl. - transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). + -- transitivity (length (on_SEC (support (!! (round gatherR2 da config))))). ++ rewrite Hsec'. reflexivity. ++ unfold on_SEC. rewrite filter_length. + changeR2. lia. * inversion h_in_pt1_lpt''. } } + intro abs. @@ -2626,7 +2624,6 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq contradiction. +++ inversion H2. ** assumption. } - * (* Valid case: SEC is a triangle *) right. split; trivial. rewrite <- Hsec'. @@ -2637,12 +2634,16 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq ++ assert (Hperm' : PermutationA equiv (support (!! (round gatherR2 da config))) (isobarycenter_3_pts ptx pty ptz :: ptx :: pty :: ptz :: nil)). { assert ((!! (round gatherR2 da config))[target (!! config)] > 0). - { apply le_lt_trans with ((!! config)[target (!! config)]); try lia. + { apply Nat.le_lt_trans with ((!! config)[target (!! config)]); try lia. rewrite increase_move_iff. exists gmove. split. - apply destination_is_target; trivial. rewrite Hmoving. now left. - - now rewrite <- moving_spec, Hmoving; left. } + - assert (hmoved_in:List.In gmove (moving gatherR2 da config)). + { rewrite Hmoving. + now left. } + apply moving_spec in hmoved_in. + apply hmoved_in. } apply (NoDupA_inclA_length_PermutationA _). - apply support_NoDupA. - apply equilateral_isobarycenter_NoDupA; trivial. @@ -2658,7 +2659,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq cut (4 + length l <= 3); try lia. change (4 + length l) with (length (pt1 :: pt2 :: pt3 :: pt4 :: l)). rewrite <- Hsec', <- Hlen, size_spec. - apply (NoDupA_inclA_length equiv_equiv). + apply (NoDupA_inclA_length setoid_equiv). - apply on_SEC_NoDupA, support_NoDupA. - unfold on_SEC. intro. rewrite (filter_InA _). intuition. } subst. @@ -2669,7 +2670,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq - rewrite <- Hsec'. apply on_SEC_NoDupA, support_NoDupA. - apply support_NoDupA. - rewrite <- Hsec'. unfold on_SEC. intro. rewrite (filter_InA _). intuition. - - rewrite <- size_spec. cbn. lia. } + - rewrite <- size_spec. rewrite Hlen. cbn. lia. } rewrite <- Hsec' in Hperm'. (* Triangle equilatéral: comme qqchose bouge et que on est encore avec 3 colonne après, une colonne s'est déplacée vers le barycentre, contradiction: @@ -2680,7 +2681,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq PermutationA equiv (pta :: ptb :: ptc :: nil) (ptx :: pty :: ptz :: nil) /\ PermutationA equiv (isobarycenter_3_pts ptx pty ptz :: pta :: ptb ::nil) (pt1 :: pt2 :: pt3 :: nil)). - { assert (hincl:=incl_clean_next da config Hclean). + { assert (hincl:=incl_clean_next config Hclean). rewrite Hsec in hincl. rewrite Hperm', Hsec' in hincl. assert (hbary : InA equiv (isobarycenter_3_pts ptx pty ptz) @@ -2722,26 +2723,44 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq end end; subst; permut_3_4. } destruct Hex as [pta [ptb [ptc [Hpermxyz Hperm]]]]. - pose (better_SEC := {| center := middle pta ptb; radius := /2 * R2.dist pta ptb |}). - assert (Hbary_strict : (R2.dist (isobarycenter_3_pts ptx pty ptz) (center better_SEC) + pose (better_SEC := {| R2.center := middle pta ptb; radius := /2 * dist pta ptb |}). + + assert (Hbary_strict : (dist (isobarycenter_3_pts ptx pty ptz) (R2.center better_SEC) < radius better_SEC)%R). { rewrite PermutationA_Leibniz in Hpermxyz. rewrite <- (isobarycenter_3_pts_compat Hpermxyz). - unfold better_SEC. simpl. rewrite R2norm_dist. unfold isobarycenter_3_pts, middle. - replace (/ 3 * (pta + (ptb + ptc)) - 1 / 2 * (pta + ptb))%R2 - with (/6 * (ptc + ptc - (pta + ptb)))%R2 by (destruct pta, ptb, ptc; simpl; f_equal; field). - rewrite R2norm_mul. rewrite Rabs_pos_eq; try lra; []. - rewrite <- R2norm_dist. - cut (R2.dist (ptc + ptc) (pta + ptb) < 3 * R2.dist pta ptb)%R; try lra; []. + unfold better_SEC. simpl. repeat rewrite R2.norm_dist. + pose (h:=@Barycenter_spec pta ptb ptc). + Transparent isobarycenter_3_pts middle. + unfold isobarycenter_3_pts, middle. + Opaque isobarycenter_3_pts middle. + replace (/ 3 * (pta + (ptb + ptc)) - 1 / 2 * (pta + ptb))%VS + with (/6 * (ptc + ptc - (pta + ptb)))%VS by (destruct pta, ptb, ptc; simpl; f_equal; field). + rewrite norm_mul. rewrite Rabs_pos_eq; try lra; []. + repeat rewrite <- norm_dist. + cut (dist (ptc + ptc) (pta + ptb) < 3 * dist pta ptb)%R. + { changeR2. lra. } + + eapply Rle_lt_trans. - - apply (R2riang_ineq _ (ptc + pta)%R2). - - setoid_rewrite R2.add_comm at 2 4. do 2 rewrite R2add_dist. + - apply (triang_ineq _ (ptc + pta)%VS). + - changeR2. + setoid_rewrite RealVectorSpace.add_comm at 2 4. + do 2 rewrite dist_translation. rewrite <- (classify_triangle_compat Hpermxyz) in Htriangle. rewrite classify_triangle_Equilateral_spec in Htriangle. destruct Htriangle as [Heq1 Heq2]. - setoid_rewrite R2.dist_sym at 1 2. rewrite Heq1, Heq2. - assert (Hle' := R2.dist_pos ptb ptc). - rewrite <- PermutationA_Leibniz in Hpermxyz. rewrite <- Hpermxyz in Hnodup. inv_nodup Hnodup. - rewrite <- R2.dist_defined in heq1. lra. } + setoid_rewrite dist_sym at 1 2. changeR2. rewrite Heq1, Heq2. + assert (Hle' := dist_nonneg ptb ptc). + rewrite <- PermutationA_Leibniz in Hpermxyz. rewrite <- Hpermxyz in Hnodup. + inv_nodup Hnodup. + assert (heq1': (dist ptb ptc <> 0)%R). + { red. + intro abs. + rewrite dist_defined in abs. + vm_compute in abs. + contradiction. } + changeR2. + lra. } assert (enclosing_circle better_SEC (isobarycenter_3_pts ptx pty ptz :: pta :: ptb :: nil)). { intros pt hin. simpl in hin. @@ -2754,7 +2773,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq - unfold better_SEC ; simpl. rewrite middle_comm. rewrite R2dist_middle. - rewrite R2.dist_sym. + rewrite dist_sym. reflexivity. } assert (better_SEC = (SEC (support (!! (round gatherR2 da config))))). { rewrite PermutationA_Leibniz in Hperm',Hperm. @@ -2795,7 +2814,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq ++ apply no_moving_same_config in Hmoving. now rewrite Hmoving. ++ assert (Hperm' : PermutationA equiv (support (!! (round gatherR2 da config))) (ptx :: pty :: ptz :: nil)). - { assert (forall x, In x (gmove :: l) -> get_location (round gatherR2 da config x) == vertex). + { assert (forall x, List.In x (gmove :: l) -> get_location (round gatherR2 da config x) == vertex). { rewrite <- Htarget. intros x H3. apply destination_is_target; auto. @@ -2817,6 +2836,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq reflexivity. -- unfold on_SEC. rewrite filter_length. + changeR2. lia. - rewrite Htarget. assumption. } @@ -2841,6 +2861,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq rewrite <- Hsec. apply on_SEC_idempotent. -- now apply dirty_next_on_SEC_same. + + (* Scalene case *) assert (Htarget := scalene_target config Hsec Htriangle). right. split; trivial. @@ -2852,7 +2873,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq remember (opposite_of_max_side ptx pty ptz) as vertex. assert (Hperm' : PermutationA equiv (support (!! (round gatherR2 da config))) (ptx :: pty :: ptz :: nil)). - { assert (forall x, In x (gmove :: l) -> get_location (round gatherR2 da config x) == vertex). + { assert (forall x, List.In x (gmove :: l) -> get_location (round gatherR2 da config x) == vertex). { rewrite <- Htarget. intros x H3. apply destination_is_target;auto. @@ -2874,7 +2895,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq reflexivity. -- unfold on_SEC. rewrite filter_length. - lia. + changeR2. lia. - subst. rewrite Htarget. assumption. } @@ -2900,8 +2921,7 @@ destruct (support (max (!! (round gatherR2 da config)))) as [| pt1 [| pt2 l]] eq rewrite <- Hsec. apply on_SEC_idempotent. -- now apply dirty_next_on_SEC_same. -Qed. *) -Admitted. +Qed. (** *** Lemmas about the generic case **) @@ -2960,7 +2980,7 @@ assert (Hex : forall id id', { intros id id' Hid Hid' Hneq Hactive. simpl in *. destruct (da.(activate) id') eqn:Hactive'; trivial; exfalso. decompose [or] Hid; decompose [or] Hid'; try subst id; try subst id'; - (now elim Hneq) || rewrite Hactive in *; changeR2; rewrite Hactive' in *; + (now contradiction Hneq) || rewrite Hactive in *; changeR2; rewrite Hactive' in *; rewrite ?Hid1, ?Hid2, ?Hid3, ?Hid4 in *; R2dec. } (* Therefore, at least three were not activated and not on the target *) assert (Hperm_id : exists id1' id2' id3' id4', @@ -3090,7 +3110,7 @@ Proof using Hssync. unfold measure at 2. destruct (support (max (!! config))) as [| pt [| pt' smax]] eqn:Hmax. - (* No robots *) - rewrite support_nil, max_is_empty in Hmax. elim (obs_non_nil _ Hmax). + rewrite support_nil, max_is_empty in Hmax. contradiction (obs_non_nil _ Hmax). - (* A majority tower *) get_case config. apply MajTower_at_forever in Hcase. @@ -3109,7 +3129,7 @@ Proof using Hssync. simpl. get_case config. rewrite (round_simplify_Majority Hcase0 gmove). - destruct (da.(activate) gmove); try reflexivity; []. now elim Hactive. + destruct (da.(activate) gmove); try reflexivity; []. contradiction diff_false_true. - (* Computing the SEC *) get_case config. clear Hmax pt pt' smax. destruct (is_clean (!! config)) eqn:Hclean. @@ -3201,7 +3221,7 @@ Proof using Hssync. unfold measure. destruct (support (max (!! (round gatherR2 da config)))) as [| ? [| ? ?]] eqn:Hmax'. * (* Absurd: no robot after one round *) - rewrite support_nil, max_is_empty in Hmax'. elim (obs_non_nil _ Hmax'). + rewrite support_nil, max_is_empty in Hmax'. contradiction (obs_non_nil _ Hmax'). * (* A majority tower after one round *) destruct (on_SEC (support (!! config))) as [| ? [| ? [| ? [| ? ?]]]]; cbn in Hle; lia || left; lia. @@ -3218,6 +3238,32 @@ Qed. End SSYNC_Results. +(* destination is independent from the demonic_action. This replace + the same_destination inside previous section. *) +Corollary same_destination_strong : forall da da' (config : configuration) id1 id2, + SSYNC_da da -> SSYNC_da da' -> + List.In id1 (moving gatherR2 da config) -> + List.In id2 (moving gatherR2 da' config) -> + round gatherR2 da config id1 == round gatherR2 da' config id2. +Proof using Type. + intros da da' config id1 id2 hss hss' Hmove1 Hmove2. changeR2. + apply WithMultiplicity.no_info. + destruct (le_lt_dec 2 (length (support (max (!! config))))) as [Hle |Hlt]. + + assert (no_Majority config). { unfold no_Majority. now rewrite size_spec. } + now repeat rewrite destination_is_target. + + rewrite moving_spec in Hmove1, Hmove2. + rewrite (round_simplify da hss _ id1) in Hmove1 |- *. + rewrite (round_simplify da' hss' _ id2) in Hmove2 |- *. + destruct (da.(activate) id1), (da'.(activate) id2); + try (now contradiction Hmove1 + contradiction Hmove2); []. + cbn zeta in *. + destruct (support (max (!! config))) as [| ? [| ? ?]] eqn:Hsupp. + - now contradiction Hmove1. + - reflexivity. + - simpl in Hlt. lia. +Qed. + + (** *** With termination, the rest of the proof is easy **) Lemma gathered_precise : forall config pt, @@ -3237,7 +3283,7 @@ Lemma not_gathered_exists : forall config pt, Proof using . intros config pt Hgather. destruct (forallb (fun x => if get_location (config x) =?= pt then true else false) names) eqn:Hall. -- elim Hgather. rewrite forallb_forall in Hall. +- contradiction Hgather. rewrite forallb_forall in Hall. intro id'. setoid_rewrite R2dec_bool_true_iff in Hall. repeat rewrite Hall; reflexivity || apply In_names. - rewrite <- negb_true_iff, existsb_forallb, existsb_exists in Hall. destruct Hall as [id' [_ Hid']]. revert Hid'. destruct_match; discriminate || now exists id'. @@ -3249,7 +3295,7 @@ Theorem OneMustMove : forall config id, ~ WithMultiplicity.invalid config -> ~ga Proof using . intros config id Hvalid Hgather. destruct (support (max (!! config))) as [| pt [| pt' lmax]] eqn:Hmax. -* elim (support_max_non_nil _ Hmax). +* contradiction (support_max_non_nil _ Hmax). * rewrite <- MajTower_at_equiv in Hmax. apply not_gathered_generalize with _ _ pt in Hgather. apply not_gathered_exists in Hgather. destruct Hgather as [gmove Hmove]. @@ -3257,7 +3303,7 @@ destruct (support (max (!! config))) as [| pt [| pt' lmax]] eqn:Hmax. rewrite (round_simplify_Majority _ Hda Hmax gmove). destruct_match. + intro Habs. apply Hmove. now rewrite <- Habs. - + now elim Hactive. + + contradiction diff_false_true. * (* No majority tower *) get_case config. destruct (is_clean (!! config)) eqn:Hclean. @@ -3276,7 +3322,7 @@ destruct (support (max (!! config))) as [| pt [| pt' lmax]] eqn:Hmax. destruct Hin as [gmove Hmove]. exists gmove. intros da Hda Hactive. rewrite active_spec in Hactive. rewrite moving_spec. rewrite (round_simplify_dirty da Hda Hmaj Hclean gmove). - destruct (da.(activate) gmove); try (now elim Hactive); []. + destruct (da.(activate) gmove); [|contradiction diff_false_true]. destruct (mem equiv_dec (get_location (config gmove)) (SECT (!! config))) eqn:Htest. - rewrite mem_true_iff, Hmove in Htest. contradiction. @@ -3327,7 +3373,7 @@ intros da config pt Hssync Hgather. rewrite (round_simplify_Majority). induction names as [| id l]. + reflexivity. + cbn -[equiv_dec]. destruct_match. - - elim Hdiff. hnf in *. subst pt'. pattern id. apply no_byz. intro g. apply Hgather. + - contradiction Hdiff. hnf in *. subst pt'. pattern id. apply no_byz. intro g. apply Hgather. - apply IHl. } rewrite H0. specialize (Hgather g1). rewrite <- Hgather. apply pos_in_config. Qed. @@ -3377,31 +3423,75 @@ destruct (gathered_at_dec config (get_location (config (Good g1)))) as [Hmove | - rewrite Heq. assumption. Qed. +(* *********** UNFAIR DEMON ************** *) + +(* Unfairness in terms of gathering: someone always moves if robots + are not gathered (and position is valid). *) +Definition unfair_da_gathering r da config := + ~WithMultiplicity.invalid config -> + (forall pt, ~gathered_at pt config) -> + moving r da config <> nil. + +Definition unfair_gathering r d config := + Stream.forever2 (Stream.instant2 (unfair_da_gathering r)) d (execute r d config). + +(* We link this definition to the usual notion of unfairness: if a + robot can move, then some robot moves *) +Definition sim_da_with_all_activated da : similarity_da. +Proof. +exists (da_with_all_activated da). +apply (proj2_sig da). +Defined. -(** Let us change the assumption over the demon, it is no longer fair - but instead activates at least a robot that should move at each round. *) -Definition OKunfair r := - Stream.forever (Stream.instant (fun da => forall config, ~WithMultiplicity.invalid config -> moving r da config <> nil)). +Lemma unfair_da_gather_impl: forall da config, + unfair_da gatherR2 (proj_sim_da da) config -> + unfair_da_gathering gatherR2 da config. +Proof using . +unfold unfair_da, unfair_da_gathering. +intros da config Hunfair Hvalid Hgather. +apply Hunfair. +destruct (OneMustMove (Good g1) Hvalid (Hgather (get_location (config (Good g1))))) + as [gmove Hgmove]. +assert (Hall : List.In gmove (active (proj_sim_da (sim_da_with_all_activated da)))). +{ unfold active. simpl. rewrite List.filter_In. auto using In_names. } +specialize (Hgmove (sim_da_with_all_activated da) + (FSYNC_SSYNC_da (da_with_all_activated_FSYNC_da da)) Hall). +rewrite moving_spec in Hgmove. +intro Habs. apply no_moving_same_config in Habs. apply Hgmove, Habs. +Qed. + +Lemma unfair_gather_impl : forall d config, + unfair gatherR2 (similarity_demon2demon d) config -> + unfair_gathering gatherR2 d config. +Proof using . +coinduction cfsd. ++ destruct H. clear -H. simpl in *. now apply unfair_da_gather_impl. ++ now destruct H. +Qed. -Theorem unfair_Gathering_in_R2 : - forall d, SSYNC (similarity_demon2demon d) -> OKunfair gatherR2 d -> WithMultiplicity.ValidSolGathering gatherR2 d. +(* Final theorem for unfair demons. *) +Theorem unfair_Gathering_in_R2 : forall d config, + SSYNC (similarity_demon2demon d) -> + unfair gatherR2 (similarity_demon2demon d) config -> + ~WithMultiplicity.invalid config -> WillGather (execute gatherR2 d config). Proof using . -intros d Hssync Hunfair config. revert d Hssync Hunfair. pattern config. +intros d config Hssync Hunfair' Hvalid. +assert (Hunfair : unfair_gathering gatherR2 d config) by now apply unfair_gather_impl. +clear Hunfair'. +revert d Hssync Hunfair Hvalid. pattern config. apply (well_founded_ind wf_lt_config). clear config. -intros config Hind d Hssync Hunfair Hok. +intros config Hind d Hssync Hunfair Hvalid. (* Are we already gathered? *) destruct (gathered_at_dec config (get_location (config (Good g1)))) as [Hmove | Hmove]. + (* If so, not much to do *) apply Stream.Now. exists (get_location (config (Good g1))). now apply gathered_at_OK. + (* Otherwise, by assumption on the demon, a robot should move - so we can use our well-founded induction hypothesis. *) + so we can u se our well-founded induction hypothesis. *) destruct Hunfair as [Hactive Hunfair], Hssync. hnf in Hactive. - apply Stream.Later, Hind. - - apply round_lt_config; auto. - - assumption. - - assumption. - - now apply never_invalid. + apply Stream.Later, Hind; trivial; try (now apply never_invalid); []. + apply round_lt_config; trivial; []. + apply Hactive; trivial; []. intros pt Habs. apply Hmove. eapply gathered_precise, Habs. Qed. End GatheringInR2. diff --git a/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v b/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v index 2146fc56871a58cd5ffcb80c752ebf59358dac07..57af3eae0f947e5a5d79407b52f6af0bada36973 100644 --- a/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v +++ b/CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v @@ -10,7 +10,6 @@ (**************************************************************************) Require Import Bool. -Require Import Arith.Div2. Require Import Lia Field. Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def. Require Import List. @@ -20,7 +19,8 @@ Require Import RelationPairs. Require Import Morphisms. Require Import Psatz. Require Import Inverse_Image. -Require Import Pactole.Setting. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Setting Pactole.Util.Fin. Require Import Pactole.Models.Flexible. Require Import Pactole.Models.NoByzantine. Require Import Pactole.Observations.SetObservation. @@ -29,7 +29,7 @@ Require Import Pactole.CaseStudies.Gathering.Definitions. Set Implicit Arguments. -Typeclasses eauto := (bfs) 10. +(* Typeclasses eauto := (bfs) 10. *) (** * The Gathering Problem **) @@ -42,11 +42,10 @@ Typeclasses eauto := (bfs) 10. (** ** Framework of the correctness proof: a finite set with at least two elements **) Section GatheringInR2. -Variable n : nat. -Hypothesis size_G : (2 <= n)%nat. +Context {k : nat} {ltc_1_k : 1 <c k}. (** There are n good robots and no byzantine ones. *) -Instance MyRobots : Names := Robots n 0. +Instance MyRobots : Names := Robots k 0. Instance NoByz : NoByzantine. Proof using . now split. Qed. @@ -85,7 +84,7 @@ Instance Robot : robot_choice (path location) := { robot_choice_Setoid := path_S (** In a flexible setting, the minimum distance that robots are allowed to move is delta. *) Variable delta : R. -Lemma inactive_proper : Proper (equiv ==> eq ==> equiv ==> equiv) +Lemma inactive_proper : Proper (equiv ==> eq ==> equiv ==> @equiv location location_Setoid) (fun (config : configuration) (id : ident) (_ : unit) => config id). Proof using . repeat intro; subst; auto. Qed. @@ -115,7 +114,8 @@ Arguments dist : simpl never. Definition path_R2 := path location. Definition paths_in_R2 : location -> path_R2 := local_straight_path. (* TODO: understand the warning here. *) -Coercion paths_in_R2 : location >-> path_R2. +#[local,warnings="-uniform-inheritance"] + Coercion paths_in_R2 : location >-> path_R2. Instance paths_in_R2_compat : Proper (@equiv _ location_Setoid ==> equiv) paths_in_R2. Proof using . intros pt1 pt2 Heq. now rewrite Heq. Qed. @@ -128,31 +128,28 @@ Proof using . intros ? ? Heq. apply no_byz_eq. intro. apply Heq. Qed. Lemma config_list_alls : forall pt, config_list (fun _ => pt) = alls pt nG. Proof using . intro. rewrite config_list_spec, map_cst. -setoid_rewrite names_length. simpl. now rewrite plus_0_r. +rewrite names_length. simpl. now rewrite Nat.add_0_r. Qed. (** Define one robot to get their location whenever they are gathered. *) -Definition g1 : G. -Proof. -exists 0%nat. abstract (pose (Hle := size_G); lia). -Defined. +Definition g1 : G := fin0. (* Definition Spect_map f s := Spect.M.fold (fun e acc => Spect.M.add (f e) acc) s Spect.M.empty. *) Lemma map_sim_elements : forall (sim : similarity location) s, PermutationA equiv (elements (map sim s)) (List.map sim (elements s)). -Proof using . intros. apply map_injective_elements; autoclass; apply Similarity.injective. Qed. +Proof using . intros. eapply map_injective_elements; autoclass; apply Similarity.injective. Qed. (** Spectra can never be empty as the number of robots is non null. *) Lemma obs_non_nil : forall config, !! config =/= empty. -Proof using size_G. +Proof using ltc_1_k. intros config Habs. specialize (Habs (get_location (config (Good g1)))). rewrite empty_spec in Habs. rewrite <- Habs. apply pos_in_config. Qed. Lemma elements_non_nil : forall config, elements (!! config) <> nil. -Proof using size_G. intro. rewrite elements_nil. apply obs_non_nil. Qed. +Proof using ltc_1_k. intro. rewrite elements_nil. apply obs_non_nil. Qed. (* We need to unfold [obs_is_ok] for rewriting *) Definition obs_from_config_spec : forall (config : configuration) pt, @@ -219,17 +216,17 @@ Instance measure_compat : Proper (equiv ==> eq) measure. Proof using . intros ? ? Heq. unfold measure. now rewrite Heq. Qed. Lemma measure_nonneg : forall config, 0 <= measure config. -Proof using size_G. +Proof using ltc_1_k. intros config. unfold measure. destruct (elements (!! config)) as [| pt l] eqn:Heq. -+ elim (elements_non_nil _ Heq). ++ contradiction (elements_non_nil _ Heq). + rewrite <- (R2_dist_defined_2 pt). apply max_dist_obs_le; rewrite Heq; now left. Qed. (** The minimum value 0 is reached only on gathered configurations. *) Lemma gathered_elements : forall config pt, gathered_at pt config <-> PermutationA (@equiv _ location_Setoid) (elements (!! config)) (pt :: nil). -Proof using size_G. +Proof using ltc_1_k. intros config pt. split; intro H. * apply NoDupA_equivlistA_PermutationA; autoclass. @@ -251,7 +248,7 @@ split; intro H. Qed. Lemma gathered_measure : forall config, measure config = 0%R <-> exists pt, gathered_at pt config. -Proof using size_G. +Proof using ltc_1_k. intros config. split; intro H. * unfold measure, max_dist_obs in *. assert (Hnil := elements_non_nil config). @@ -268,7 +265,7 @@ intros config. split; intro H. cut (length (pt' :: l) = length (x :: nil)); try (simpl; lia). f_equiv; eauto. } subst. rewrite PermutationA_1 in Hperm; autoclass; []. - elim H2. left. + contradiction H2. left. cbn in H. do 2 rewrite R2_dist_defined_2 in H. cbn in H. setoid_rewrite (Rmax_comm 0)%R in H. rewrite (Rmax_left 0 0)%R in H; try reflexivity; []. rewrite dist_sym in H. repeat (rewrite (Rmax_left (dist pt' pt) 0) in H; try apply dist_nonneg). @@ -323,7 +320,10 @@ assert (Hperm : PermutationA equiv (List.map sim (elements (!! config))) (elements (!! (map_config sim config)))). { rewrite <- map_injective_elements, obs_from_config_map, obs_from_config_ignore_snd; autoclass; reflexivity || apply Bijection.injective. } +Typeclasses eauto := (bfs) 10. rewrite (obs_from_config_ignore_snd origin). +Typeclasses eauto := (dfs). + change (lift (existT _ ?x _)) with x. change get_location with (@Datatypes.id location). unfold Datatypes.id. simpl pgm. unfold ffgatherR2_pgm. changeR2. @@ -342,7 +342,6 @@ apply (update_compat config); auto. + rewrite <- (map_config_id config) at 2. rewrite map_config_merge. - f_equiv. intros x y Hxy. simpl. now rewrite Bijection.retraction_section. - - auto. - simpl. repeat intro. now subst. + unfold lift_path; cbn -[straight_path isobarycenter]. intro. now rewrite Bijection.retraction_section, Hda. @@ -356,7 +355,7 @@ Theorem round_lt_config : forall da config, FSYNC_da da -> delta <= measure config -> measure (round ffgatherR2 da config) <= measure config - delta. -Proof using size_G. +Proof using ltc_1_k. intros da config Hdelta HFSync Hnotdone. set (elems := elements (!! config)). set (C := isobarycenter elems). @@ -573,7 +572,7 @@ Theorem round_last_step : forall da config, FSYNC_da da -> measure config <= delta -> measure (round ffgatherR2 da config) == 0. -Proof using size_G. +Proof using ltc_1_k. intros da config Hdelta HFSync Hlt. unfold measure. set (elems := (elements (!! config))). @@ -633,7 +632,7 @@ Proof using . unfold lt_config. apply wf_inverse_image, lt_wf. Qed. Lemma lt_config_decrease : 0 < delta -> forall config1 config2, measure config1 <= measure config2 - delta -> lt_config delta config1 config2. -Proof using size_G. +Proof using ltc_1_k. intros Hdelta config1 config2 Hle. unfold lt_config. rewrite <- Z2Nat.inj_lt. + apply Zup_lt. field_simplify; try lra. unfold Rdiv. apply Rmult_le_compat. @@ -653,7 +652,7 @@ Qed. Lemma gathered_precise : forall config pt, gathered_at pt config -> forall id, gathered_at (config id) config. -Proof using size_G. +Proof using . intros config pt Hgather id id'. transitivity pt. - apply Hgather. - symmetry. apply (no_byz id), Hgather. @@ -668,7 +667,7 @@ Lemma not_gathered_exists : forall config pt, Proof. intros config pt Hgather. destruct (forallb (fun x => R2dec_bool (config x) pt) names) eqn:Hall. -- elim Hgather. rewrite forallb_forall in Hall. +- contradiction Hgather. rewrite forallb_forall in Hall. intro id'. setoid_rewrite R2dec_bool_true_iff in Hall. hnf. repeat rewrite Hall; trivial; apply In_names. - rewrite <- negb_true_iff, existsb_forallb, existsb_exists in Hall. @@ -679,7 +678,7 @@ Qed. *) Lemma gathered_at_elements : forall config pt, gathered_at pt config -> PermutationA equiv (elements (!! config)) (pt :: nil). -Proof using size_G. +Proof using ltc_1_k. intros config pt Hgather. apply NoDupA_equivlistA_PermutationA; autoclass. + apply elements_NoDupA. @@ -694,7 +693,7 @@ Qed. Lemma gathered_at_forever : forall da config pt, FSYNC_da da -> gathered_at pt config -> gathered_at pt (round ffgatherR2 da config). -Proof using size_G. +Proof using ltc_1_k. intros da config pt Hda Hgather g. rewrite round_simplify; trivial; []. cbn zeta. rewrite (gathered_at_elements Hgather), isobarycenter_singleton. @@ -703,7 +702,7 @@ Qed. Lemma gathered_at_OK : forall d conf pt, FSYNC (similarity_demon2demon d) -> gathered_at pt conf -> Gather pt (execute ffgatherR2 d conf). -Proof using size_G. +Proof using ltc_1_k. cofix Hind. intros d conf pt [] Hgather. constructor. + clear Hind. simpl. assumption. + rewrite execute_tail. apply Hind; now try apply gathered_at_forever. @@ -738,7 +737,7 @@ Qed. (** The final theorem. *) Theorem FSGathering_in_R2 : forall d, delta > 0 -> FSYNC (similarity_demon2demon d) -> FullSolGathering ffgatherR2 d. -Proof using size_G. +Proof using ltc_1_k. intros d Hdelta HFS config. revert d HFS. pattern config. apply (well_founded_ind (wf_lt_config Hdelta)). clear config. intros config Hind [da d] [Hda HFS]. diff --git a/CaseStudies/Gathering/InR2/Peleg.v b/CaseStudies/Gathering/InR2/Peleg.v index 80d03a76172ba3d63b6a964b2222758b9f84ee1d..eba8f8d98228828c3d9318eb99a94599848251d5 100644 --- a/CaseStudies/Gathering/InR2/Peleg.v +++ b/CaseStudies/Gathering/InR2/Peleg.v @@ -10,7 +10,6 @@ (**************************************************************************) Require Import Bool. -Require Import Arith.Div2. Require Import Lia Field. Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def. Require Import List. @@ -21,6 +20,7 @@ Require Import Morphisms. Require Import Psatz. Require Import Inverse_Image. Require Import Pactole.Setting. +Require Import Pactole.Util.Bijection. Require Import Pactole.Models.Flexible. Require Import Pactole.Models.NoByzantine. Require Import Pactole.Observations.MultisetObservation. @@ -31,7 +31,7 @@ Require Import Pactole.CaseStudies.Gathering.Definitions. Import Permutation. Set Implicit Arguments. -Typeclasses eauto := (bfs) 10. +(* Typeclasses eauto := (bfs) 10. *) (** * The Gathering Problem **) @@ -111,7 +111,8 @@ Arguments dist : simpl never. (* The robot trajectories are straight paths. *) Definition path_R2 := path location. Definition paths_in_R2 : location -> path_R2 := local_straight_path. -Coercion paths_in_R2 : location >-> path_R2. +#[local,warnings="-uniform-inheritance"] + Coercion paths_in_R2 : location >-> path_R2. Instance paths_in_R2_compat : Proper (@equiv _ location_Setoid ==> equiv) paths_in_R2. Proof using . intros pt1 pt2 Heq. now rewrite Heq. Qed. @@ -125,8 +126,11 @@ Proof using . intros. apply no_byz_eq. intro. now apply WithMultiplicity.no_info Lemma config_list_alls : forall pt, config_list (fun _ => pt) = alls pt nG. Proof using . intro. rewrite config_list_spec, map_cst. -setoid_rewrite names_length. simpl. now rewrite plus_0_r. +Typeclasses eauto := (bfs) 10. +setoid_rewrite names_length. simpl. now rewrite Nat.add_0_r. Qed. +Typeclasses eauto := (dfs). + (** Define one robot to get their location whenever they are gathered. *) Definition g1 : G. @@ -189,7 +193,7 @@ Lemma max_dist_obs_le : InA equiv pt0 (support s) -> InA equiv pt1 (support s) -> dist pt0 pt1 <= max_dist_obs s. -Proof using . intros. now apply max_dist_list_list_le. Qed. +Proof using . intros. now apply @max_dist_list_list_le. Qed. Lemma max_dist_obs_ex : forall (s : observation), @@ -214,7 +218,7 @@ Lemma measure_nonneg : forall config, 0 <= measure config. Proof using size_G. intros config. unfold measure. destruct (support (!! config)) as [| pt l] eqn:Heq. -+ rewrite support_nil in Heq. elim (obs_non_nil _ Heq). ++ rewrite support_nil in Heq. contradiction (obs_non_nil _ Heq). + rewrite <- (R2_dist_defined_2 pt). apply max_dist_obs_le; rewrite Heq; now left. Qed. @@ -258,7 +262,7 @@ intros config. split; intro H. cut (length (pt' :: l) = length (x :: nil)); try (simpl; lia). f_equiv; eauto. } subst. rewrite PermutationA_1 in Hperm; autoclass; []. - elim H2. left. + contradiction H2. left. cbn in H. do 2 rewrite R2_dist_defined_2 in H. rewrite (Rmax_left 0 0) in H; try reflexivity; []. rewrite dist_sym in H. rewrite (Rmax_comm 0) in H. @@ -329,7 +333,9 @@ assert (Hperm : PermutationA (equiv * eq)%signature (precondition_satisfied da config g))) config)))). { rewrite <- map_injective_elements, obs_from_config_map, obs_from_config_ignore_snd; autoclass; reflexivity || apply Bijection.injective. } -rewrite obs_from_config_ignore_snd. +Typeclasses eauto := (bfs) 10. +rewrite @obs_from_config_ignore_snd. +Typeclasses eauto := (dfs). simpl pgm. unfold ffgatherR2_pgm. changeR2. remember (elements (!! config)) as E. apply (@PermutationA_map _ _ _ (equiv * eq)%signature _ (fun xn => (fst xn, INR (snd xn)))) in Hperm; @@ -362,7 +368,6 @@ apply get_location_compat, update_compat; auto. (lift (existT precondition (frame_choice_bijection sim) (precondition_satisfied da config g)) x))); try reflexivity; []. rewrite 2 get_location_lift. simpl. rewrite Bijection.retraction_section. apply Hxy. - - autoclass. - apply lift_compat. intros x y Hxy. now rewrite Hxy. + Admitted. (* Peleg's gathering in FSYNC: round_simplify -> hypothesis missing on the demon *) @@ -708,7 +713,7 @@ Lemma not_gathered_exists : forall config pt, Proof. intros config pt Hgather. destruct (forallb (fun x => R2dec_bool (config x) pt) names) eqn:Hall. -- elim Hgather. rewrite forallb_forall in Hall. +- contradiction Hgather. rewrite forallb_forall in Hall. intro id'. setoid_rewrite R2dec_bool_true_iff in Hall. hnf. repeat rewrite Hall; trivial; apply In_names. - rewrite <- negb_true_iff, existsb_forallb, existsb_exists in Hall. diff --git a/CaseStudies/Gathering/InR2/Viglietta.v b/CaseStudies/Gathering/InR2/Viglietta.v index 085802168342434894cbe81ee1650203b43aa45e..fac80ea1d715e1461deb24dd767be53ec83f4c52 100644 --- a/CaseStudies/Gathering/InR2/Viglietta.v +++ b/CaseStudies/Gathering/InR2/Viglietta.v @@ -10,7 +10,7 @@ - me=B, you=B: stay, light goes to A *) Require Import Lia. -Require Import Pactole.Setting. +Require Import Pactole.Setting Pactole.Util.Fin. Require Import Pactole.Spaces.R. Require Import Pactole.Spaces.R2. Require Import Pactole.Models.Similarity. @@ -18,7 +18,6 @@ Require Import Pactole.Models.NoByzantine. Require Import Pactole.CaseStudies.Gathering.Definitions. (* Helping typeclass resolution avoid infinite loops. *) -Typeclasses eauto := (bfs). (* Avoid problems with previous instances. *) Remove Hints eq_setoid : typeclass_instances. @@ -34,16 +33,16 @@ Instance NoByz : NoByzantine. Proof using . now split. Qed. Notation lt2 := (fun x => x < 2)%nat. -Definition lt02 : (0 < 2)%nat := ltac:(abstract lia). +Definition lt02 : (0 < 2)%nat := ltac:(abstract auto). Definition lt12 : (1 < 2)%nat := ltac:(abstract lia). -Definition r0 : G := exist lt2 0%nat lt02. -Definition r1 : G := exist lt2 1%nat lt12. +Definition r0 : G := Fin lt02. +Definition r1 : G := Fin lt12. Lemma id_case : forall id, id = Good r0 \/ id = Good r1. Proof using . intros [[[| [| ?]] ?] | []]; simpl; solve [ exfalso; lia - | now left + right; f_equal; apply eq_proj1 ]. + | now left + right; f_equal; apply fin2natI ]. Qed. (** The space is R², so that we can define similarities and they respect middle points. *) @@ -211,7 +210,7 @@ assert (Hpt1 : pt1 == pt) by now rewrite <- Hgather, Hr0. assert (Hpt2 : pt2 == pt) by now rewrite <- Hgather, Hr1. assert (Hg : g = r0 \/ g = r1). { destruct g as [[| [| ?]] ?]. - - left. unfold r0. f_equal. apply le_unique. + - left. unfold r0. unfold fin0. f_equal. apply le_unique. - right. unfold r1. f_equal. apply le_unique. - exfalso. lia. } destruct Hg; subst g. @@ -275,7 +274,7 @@ destruct (gathered_at_dec config (get_location (config (Good r0)))) as [| Hgathe * right. rewrite round_simplify; trivial; []. assert (Hone_active : activate da (Good r0) = true \/ activate da (Good r1) = true). { destruct (activate da (Good r0)) eqn:Hr0, (activate da (Good r1)) eqn:Hr1; auto; []. - elim Hmove. apply incl_nil. cut (active da = nil). + contradiction Hmove. apply incl_nil. cut (active da = nil). - intro Hactive. rewrite <- Hactive. apply moving_active. intros id Hid config'. reflexivity. - unfold active, names. simpl. @@ -369,9 +368,9 @@ destruct (config (Good r0)) as [pt1 l1] eqn:Hr0, assert (Hpt : pt1 =/= pt2). { intro Habs. rewrite Habs in Hr0. apply Hnotgather. exists (get_location (config (Good r0))). intros [[| [| ?]] ?]; simpl. - + unfold r0. do 4 f_equal. now apply eq_proj1. + + unfold r0. do 4 f_equal. now apply fin2natI. + transitivity (Datatypes.id (fst (config (Good r1)))). - - unfold r1. simpl. do 4 f_equal. now apply eq_proj1. + - unfold r1. simpl. do 4 f_equal. now apply fin2natI. - now rewrite Hr0, Hr1. + exfalso. lia. } destruct l1 eqn:Hl1; [| destruct l2 eqn:Hl2]. @@ -457,7 +456,7 @@ destruct (gathered_at_dec config (get_location (config (Good r0)))) as [Hgather apply (Hind (measure (round rendezvous (Stream.hd d) config))). ++ subst n. destruct (round_measure (Stream.hd d) config) as [[pt Hpt] | ?]; trivial; [|]. -- apply Hsim. - -- elim Hgather'. now rewrite (Hpt r0). + -- contradiction Hgather'. now rewrite (Hpt r0). ++ apply Hfair. ++ apply Hsim. ++ reflexivity. diff --git a/CaseStudies/Gathering/WithMultiplicity.v b/CaseStudies/Gathering/WithMultiplicity.v index a9903825a3fc4c8b636f2e7e2d44c5de09abb485..e2ca236592e232f6827194b0e38375bcf0b23bb3 100644 --- a/CaseStudies/Gathering/WithMultiplicity.v +++ b/CaseStudies/Gathering/WithMultiplicity.v @@ -27,7 +27,7 @@ Close Scope R_scope. Close Scope VectorSpace_scope. Set Implicit Arguments. Typeclasses eauto := (bfs) 5. -Require Even. +(* Require Even. *) (** Gathering Definitions specific to a setting with multiplicities, i.e. a multiset observation. *) @@ -36,7 +36,7 @@ Section MultisetGathering. (** Here, we restrict the state to only contain the location. *) Context `{Location}. (* TODO: add the existence of a similarity here *) -Instance Info : State location := OnlyLocation (fun _ => True). +#[export]Instance Info : State location := OnlyLocation (fun _ => True). Context {VS : RealVectorSpace location}. Context {RMS : RealMetricSpace location}. Context `{Names}. @@ -46,7 +46,7 @@ Context `{inactive_choice}. Context {UpdFun : update_function _ _ _}. Context {InaFun : inactive_function _}. -Global Existing Instance multiset_observation. +#[export] Existing Instance multiset_observation. Notation "!! config" := (@obs_from_config location _ _ _ multiset_observation config origin : observation) (at level 10). @@ -102,12 +102,12 @@ rewrite <- (@cardinal_total_sub_eq _ _ _ _ _ (add pt2 (Nat.div2 nG) (singleton p + rewrite size_add. destruct (In_dec pt2 (singleton pt1 (Nat.div2 nG))) as [Hin | Hin]. - exfalso. rewrite In_singleton in Hin. - destruct Hin. now elim Hdiff. + destruct Hin. now contradiction Hdiff. - rewrite size_singleton; trivial; []. apply Exp_prop.div2_not_R0. apply HsizeG. - apply Exp_prop.div2_not_R0. apply HsizeG. + intro pt. destruct (pt =?= pt2) as [Heq2 | Heq2], (pt =?= pt1) as [Heq1 | Heq1]. - - rewrite Heq1, Heq2 in *. now elim Hdiff. + - rewrite Heq1, Heq2 in *. now contradiction Hdiff. - rewrite add_spec, singleton_spec. do 2 destruct_match; try contradiction; []. simpl. @@ -143,12 +143,12 @@ destruct Hconfig as [Heven [Hge2 [pt1' [pt2' [Hdiff [Hpt1' Hpt2']]]]]]. assert (Hcase : pt1' == pt1 /\ pt2' == pt2 \/ pt1' == pt2 /\ pt2' == pt1). { assert (Hin1 : InA equiv pt1' (pt1 :: pt2 :: nil)). { rewrite <- Hsupp, support_spec. unfold In. rewrite Hpt1'. - destruct nG as [| [| nG]]; simpl; lia. } + destruct nG as [| [| nG]]; simpl; solve [ auto with arith | lia]. } assert (Hin2 : InA equiv pt2' (pt1 :: pt2 :: nil)). { rewrite <- Hsupp, support_spec. unfold In. rewrite Hpt2'. - destruct nG as [| [| nG]]; simpl; lia. } + destruct nG as [| [| nG]]; simpl; solve [ auto with arith | lia]. } rewrite 2 InA_cons, InA_nil in Hin1, Hin2. clear -Hin1 Hin2 Hdiff. - decompose [or] Hin1; decompose [or] Hin2; tauto || elim Hdiff; etransitivity; eauto. } + decompose [or] Hin1; decompose [or] Hin2; tauto || contradiction Hdiff; etransitivity; eauto. } split. + intro. apply Hdiff. decompose [and or] Hcase; repeat (etransitivity; eauto; symmetry). @@ -158,7 +158,7 @@ split. destruct Hcase as [[Heq1 Heq2] | [Heq1 Heq2]]; rewrite Heq1 in *; rewrite Heq2 in *; try match goal with H : pt == _ |- _ => rewrite H in *; clear H end; - rewrite ?Hpt1', ?Hpt2'; lia || now elim Hdiff. + rewrite ?Hpt1', ?Hpt2'; lia || now contradiction Hdiff. - rewrite cardinal_add, cardinal_singleton, cardinal_obs_from_config, even_div2; auto; lia. Qed. @@ -171,8 +171,8 @@ destruct (invalid_strengthen HnB Hinvalid) as [pta [ptb Hdiff Hobs]]. rewrite Hobs, add_In, In_singleton in Hin1, Hin2, Hin3. destruct Hin1 as [[] | []], Hin2 as [[] | []], Hin3 as [[] | []]; solve [ etransitivity; eauto - | elim Hdiff13; etransitivity; eauto - | elim Hdiff23; etransitivity; eauto ]. + | contradiction Hdiff13; etransitivity; eauto + | contradiction Hdiff23; etransitivity; eauto ]. Qed. Arguments invalid_same_location _ config {pt1} {pt2} pt3 _ _ _ _ _. @@ -210,9 +210,9 @@ destruct (n1 =?= n2) as [Hn | Hn]. eapply proj1. rewrite <- elements_spec. rewrite Helem. now right; left. * right. - intro Hvalid. elim Hn. + intro Hvalid. contradiction Hn. assert (Hhalf : 0 < Nat.div2 nG). - { destruct Hvalid as [_ [Hle _]]. destruct nG as [| [| ?]]; simpl; lia. } + { destruct Hvalid as [_ [Hle _]]. destruct nG as [| [| ?]]; simpl; solve [ auto with arith | lia]. } destruct (invalid_strengthen HnB Hvalid) as [pt1' [pt2' Hdiff Hobs]]. assert (Hperm : PermutationA eq_pair ((pt1, n1) :: (pt2, n2) :: nil) ((pt1', Nat.div2 nG) :: (pt2', Nat.div2 nG) :: nil)). @@ -221,7 +221,7 @@ destruct (n1 =?= n2) as [Hn | Hn]. constructor. + split; simpl; reflexivity || lia. + destruct_match. - - elim Hdiff. hnf in * |-; simpl in *. auto. + - contradiction Hdiff. hnf in * |-; simpl in *. auto. - reflexivity. } rewrite PermutationA_2 in Hperm; autoclass; []. clear -Hperm. destruct Hperm as [[[] []] | [[] []]]; compute in *; congruence. diff --git a/CaseStudies/LifeLine/Algorithm.v b/CaseStudies/LifeLine/Algorithm.v new file mode 100644 index 0000000000000000000000000000000000000000..02b9fc67112ff43d8856bc9f68cd56f72d7608ba --- /dev/null +++ b/CaseStudies/LifeLine/Algorithm.v @@ -0,0 +1,1890 @@ +Require Import Utf8. +Require Import Bool. +Require Import Lia Field. +Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def Lra. +Require Import List SetoidList SetoidDec. +Require Import Relations RelationPairs Morphisms. +Require Import Inverse_Image. +Require Import Decidable. +Require Import Pactole.Util.Preliminary. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Setting. +Require Import Pactole.Observations.SetObservation. +Require Import Pactole.Spaces.R2. +Require Import Pactole.Spaces.Isometry. +Require Import Pactole.Models.Rigid. +Require Import Pactole.Models.Isometry. +Require Import Pactole.Util.Fin. +Import Permutation. +Import Datatypes. (* to recover [id] *) +Close Scope VectorSpace_scope. +Open Scope R_scope. +Set Implicit Arguments. + +(* TODO: rename Dmax *) +(* Keeping connection from a fixed base to a scouting robot. All robots +(including the scout) share the same speed (max distance D in one round) +and the same detection radius (Dmax). The existence of a family +of solution is shown whenever Dmax > 7*D. No optimality is proved. *) + +Section PursuitSetting. + +(* We have a given finite number n of robot. The final lemma will require +that there are enough robots during the execution. *) +Variable n: nat. + +(** All robots are non byzantine, except the scout + (considered as a Byzantine as it does not follow the protocol) *) +Instance Identifiers : Names := Robots n 1. +Definition scout_B : B := Fin Nat.lt_0_1. +Definition scout : ident := Byz scout_B. + +(** The only Byzantine robot is the scout. *) +Lemma byz_is_scout : ∀ b : B, b = scout_B. +Proof using . + intros [b Hb]. change B with (fin 1) in *. apply fin2natI. cbn. + destruct b as [|b]. reflexivity. inv Hb. inv H0. +Qed. + +(** The space is the Euclidean plane *) +Instance Loc : Location := make_Location R2. +Instance VS : RealVectorSpace location := R2_VS. +Instance ES : EuclideanSpace location := R2_ES. +Remove Hints R2_VS R2_ES : typeclass_instances. +Local Opaque dist. +Local Opaque inverse. + +Ltac changeR2 := + change R2 with location in *; + change R2_Setoid with location_Setoid in *; + change R2_EqDec with location_EqDec in *; + change R2_VS with VS in *; + change R2_ES with ES in *. + +(** The robot state is the location + 2 booleans and an integer: alive, light, id *) +(* NB: To avoid confusion with [Names.ident], I use [id]. *) +Record state := { + loc : location; + id : nat; + alive : bool; + light : bool}. +(* Each robots have a light, which state is on or off *) +(* A robot will be alive or not (the latter meaning it has withdrawn +from the task). No collision with dead robots are considered *) +(* A robot may be still at base, waiting to be needed by the task. No +collision implying non launched robot is considered either. *) + +Notation "'{{' x 'with' 'loc' ':=' y '}}'" := + ({| loc := y; id := x.(id); alive := x.(alive); light := x.(light) |}). +Notation "'{{' x 'with' 'alive' ':=' y '}}'" := + ({| loc := x.(loc); id := x.(id); alive := y; light := x.(light) |}). +Notation "'{{' x 'with' 'light' ':=' y '}}'" := + ({| loc := x.(loc); id := x.(id); alive := x.(alive); light := y |}). + +#[refine] +Instance state_Setoid : Setoid state := { + equiv := fun st1 st2 => st1.(loc) == st2.(loc) + /\ st1.(id) = st2.(id) + /\ st1.(alive) = st2.(alive) + /\ st1.(light) = st2.(light) }. +Proof using . split. ++ intros []; hnf; repeat split. ++ intros []; hnf; repeat split; symmetry; intuition. ++ intros []; hnf; repeat split; simpl in *; etransitivity; intuition auto; congruence. +Defined. + +#[refine] +Instance state_EqDec : EqDec state_Setoid := + fun st1 st2 => if Nat.eq_dec st1.(id) st2.(id) then + if bool_dec st1.(alive) st2.(alive) then + if bool_dec st1.(light) st2.(light) then + if st1.(loc) =?= st2.(loc) then left _ + else right _ else right _ else right _ else right _. +Proof using . ++ abstract (unfold equiv; cbn -[equiv]; repeat (split; trivial)). ++ abstract (revert_all; intros Hneq [Habs _]; apply Hneq, Habs). ++ abstract (intros [? [? []]]; contradiction). ++ abstract (intros [? [? []]]; contradiction). ++ abstract (intros [? [? []]]; contradiction). +Defined. + +#[refine] +Instance St : State state := { + get_location := fun st => st.(loc); + precondition := fun _ => True; + lift := fun f st => {| loc := projT1 f st.(loc); id := st.(id); + alive := st.(alive); light := st.(light) |} }. +Proof using . ++ abstract (intros ? []; repeat split; try reflexivity). ++ abstract (now intros ? ? [Heq _]). ++ abstract (intros f1 f2 Hf st1 st2 [? [? []]]; unfold equiv; cbn -[equiv]; auto). +Defined. + +Instance loc_compat : Proper (equiv ==> equiv) loc. +Proof using . intros [] [] [Heq _]. apply Heq. Qed. + +Instance id_compat : Proper (equiv ==> eq) id. +Proof using . intros [] [] [_ [Heq _]]. apply Heq. Qed. + +Instance alive_compat : Proper (equiv ==> eq) alive. +Proof using . intros [] [] [_ [_ [Heq _]]]. apply Heq. Qed. + +Instance light_compat : Proper (equiv ==> eq) light. +Proof using . intros [] [] [_ [_ [_ Heq]]]. apply Heq. Qed. + +(* Trying to avoid notation problem with implicit arguments *) +Notation "!! config" := + (@obs_from_config location _ _ _ set_observation config origin) (at level 10). + +(* Max distance D traveled by a robot between two rounds. This defines + the "collision zone": any two robots closer than that are immediately + considered colliding. *) +Variable D : R. +Hypothesis D_pos : D > 0. + +(** Max detection radius: robots inside this radius are in range, the others are not seen. *) +Variable Dmax : R. +Hypothesis Dmax_6D : Dmax >= 6*D. + +(* In the problem of connection, the free robot (numbered 0) is + launched and alive. By definition of the problem. + -> This is wrong as it is decided by the demain. But we don't need it. *) +(* Hypothesis byz_alive : forall config, (config scout).(alive) = true. *) + +(* Dp is the distance where a robot is still visible but from which it + becomes able to leave the visibility radius at next round. *) +Definition Dp(*ursuit*) := Dmax - D. + +(** The observation is the set of location and light values + of alive robots within range with lower identifier. *) +Definition target : Type := location * bool. +Definition tgt_loc : target -> location := fst. +Definition tgt_light : target -> bool := snd. +Instance target_Setoid : Setoid target := prod_Setoid _ (eq_setoid _). + +Definition mk_tgt : state -> target := fun st => (st.(loc), st.(light)). + +Instance tgt_loc_compat : Proper (equiv ==> equiv) tgt_loc. +Proof using . intros [] [] [Heq _]. apply Heq. Qed. + +Instance tgt_light_compat : Proper (equiv ==> equiv) tgt_light. +Proof using . intros [] [] [_ Heq]. apply Heq. Qed. + +Instance mk_tgt_compat : Proper (equiv ==> equiv) mk_tgt. +Proof using . intros [] [] [? [_ [_ ?]]]. now split. Qed. + +Notation "X <=? Y" := ((Rle_bool X Y)%R) (at level 70): R_scope. + +Definition visible st e := + e.(alive) && (dist e.(loc) st.(loc) <=? Dmax)%R + && (e.(id) <? st.(id)). + + +#[refine] +Instance Obs : Observation := {| + observation := set target; + observation_Setoid := Set_Setoid _; + observation_EqDec := @Set_EqDec _ target_Setoid _ _ FSetList.SetListFacts; + obs_from_config config st := + SetObservation.make_set + (List.map mk_tgt + (List.filter (fun x => x.(alive) && ((dist x.(loc) st.(loc)) <=? Dmax) + && (x.(id) <? st.(id))) + (config_list config))); + obs_is_ok s config st := + forall loclit, In loclit s + <-> exists r, let st' := config r in + loclit == (st'.(loc), st'.(light)) + ∧ st'.(alive) = true ∧ dist st'.(loc) st.(loc) <= Dmax + ∧ (st'.(id) < st.(id))%nat |}. +Proof. all:autoclass. +* intros config1 config2 Hconfig st1 st2 Hst. + apply SetObservation.make_set_compat, eqlistA_PermutationA_subrelation. + eapply map_eqlistA_compat with equiv; autoclass; []. f_equiv. + + intros ? ? Heq. now rewrite Heq, Hst. + + now rewrite (config_list_compat Hconfig). +* intros config st x. + assert (Proper (equiv ==> eq) + (fun x => alive x && Rle_bool (dist (loc x) (loc st)) Dmax && (id x <? id st))). + { intros [] [] [Heq1 [Heq2 [Heq3 _]]]. cbn -[dist] in *. now rewrite Heq1, Heq2, Heq3. } + rewrite SetObservation.make_set_spec, InA_map_iff; autoclass; []. + setoid_rewrite filter_InA; trivial; []. do 2 setoid_rewrite Bool.andb_true_iff. + setoid_rewrite Rle_bool_true_iff. setoid_rewrite Nat.ltb_lt. + split; intro Hex. + + destruct Hex as [st' [Heq [Hin Hst']]]. + rewrite config_list_spec, InA_map_iff in Hin; autoclass; []. + destruct Hin as [r [Hr _]]. + exists r. cbn zeta. rewrite <- Heq, <- Hr. split; try reflexivity; []. + now rewrite Hr. + + destruct Hex as [r [Heq Hr]]. + exists (config r). rewrite Heq. split; try reflexivity; []. + change state_Setoid with State.state_Setoid. rewrite config_list_InA. + now split; try (exists r). +Defined. + +Notation dist := (@dist location location_Setoid location_EqDec VS _). +Notation obs_from_config := (@obs_from_config state Loc St Identifiers Obs). + +Lemma obs_from_config_spec : forall config st tgt, + (tgt ∈ obs_from_config config st)%set + <-> ∃ r, let st' := config r in + tgt == (loc st', light st') + ∧ alive st' = true ∧ dist (loc st') (loc st) <= Dmax ∧ (id st' < id st)%nat. +Proof using . +intros config st tgt. +assert (Hobs := obs_from_config_spec config st tgt). apply Hobs. +Qed. + +Definition iso_precondition (iso : isometry location) : precondition iso := I. + +Definition lift_target (f : location -> location) : target -> target := + fun tgt => (f (tgt_loc tgt), tgt_light tgt). + +Instance lift_target_compat : Proper ((equiv ==> equiv) ==> equiv ==> equiv) lift_target. +Proof using . intros f g Hfg st1 st2 []. split; trivial; []. now apply Hfg. Qed. + +Lemma obs_from_config_map : forall (iso : isometry location) Piso config pt, + map (lift_target iso) (obs_from_config config pt) + == obs_from_config (map_config (lift (existT precondition iso Piso)) config) + (lift (existT precondition iso Piso) pt). +Proof using . +repeat intro. rewrite map_spec; autoclass; []. +setoid_rewrite obs_from_config_spec. +split; intro Hex. ++ destruct Hex as [tgt [[r [Heq [Halive [Hdist Hid]]]] Heq']]. + exists r. split; [| repeat split]. + - rewrite Heq', Heq. cbn -[equiv]. now split. + - cbn. assumption. + - cbn. changeR2. now rewrite dist_prop. + - cbn. assumption. ++ destruct Hex as [r [Heq [Halive [Hdist Hid]]]]. cbn -[equiv] in *. + exists (mk_tgt (config r)). split; try (exists r; split; [| repeat split]; auto). + - now rewrite dist_prop in Hdist. + - apply Heq. +Qed. + +Lemma in_range_dist : forall config st tgt, + (tgt ∈ obs_from_config config st)%set -> dist (tgt_loc tgt) st.(loc) <= Dmax. +Proof using . +intros config st tgt Hin. rewrite obs_from_config_spec in Hin. +destruct Hin as [r [Heq [Halive [Hdist Hid]]]]. rewrite Heq. apply Hdist. +Qed. + +(** The robot action consists in dying, turning its light on or moving with its light off *) +Inductive action := + | Die + | LightOn + | Move (loc : location). + +#[refine] +Instance action_Setoid : Setoid action := { + equiv := fun act1 act2 => match act1, act2 with + | Die, Die | LightOn, LightOn => True + | Move loc1, Move loc2 => loc1 == loc2 + | _, _ => False + end}. +Proof. split. ++ abstract (intros []; auto). ++ abstract (intros [] []; cbn -[equiv]; auto). ++ abstract (intros [] [] [] ? ?; cbn -[equiv] in *; auto; etransitivity; eauto). +Defined. + +Instance Action : robot_choice action := { robot_choice_Setoid := action_Setoid }. + +(** The demon's only choices are the location of the scout and the local frame of reference. + In particular, it does not perform any choice to update the state of a robot *) +Instance Upt : update_choice unit := NoChoice. +Instance Ina : inactive_choice unit := NoChoiceIna. + +(** The moves are rigid and are of distance at most D *) +#[refine] +Instance InaFun : inactive_function unit := {inactive := fun config id _ => config id }. +Proof. abstract (intros ? ? Heq ? ? ? ? ? ?; subst; apply Heq). Defined. + +#[refine] +Instance UpdFun : update_function action (isometry location) unit := { + update := fun config g _ act _ => + match act with + | Die => {{ config (Good g) with alive := false }} + | LightOn => {{ config (Good g) with light := true }} + | Move pt => if Rle_dec (dist pt origin) D then + {| loc := pt; + id := (config (Good g)).(id); + alive := (config (Good g)).(alive); + light := false |} + else config (Good g) + end }. +Proof. +intros config1 config2 Hconfig ? g ? ? ? ? [] [] Hact ? ? ?; subst; try tauto; [| |]. ++ repeat split; cbn -[dist equiv]; simpl in Hact; auto; f_equiv; apply Hconfig. ++ repeat split; cbn -[dist equiv]; simpl in Hact; auto; f_equiv; apply Hconfig. ++ unfold equiv in Hact; cbn -[equiv] in Hact; rewrite Hact. + destruct_match; try apply Hconfig; []. + repeat split; trivial; cbn -[equiv]; f_equiv; apply Hconfig. +Defined. + +Lemma rigid_update : forall config (frame : isometry location) g act choice, + get_location (update config g frame act choice) + == match act with + | Die => get_location (config (Good g)) + | LightOn => get_location (config (Good g)) + | Move pt => if Rle_dec (dist pt origin) D then pt else get_location (config (Good g)) + end. +Proof using . intros config frame g [] ?; cbn -[dist equiv]; try destruct_match; reflexivity. Qed. + + +(** ** Definition of the abstract robogram **) + +(** *** Parameters that the robogram relies on **) + +Class Param := { + (* which robot to follow? *) + choose_target : observation -> target; + (* where to go to follow the target? *) + choose_new_pos : observation -> location -> location; + (* Compatibility properties *) + choose_target_compat : Proper (equiv ==> equiv) choose_target; + choose_new_pos_compat : Proper (equiv ==> equiv ==> equiv) choose_new_pos; + (* Specifications *) + choose_target_spec : forall obs, + obs =/= {}%set -> + let tgt := choose_target obs in + (* the target must be in range *) + In tgt obs + (* the target must preferably have its light off *) + ∧ (tgt_light tgt = true -> forall r, In r obs -> tgt_light r = true) + (* the target must preferably be close *) + ∧ (tgt_light tgt = true -> dist origin (tgt_loc tgt) > Dp -> + forall r, In r obs -> dist origin (tgt_loc r) > Dp); + (* The location should be reachable in a round and be within range Dp of the target. *) + choose_new_pos_spec : forall obs pt, + dist pt origin <= Dmax -> + let new := choose_new_pos obs pt in + dist new pt <= Dp /\ dist new origin <= D }. + +Global Existing Instance choose_target_compat. +Global Existing Instance choose_new_pos_compat. + + +(** Is it safe to move? *) +Definition move_to obs pt := + negb (exists_ (fun x => Rle_bool (dist (tgt_loc x) pt) (2*D)) obs). + +Lemma move_to_false : forall obs pt, + move_to obs pt = false <-> ∃ other, (other ∈ obs)%set ∧ dist (tgt_loc other) pt <= 2*D. +Proof using . +intros obs pt. unfold move_to. rewrite Bool.negb_false_iff. +rewrite exists_spec. ++ split; intros [tgt [Hin Hdist]]; + exists tgt; split; trivial; now rewrite Rle_bool_true_iff in *. ++ intros tgt1 tgt2 Htgt. now rewrite Htgt. +Qed. + +Lemma move_to_true : forall obs pt, + move_to obs pt = true <-> forall x, In x obs -> dist (tgt_loc x) pt > 2*D. +Proof using . +intros obs pt. rewrite <- not_false_iff_true, move_to_false. split; intro Hspec. ++ intros tgt Hin. destruct (Rlt_le_dec (2*D) (dist (tgt_loc tgt) pt)) as [Hlt | Hle]. + - assumption. + - contradiction Hspec. now exists tgt. ++ intros [tgt [Hin Habs]]. apply Hspec in Hin. + apply (Rlt_irrefl (2*D)). eapply Rlt_le_trans; eauto. +Qed. + +Instance move_to_compat : Proper (equiv ==> equiv ==> eq) move_to. +Proof using . +intros obs1 obs2 Hobs tgt1 tgt2 Htgt. +destruct (move_to obs2 tgt2) eqn:Hex. ++ rewrite move_to_true in *. intro x. rewrite Hobs, Htgt. now revert x. ++ rewrite move_to_false in *. setoid_rewrite Hobs. setoid_rewrite Htgt. assumption. +Qed. + +Lemma move_to_isometry_compat : forall (iso : isometry R2) obs pt, + move_to (map (lift_target iso) obs) (iso pt) = move_to obs pt. +Proof using . +intros iso obs pt. destruct (move_to obs pt) eqn:Hcase. ++ rewrite move_to_true in *. + intros x' Hx'. rewrite map_spec in Hx'; autoclass; []. + destruct Hx' as [x [Hin Hx]]. rewrite Hx. cbn -[dist]. rewrite dist_prop. auto. ++ rewrite move_to_false in *. destruct Hcase as [tgt [Hin Htgt]]. + exists (iso (tgt_loc tgt), tgt_light tgt). cbn -[dist]. rewrite dist_prop. + split; trivial; []. rewrite map_spec; autoclass. +Qed. + +(** If a robot is too close, the observing robot should die *) + +Definition should_die obs pt := + is_empty obs || + exists_ (fun x => Rle_bool (dist (tgt_loc x) pt) D) obs. + +Lemma should_die_true : forall obs pt, + should_die obs pt = true + <-> obs == {}%set ∨ ∃ tgt, (tgt ∈ obs)%set /\ dist (tgt_loc tgt) pt <= D. +Proof using . +intros obs pt. unfold should_die. rewrite Bool.orb_true_iff, is_empty_spec, exists_spec. ++ split; intros [Hempty | [tgt [Hin Hdist]]]; + (now left) || right; exists tgt; split; trivial; now rewrite Rle_bool_true_iff in *. ++ intros tgt1 tgt2 Htgt. now rewrite Htgt. +Qed. + +Lemma should_die_false : forall obs pt, + should_die obs pt = false + <-> (∃ tgt, (tgt ∈ obs)%set) ∧ ∀ tgt, (tgt ∈ obs)%set -> dist (tgt_loc tgt) pt > D. +Proof using . +intros obs pt. rewrite <- not_true_iff_false, should_die_true. split; intro Hspec. +* destruct (choose obs) as [elt |] eqn:Hchoose. + + split; try (now exists elt; apply choose_1); []. + intros tgt Hin. destruct (Rlt_le_dec D (dist (tgt_loc tgt) pt)) as [Hlt | Hle]. + - assumption. + - contradiction Hspec. right. now exists tgt. + + contradiction Hspec. left. intro x. apply choose_2 in Hchoose. specialize (Hchoose x). + split; intro Habs. + - contradiction. + - now apply empty_spec in Habs. +* intros [Hempty | [tgt [Hin Habs]]]. + + destruct Hspec as [[tgt Htgt] _]. + specialize (Hempty tgt). rewrite empty_spec in Hempty. now apply Hempty. + + apply Hspec in Hin. apply (Rlt_irrefl D). eapply Rlt_le_trans; eauto. +Qed. + +Instance should_die_compat : Proper (equiv ==> equiv ==> eq) should_die. +Proof using . +intros obs1 obs2 Hobs pt1 pt2 Hpt. +destruct (should_die obs2 pt2) eqn:Hex; +rewrite should_die_true in * || rewrite should_die_false in *; +setoid_rewrite Hobs; setoid_rewrite Hpt; assumption. +Qed. + +Lemma should_die_map : forall (iso : isometry location) obs pt, + should_die (map (lift_target iso) obs) (iso pt) = should_die obs pt. +Proof using . +intros iso obs pt. destruct (should_die obs pt) eqn:Hcase. ++ rewrite should_die_true in *. destruct Hcase as [Hcase | [tgt [Hin Hdist]]]. + - left. intro x. rewrite map_spec, empty_spec; autoclass; []. + split; try tauto; []. intros [tgt [Hin Heq]]. + revert Hin. rewrite Hcase. apply empty_spec. + - right. exists (lift_target iso tgt). cbn. rewrite map_spec, dist_prop; autoclass. ++ rewrite should_die_false in *. destruct Hcase as [[tgt Htgt] Hall]. split. + - exists (lift_target iso tgt). rewrite map_spec; autoclass. + - intros x Hin. rewrite map_spec in Hin; autoclass; []. destruct Hin as [x' [Hin' Heq]]. + rewrite Heq. cbn. rewrite dist_prop. now apply Hall. +Qed. + +Context {robogram_args : Param}. + +Corollary target_in_range : forall obs, obs =/= {}%set -> + In (choose_target obs) obs. +Proof using . apply choose_target_spec. Qed. + +Corollary target_has_lower_id : forall st config, + let obs := obs_from_config config st in + let tgt := choose_target obs in + obs =/= {}%set -> + exists r, mk_tgt r == tgt /\ (r.(id) < st.(id))%nat. +Proof using . +intros st config obs tgt Hobs. +assert (Hin := target_in_range Hobs). +assert (Hspec := obs_from_config_spec config st (choose_target (obs_from_config config st))). +unfold obs in Hin. rewrite Hspec in Hin. destruct Hin as [r [Heq [Halive [Hloc Hid]]]]. +exists (config r). now split. +Qed. + +Corollary target_is_alive : forall st config, + let obs := obs_from_config config st in + let tgt := choose_target obs in + obs =/= {}%set -> + exists r, mk_tgt r == choose_target obs /\ alive r = true. +Proof using . +intros st config obs tgt Hobs. +assert (Hin := target_in_range Hobs). +assert (Hspec := obs_from_config_spec config st (choose_target (obs_from_config config st))). +unfold obs in Hin. rewrite Hspec in Hin. destruct Hin as [r [Heq [Halive [Hloc Hid]]]]. +exists (config r). now split. +Qed. + +Corollary target_light_on : forall obs, + obs =/= {}%set -> tgt_light (choose_target obs) = true -> + forall r, In r obs -> tgt_light r = true. +Proof using . apply choose_target_spec. Qed. + +Corollary target_is_close : forall obs, obs =/= {}%set -> + tgt_light (choose_target obs) = true -> + dist origin (tgt_loc (choose_target obs)) > Dp -> + forall r, In r obs -> dist origin (tgt_loc r) > Dp. +Proof using . apply choose_target_spec. Qed. + + +(** ** Specification of the problem **) + +Definition no_collision config := forall r r' : ident, r <> r' -> + (config r).(alive) = true -> (config r').(alive) = true -> + get_location (config r) =/= get_location (config r'). + +Definition connected_path config := + forall g, (config (Good g)).(alive) = true -> + ∃ r', (config r').(alive) = true + ∧ dist (get_location (config r')) (get_location (config (Good g))) <= Dmax + ∧ ((config r').(id) < (config (Good g)).(id))%nat. + +Lemma connected_path_iff_obs_non_empty : forall config, + connected_path config <-> forall g, (config (Good g)).(alive) = true -> + obs_from_config config (config (Good g)) =/= {}%set. +Proof using . +intro config. pose (obs g := obs_from_config config (config (Good g))). +split. ++ intros Hconnect g Halive. fold (obs g). apply Hconnect in Halive. + destruct Halive as [r [Halive' [Hdist Hid]]]. + assert (Hin : (mk_tgt (config r) ∈ (obs g))%set). + { unfold obs. rewrite obs_from_config_spec. exists r. now repeat split. } + intro Habs. rewrite Habs in Hin. revert Hin. apply empty_spec. ++ intros Hobs g Halive. apply Hobs in Halive. fold (obs g) in Halive. + destruct (choose_dec (obs g)). + - unfold obs in Hin. rewrite obs_from_config_spec in Hin. + destruct Hin as [r [Heq [Halive' [Hdist Hid]]]]. + exists r. now repeat split. + - contradiction Halive. intro x. rewrite empty_spec. split; tauto || apply Hempty. +Qed. + +Lemma connected_path_iso_compat : + forall (iso : isometry location) (Hiso : @precondition _ _ St iso) config, + connected_path (map_config (lift (existT precondition iso Hiso)) config) + <-> connected_path config. +Proof using . +intros iso Hiso config. rewrite 2 connected_path_iff_obs_non_empty. +simpl alive. unfold map_config at 2. setoid_rewrite <- obs_from_config_map. +split; intros Hobs g Halive; apply Hobs in Halive; clear Hobs. ++ intro Habs. apply Halive. rewrite Habs. reflexivity. ++ intro Habs. apply Halive. intro x. rewrite empty_spec. split; try tauto; []. + intro Hin. specialize (Habs (lift_target iso x)). + rewrite map_spec, empty_spec in Habs; autoclass; []. + rewrite <- Habs. eauto. +Qed. + + +(** ** The abstract robogram **) + +(** The abstract roborgram (protocol) is defined as follows, + it uses the functions described above: two generic ones and two concrete ones. *) +Definition rbg_fun obs : action := + if should_die obs origin then Die else + let target := choose_target obs in + let new_pos := choose_new_pos obs (tgt_loc target) in + if move_to obs new_pos then Move new_pos else LightOn. + +Instance rbg_compat : Proper (equiv ==> equiv) rbg_fun. +Proof using . +intros obs1 obs2 Hobs. unfold rbg_fun. +repeat (rewrite Hobs; try destruct_match; try reflexivity). +Qed. + +Definition rbg : robogram := {| pgm := rbg_fun |}. + +(* Assumption about the environment: + - local frame of reference are centred on the observer + - it is FSYNC + - the scout moves at most at speed D + - the constant part of the scout state keep good values *) +Definition da_assumption (da : demonic_action) := + let reloc cfg := (da.(relocate_byz) cfg scout_B) in + isometry_da_prop da + ∧ FSYNC_da da + ∧ (forall cfg, dist (reloc cfg).(loc) (cfg scout).(loc) <= D) + ∧ (forall cfg, (reloc cfg).(alive) = true) + ∧ (forall cfg, (reloc cfg).(light) = false) + ∧ forall cfg, (reloc cfg).(id) = 0%nat. + +Theorem round_simplify : forall da, da_assumption da -> forall config g, + round rbg da config (Good g) + == let new_frame := change_frame da config g in + let local_config := + map_config (fun st => {{ st with loc := new_frame st.(loc) }}) config in + let local_state := {{ config (Good g) with loc := new_frame (config (Good g)).(loc) }} in + let obs := obs_from_config local_config local_state in + let local_action := rbg obs in + let new_local_state := update local_config g new_frame local_action tt in + {{ new_local_state with loc := new_frame â»Â¹ new_local_state.(loc) }}. +Proof using . intros da [_ [Hactive _]] config g. unfold round. now rewrite Hactive. Qed. + +Ltac simplify_round_at rbg da config g := + let Hround := fresh "Hround" in let new_frame := fresh "new_frame" in + let local_config := fresh "local_config" in let local_state := fresh "local_state" in + let obs := fresh "obs" in + assert (Hround := round_simplify (da := da) ltac:(auto) config g); + pose (new_frame := change_frame da config g); + pose (local_config := map_config (λ st : state, {{st with loc := new_frame (loc st)}}) config); + pose (local_state := {{config (Good g) with loc := new_frame (loc (config (Good g)))}}); + pose (obs := obs_from_config local_config local_state); + change ( + let new_frame := change_frame da config g in + let local_config := + map_config (λ st : state, {{st with loc := new_frame (loc st)}}) config in + let local_state := {{config (Good g) with loc := new_frame (loc (config (Good g)))}} in + let obs := obs_from_config local_config local_state in + let local_action := rbg obs in + let new_local_state := update local_config g new_frame local_action tt in + {{new_local_state with loc := (new_frame â»Â¹) (loc new_local_state)}}) + with ( + let local_action := rbg obs in + let new_local_state := update local_config g new_frame local_action tt in + {{new_local_state with loc := (new_frame â»Â¹) (loc new_local_state)}}) in Hround; + rewrite Hround in *. + +Ltac simplify_round := + lazymatch goal with + | |- context [round ?rbg ?da ?config (Good ?g)] => + simplify_round_at rbg da config g + | H : context [round ?rbg ?da ?config (Good ?g)] |- _ => + simplify_round_at rbg da config g + end. + + +Lemma update_simplify : forall da, da_assumption da -> + forall config, connected_path config -> + forall g, alive (config (Good g)) = true -> + let new_frame := change_frame da config g in + let local_config := map_config (λ st : state, {{st with loc := new_frame (loc st)}}) config in + let local_state := {{config (Good g) with loc := new_frame (loc (config (Good g)))}} in + let obs := obs_from_config local_config local_state in + exists pf, + Rle_dec (dist (choose_new_pos obs (tgt_loc (choose_target obs))) 0) D = left pf. +Proof using . +intros da Hda config Hpath g Hg new_frame local_config local_state obs. +destruct (Rle_dec (dist (choose_new_pos obs (tgt_loc (choose_target obs))) 0) D); eauto; []. +revert_one not. intro Habs. contradiction Habs. +assert (Hobs : obs =/= {}%set). +{ rewrite connected_path_iff_obs_non_empty in Hpath. apply Hpath in Hg. unfold complement in *. + cbn -[equiv] in Hg. + rewrite <- (map_empty_iff_empty (lift_target new_frame)) in Hg; autoclass; []. + rewrite (obs_from_config_map _ I) in Hg. apply Hg. } +assert (Hloc : get_location (local_config (Good g)) == 0%VS). +{ now rewrite <- center_prop, (proj1 Hda). } +assert (Htgt : dist (tgt_loc (choose_target obs)) 0 <= Dmax). +{ assert (Hrange := target_in_range Hobs). + unfold obs in Hrange. rewrite obs_from_config_spec in Hrange. fold obs in Hrange. + destruct Hrange as [r [Heq [Halive [Hdist Hid]]]]. rewrite Heq, <- Hloc. apply Hdist. } +apply (proj2 (choose_new_pos_spec obs (tgt_loc (choose_target obs)) Htgt)). +Qed. + +Ltac simplify_update_at da Hda config Hpath g Hg new_frame local_config local_state obs := + let Hpf := fresh "Hpf" in let Hrew := fresh "Hrew" in + destruct (update_simplify Hda Hpath g Hg) as [Hpf Hrew]; + revert Hpf Hrew; + changeR2; fold new_frame; fold local_config; fold local_state; fold obs; + intros Hpf Hrew; rewrite Hrew in *. + +Ltac simplify_update := + lazymatch goal with + | Hda : da_assumption ?da, + new_frame := change_frame ?da ?config ?g, + obs := obs_from_config ?local_config ?local_state, + Hpath : connected_path ?config, + Hg : alive (?config (Good ?g)) = true + |- context[Rle_dec _ D] => + simplify_update_at da Hda config Hpath g Hg new_frame local_config local_state obs + | Hda : da_assumption ?da, + new_frame := change_frame ?da ?config ?g, + obs := obs_from_config ?local_config ?local_state, + Hpath : connected_path ?config, + Hg : alive (?config (Good ?g)) = true, + H : context[Rle_dec _ D] |- _ => + simplify_update_at da Hda config Hpath g Hg new_frame local_config local_state obs + end. + +Lemma move_is_to_target : forall obs pt, rbg obs = Move pt -> + pt = choose_new_pos obs (tgt_loc (choose_target obs)). +Proof using . +intros obs pt. cbn. unfold rbg_fun. +now repeat destruct_match; intro Heq; inv Heq. +Qed. + +(** *** Other invariants **) +(* +(** Alive robots have an observable robot in range and no robot too close. + This is actually true after a round but not necessarily for the initial configuration, + so we add it as an invariant. *) +Definition alive_spec (config : configuration) := forall g, + (config (Good g)).(alive) = true + <-> (∃ r, dist (get_location (config (Good g))) (get_location (config r)) <= Dmax + ∧ ((config r).(id) < (config (Good g)).(id))%nat + ∧ (config r).(alive) = true) + ∧ ∀ r, ((config r).(id) < (config (Good g)).(id))%nat -> (config r).(alive) = true -> + dist (get_location (config (Good g))) (get_location (config r)) > D. +*) + +(* NB: as [rbg] only works in a local frame of reference, + we use [should_die] rather than rbg directly. + Another option would be to first change the frame of reference to a local one. *) + +(** The scout is always alive, always has id 0, and always has its light off. *) +Definition scout_inv config := + exists pt, config scout == {| loc := pt; id := 0; alive := true; light := false |}. + +(** Ids are always distinct *) +Definition distinct_id (config : configuration) := + forall r r', r ≠r' -> (config r).(id) ≠(config r').(id). + +(* TODO: can be strenghten for all robots as the scout cannot die + -> Requires alive (config' r) = false instead of should_die *) +(** If a relay robot dies, it has its light on *) +Definition killed_light_on config := forall g, + let obs := obs_from_config config (config (Good g)) in + (* if g dies, *) + (config (Good g)).(alive) = true -> + should_die obs (config (Good g)).(loc) = true -> + (* it has its light on *) + (config (Good g)).(light) = true. + +(** If a relay robot dies, it is because of a robot with its light off. *) +Definition killer_light_off config := forall g, + let obs := obs_from_config config (config (Good g)) in + (* if g dies, *) + (config (Good g)).(alive) = true -> + should_die obs (config (Good g)).(loc) = true -> + (* it is because of r' which has its light off *) + exists tgt, (tgt ∈ obs)%set + ∧ dist (config (Good g)).(loc) (tgt_loc tgt) <= D + ∧ tgt_light tgt = false. + +(** Every relay robot has an observable robot at most Dp away. *) +Definition exists_at_less_than_Dp config := forall g, alive (config (Good g)) = true -> + let obs := obs_from_config config (config (Good g)) in + (* if all robots within range have their light on *) + For_all (fun x => tgt_light x = true) obs -> + (* then one of them is at most Dp away *) + exists tgt, (tgt ∈ obs)%set ∧ dist (tgt_loc tgt) (config (Good g)).(loc) <= Dp. + +(* The specification of [choose_target] and the property [exists_at_less_than_Dp] + entail that if the target has its light on, then it is at most Dp away. *) +Lemma target_light_on_at_most_Dp : forall local_config, + exists_at_less_than_Dp local_config -> + forall g, alive (local_config (Good g)) = true -> + (local_config (Good g)).(loc) == origin -> + let obs := obs_from_config local_config (local_config (Good g)) in + obs =/= {}%set -> tgt_light (choose_target obs) = true -> + dist (local_config (Good g)).(loc) (tgt_loc (choose_target obs)) <= Dp. +Proof using . +intros local_config Hex g Halive Hg obs Hobs Hlight. +assert (Hall := target_light_on Hobs Hlight). +specialize (Hex g Halive). cbn zeta in Hex. fold obs in Hex, Hall. +destruct (Hex Hall) as [tgt [Hin Hrange]]. +destruct (Rle_or_lt (dist (loc (local_config (Good g))) (tgt_loc (choose_target obs))) Dp) + as [Hle | Hlt]; trivial; []. +rewrite Hg in Hlt, Hrange. +assert (Hdist := target_is_close Hobs Hlight Hlt). +fold obs in Hdist. apply Hdist in Hin. rewrite dist_sym in Hin. lra. +Qed. + +Lemma exists_at_less_than_Dp_map : forall (iso : isometry location) config, + exists_at_less_than_Dp config -> + exists_at_less_than_Dp + (map_config (lift (existT precondition iso (iso_precondition iso))) config). +Proof using . +intros iso config Hex g Halive obs' Hall. subst obs'. +unfold map_config in Hall at 2. unfold map_config at 2. +assert (Proper (equiv ==> impl) (For_all (λ x : target, tgt_light x = true))). +{ apply For_all_compat. intros x y Hxy. now rewrite Hxy. } +rewrite <- (obs_from_config_map iso) in Hall. +rewrite For_all_map in Hall; autoclass; try (now intros x y Hxy; rewrite Hxy); []. +specialize (Hex g Halive). cbn zeta in Hex. apply Hex in Hall. +destruct Hall as [tgt [Hin Hdist]]. +exists (lift_target iso tgt). split. ++ rewrite <- obs_from_config_map, map_spec; autoclass. ++ cbn. changeR2. now rewrite dist_prop. +Qed. + +Lemma killed_light_on_map : forall (iso : isometry location) (Hiso : @precondition _ _ St iso), + forall config, killed_light_on config -> + killed_light_on (map_config (lift (existT precondition iso Hiso)) config). +Proof using . +intros iso Hiso config Hkilled r obs Halive Hdie. cbn -[equiv] in Hdie |- *. +apply Hkilled; try apply Halive; []. unfold obs in Hdie. unfold map_config in Hdie at 2. +now rewrite <- obs_from_config_map, should_die_map in Hdie. +Qed. + +Lemma killer_light_off_map : forall (iso : isometry location) (Hiso : @precondition _ _ St iso), + forall config, killer_light_off config -> + killer_light_off (map_config (lift (existT precondition iso Hiso)) config). +Proof using . +intros iso Hiso config Hkiller r obs Halive Hdie. cbn -[equiv] in Hdie |- *. +unfold obs in Hdie. unfold map_config in Hdie at 2. +rewrite <- obs_from_config_map, should_die_map in Hdie. +apply Hkiller in Hdie; try apply Halive; []. destruct Hdie as [tgt [Hin [Hdist Hlight]]]. +exists (lift_target iso tgt). repeat split. ++ unfold obs. unfold map_config at 2. rewrite <- obs_from_config_map, map_spec; autoclass. ++ cbn. now rewrite dist_prop. ++ cbn. assumption. +Qed. + +Lemma scout_empty_obs : forall config, scout_inv config -> + obs_from_config config (config scout) == {}%set. +Proof using . +intros config [pt Hscout] x. rewrite empty_spec, obs_from_config_spec. +split; try tauto; []. intros [r [_ [_ [_ Hid]]]]. +rewrite Hscout in Hid. cbn in Hid. lia. +Qed. + +Lemma should_die_indeed_dies : forall g config, + should_die (obs_from_config config (config (Good g))) (config (Good g)).(loc) = true -> + forall da, da_assumption da -> (round rbg da config (Good g)).(alive) = false. +Proof using . +intros g config Hdie da Hda. simplify_round. cbn. unfold rbg_fun. +cut (should_die obs origin = true); try (now intro Heq; rewrite Heq); []. +unfold obs. rewrite <- (center_prop new_frame). +unfold new_frame. rewrite (proj1 Hda). +rewrite <- (should_die_map new_frame) in Hdie. rewrite <- Hdie. +f_equiv. rewrite (obs_from_config_map _ I). f_equiv. +Qed. + +Lemma dead_means_Good : forall config, scout_inv config -> + forall r, alive (config r) = false -> exists g, r = Good g. +Proof using . +intros config [pt Hscout] [g | b] Hdead; eauto; []. +exfalso. rewrite (byz_is_scout b), Hscout in Hdead. discriminate. +Qed. + + +Section InvariantProofs. + +Variable da : demonic_action. +Hypothesis Hda : da_assumption da. + +Variable config : configuration. +Local Definition config' := round rbg da config. + +Lemma da_iso : isometry_da_prop da. +Proof using Hda. apply Hda. Qed. + +Lemma da_FSYNC : FSYNC_da da. +Proof using Hda. apply Hda. Qed. + +Lemma da_relocate_scout : + dist (da.(relocate_byz) config scout_B).(loc) (config scout).(loc) <= D. +Proof using Hda. apply Hda. Qed. + +Lemma da_scout_alive : (da.(relocate_byz) config scout_B).(alive) = true. +Proof using Hda. apply Hda. Qed. + +Lemma da_scout_light : (da.(relocate_byz) config scout_B).(light) = false. +Proof using Hda. apply Hda. Qed. + +Lemma da_scout_id : (da.(relocate_byz) config scout_B).(id) = 0%nat. +Proof using Hda. apply Hda. Qed. + +(** The scout's state *) +Hypothesis scout_has_inv : scout_inv config. + +Corollary scout_alive : (config scout).(alive) = true. +Proof using scout_has_inv. destruct scout_has_inv as [pt Hpt]. now rewrite Hpt. Qed. + +Corollary scout_light : (config scout).(light) = false. +Proof using scout_has_inv. destruct scout_has_inv as [pt Hpt]. now rewrite Hpt. Qed. + +Corollary scout_id : (config scout).(id) = 0%nat. +Proof using scout_has_inv. destruct scout_has_inv as [pt Hpt]. now rewrite Hpt. Qed. + +(** id should be unique to properly identify a robot *) +Hypothesis id_uniq : distinct_id config. + +Corollary id_unique : forall r r', r ≠r' <-> (config r).(id) ≠(config r').(id). +Proof using id_uniq. intros. split; auto; intuition subst; tauto. Qed. + +Lemma forever_same_id : forall r, (config' r).(id) = (config r).(id). +Proof using Hda scout_has_inv. +intros [g | b]; unfold config'. ++ simplify_round. destruct (rbg obs); cbn; try destruct_match; reflexivity. ++ unfold round. rewrite da_FSYNC. rewrite (byz_is_scout b). fold scout. + now rewrite scout_id, da_scout_id. +Qed. + +Corollary id'_unique : forall r r', r ≠r' <-> (config' r).(id) ≠(config' r').(id). +Proof using Hda id_uniq scout_has_inv. intros. rewrite 2 forever_same_id. apply id_unique. Qed. + +Lemma scout_invariant : scout_inv config -> scout_inv config'. +Proof using Hda. +intros [pt Hscout]. exists (config' scout).(loc). +unfold config', round. +destruct_match; cbn; changeR2; [|]. +- now rewrite da_scout_alive, da_scout_light, da_scout_id. +- now rewrite Hscout. +Qed. + +Lemma distinct_id_invariant : distinct_id config'. +Proof using Hda id_uniq scout_has_inv. unfold distinct_id. apply id'_unique. Qed. + + +(** Why is a robot dead after a round? *) +Lemma why_dead : forall r, + (config' r).(alive) = false -> + (** either it was already dead *) + (config r).(alive) = false + (** or there is an alive robot too close *) + ∨ (exists r', ((config r').(id) < (config r).(id))%nat + ∧ (config r').(alive) = true + ∧ dist (config r).(loc) (config r').(loc) <= D) + (** or there is no alive robot in range *) + ∨ (forall r', ((config r').(id) < (config r).(id))%nat -> + (config r').(alive) = true -> + dist (config r').(loc) (config r).(loc) > Dmax). +Proof using Hda scout_has_inv. +intros r Hdead. +assert (Hg : exists g, r = Good g). +{ destruct r as [g | b]; eauto; []. exfalso. revert Hdead. rewrite (byz_is_scout b). + unfold config', round. destruct_match. + - rewrite da_scout_alive. discriminate. + - cbn. change (Byz scout_B) with scout. rewrite scout_alive. discriminate. } +destruct Hg as [g ?]. subst r. +unfold config' in Hdead. simplify_round. +destruct (rbg obs) eqn:Hrbg. ++ assert (Hdie : should_die obs origin = true). + { revert Hrbg. cbn. unfold rbg_fun. changeR2. repeat destruct_match; trivial; discriminate. } + rewrite should_die_true in Hdie. destruct Hdie as [Hdie | Hdie]. + - do 2 right. intros r Hid Halive. + assert (Hobs := obs_from_config_spec local_config local_state). + unfold obs in Hdie. setoid_rewrite Hdie in Hobs. setoid_rewrite empty_spec in Hobs. + specialize (Hobs (mk_tgt (local_config r))). cbn -[dist equiv] in Hobs. + setoid_rewrite dist_prop in Hobs. + apply Rnot_le_gt. intro. rewrite Hobs. exists r. intuition. + - right; left. destruct Hdie as [tgt [Hin Hdist]]. unfold obs in Hin. + rewrite obs_from_config_spec in Hin. destruct Hin as [r [Heq [Halive [Hle Hid]]]]. + exists r. repeat split; trivial; []. + rewrite Heq in Hdist. cbn -[dist] in Hdist. + rewrite <- (center_prop (change_frame da config g)), dist_prop in Hdist. + rewrite da_iso, dist_sym in Hdist. apply Hdist. ++ left. cbn in Hdead. assumption. ++ left. cbn in Hdead. revert Hdead. destruct_match; auto. +Qed. + +Corollary dies_as_it_should : forall r, + (config r).(alive) = true -> + (config' r).(alive) = false -> + should_die (obs_from_config config (config r)) (config r).(loc) = true. +Proof using Dmax_6D D_pos config da Hda scout_has_inv. +intros r Halive Halive'. rewrite should_die_true. +destruct (why_dead r Halive') + as [Habs | [[r' [Hid [Halive_r' Hdist]]] | Hall]]; try congruence; [|]. ++ right. exists (mk_tgt (config r')). split. + - rewrite obs_from_config_spec. exists r'. repeat split; trivial; []. + rewrite dist_sym. lra. + - now rewrite dist_sym. ++ left. intro x. rewrite empty_spec. split; try tauto; []. + rewrite obs_from_config_spec. intros [r' [Heq [Halive_r' [Hdist Hr']]]]. + specialize (Hall r' Hr' Halive_r'). lra. +Qed. + +Corollary dies_iff_should_die : forall g, alive (config (Good g)) = true -> + alive (config' (Good g)) = false + <-> should_die (obs_from_config config (config (Good g))) (config (Good g)).(loc) = true. +Proof using D_pos Dmax_6D Hda scout_has_inv. +intros g Halive. split; intro Hdead. +- now apply dies_as_it_should. +- now apply should_die_indeed_dies. +Qed. + +Lemma find_killer : + killed_light_on config -> + killer_light_off config -> + forall r, alive (config r) = true -> alive (config' r) = false -> + exists r', (id (config r') < id (config r))%nat + ∧ alive (config' r') = true + ∧ light (config r') = false + ∧ dist (loc (config r)) (loc (config r')) <= D. +Proof using D_pos Dmax_6D Hda scout_has_inv. +intros Hkilled Hkiller r Halive_r Hdead_r. +destruct (dead_means_Good (scout_invariant scout_has_inv) _ Hdead_r) as [g ?]. subst r. +apply dies_as_it_should, Hkiller in Hdead_r; trivial; []. +destruct Hdead_r as [tgt [Hin [Hdist Hlight]]]. +rewrite obs_from_config_spec in Hin. destruct Hin as [r' [Htgt [Halive [Hdist' Hid]]]]. +rewrite Htgt in *. cbn in Hlight. +exists r'. repeat split; trivial; []. +rewrite <- Bool.not_false_iff_true. intro Habs. +destruct (dead_means_Good (scout_invariant scout_has_inv) _ Habs) as [g' ?]. subst r'. +apply dies_as_it_should, Hkilled in Habs; trivial; []. congruence. +Qed. + +Lemma moving_means_light_off : forall r, + (config' r).(alive) = true -> + (config' r).(loc) =/= (config r).(loc) -> + (config' r).(light) = false. +Proof using Hda scout_has_inv. +intros [g | b] Halive Hmoving. +* unfold config' in *. simplify_round. + destruct (rbg obs). + + cbn in Halive. discriminate. + + cbn -[equiv] in Hmoving. contradiction Hmoving. apply retraction_section. + + cbn -[equiv] in Hmoving |- *. destruct_match. + - reflexivity. + - contradiction Hmoving. cbn. now rewrite simpl_inverse_l. +* rewrite (byz_is_scout b). apply scout_invariant in scout_has_inv. + destruct scout_has_inv as [pt Hpt]. now rewrite Hpt. +Qed. + +Corollary light_on_same_location : forall r, alive (config' r ) = true -> + light (config' r) = true -> loc (config' r) == loc (config r). +Proof using Hda scout_has_inv. +intros r Halive Hlight. +destruct (equiv_dec (loc (config' r)) (loc (config r))); trivial; []. +exfalso. revert Hlight. rewrite moving_means_light_off; auto. +Qed. + +Lemma alive_before : forall r, + (config' r).(alive) = true -> (config r).(alive) = true. +Proof using Hda scout_has_inv. +intros [g | b]; unfold config'. ++ simplify_round. destruct (rbg obs). + - cbn. discriminate. + - cbn. auto. + - cbn. destruct_match; auto. ++ intros _. rewrite (byz_is_scout b). apply scout_alive. +Qed. + +(** Because of the update function, robots can move at most D in a round. *) +Lemma move_at_most_D : forall r, + dist (get_location (config' r)) (get_location (config r)) <= D. +Proof using D_pos Hda. +intros [g | b]; unfold config'. +* simplify_round. destruct (rbg obs); cbn -[dist]. + + rewrite simpl_inverse_l, dist_same. now apply Rlt_le. + + rewrite simpl_inverse_l, dist_same. now apply Rlt_le. + + destruct_match. + - cbn -[dist]. rewrite <- da_iso. + assert (Hcenter := center_prop new_frame). + rewrite Inversion in Hcenter. fold new_frame. rewrite <- Hcenter. + changeR2. change (retraction new_frame) with (section (new_frame â»Â¹)). + rewrite dist_prop. assumption. + - cbn -[dist]. rewrite simpl_inverse_l, dist_same. now apply Rlt_le. +* unfold round. rewrite da_FSYNC. cbn -[dist]. + rewrite (byz_is_scout b). apply da_relocate_scout. +Qed. + +Lemma wlog_lt_id : forall P, + (forall r r', P r r' <-> P r' r) -> + (forall r r', ((config r).(id) < (config r').(id))%nat -> P r r') -> + forall r r', r <> r' -> P r r'. +Proof using id_uniq. +intros P Hsym HP r r' Hdiff. +rewrite id_unique, Nat.lt_gt_cases in Hdiff. +destruct Hdiff. +- now apply HP. +- now apply Hsym, HP. +Qed. + +Lemma no_collision_invariant : connected_path config -> + no_collision config -> no_collision config'. +Proof using D_pos Dmax_6D Hda id_uniq scout_has_inv. +intros Hpath Hcollision r r' Hdiff. +pattern r, r'. +match goal with |- ?P _ _ => assert (Hsym : forall r r', P r r' ↔ P r' r) by (intuition auto with *) end. +cbn beta. apply (wlog_lt_id _ Hsym); trivial; []. clear r r' Hdiff Hsym. +intros r r' Hlt Halive Halive' Heq'. +(* The robot with highest id cannot be the scout. *) +assert (Hr' : exists g, r' = Good g). +{ destruct r' as [g | b]; eauto; []. + rewrite (byz_is_scout b) in *. fold scout in Hlt. rewrite scout_id in Hlt. lia. } +destruct Hr' as [g ?]. subst r'. +(* To collide, both robots must have been at most 2*D away. *) +assert (Hg_alive : alive (config (Good g)) = true) by apply alive_before, Halive'. +assert (Hdist : dist (get_location (config r)) + (get_location (config (Good g))) <= 2 * D). +{ etransitivity; [apply RealMetricSpace.triang_ineq |]. + replace (2 * D) with (D + D) by lra. apply Rplus_le_compat. + + rewrite dist_sym. apply move_at_most_D. + + rewrite Heq'. apply move_at_most_D. } +(* To collide, r was at most D away from the new position of g *) +assert (Hnext : dist (get_location (config r)) (get_location (config' (Good g))) <= D). +{ rewrite <- Heq', dist_sym. apply move_at_most_D. } +(* Thus, g cannot have moved *) +assert (Hg : get_location (config' (Good g)) == get_location (config (Good g))). +{ unfold config' in Halive', Heq' |- *. simplify_round. destruct (rbg obs) eqn:Hrbg. + + cbn in Halive'. discriminate. + + cbn -[equiv]. apply simpl_inverse_l. + + assert (Haction := Hrbg). unfold rbg, rbg_fun in Haction. cbn -[equiv] in Haction. + destruct (should_die obs) eqn:Hdie; try discriminate; []. + destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs)))) eqn:Hmove; + try discriminate; []. inv Haction. + rewrite move_to_true in Hmove. + assert (Hobs : (mk_tgt (local_config r) ∈ obs)%set). + { unfold obs. rewrite obs_from_config_spec. exists r. repeat split. + - cbn. now apply alive_before. + - cbn. rewrite dist_prop. transitivity (2*D); trivial; []. lra. + - cbn. apply Hlt. } + specialize (Hmove _ Hobs). simpl tgt_loc in Hmove. + assert (Hloc' : get_location (config' (Good g)) + == (new_frame â»Â¹) (choose_new_pos obs (tgt_loc (choose_target obs)))). + { unfold config'. rewrite Hround. cbn -[equiv rbg]. simplify_update. reflexivity. } + rewrite <- dist_swap, <- Hloc', dist_sym in Hmove. + rewrite <- Hround in Hnext. fold config' in Hnext. change loc with get_location in *. + exfalso. clear -Hmove Hnext D_pos. lra. } +(* Then r was actually at most D away from g *) +rewrite Hg in Hnext. +(* In that case, g should have died. *) +unfold config' in Halive'. simplify_round. +assert (Hdie : should_die obs origin = true). +{ rewrite should_die_true. right. exists (mk_tgt (local_config r)). split. + + unfold obs. rewrite obs_from_config_spec. exists r. repeat split. + - cbn. now apply alive_before. + - cbn. rewrite dist_prop. transitivity D; trivial; []. lra. + - cbn. assumption. + + cbn. rewrite <- dist_swap, <- (center_prop new_frame), + (compose_inverse_l new_frame (center new_frame)). + unfold new_frame. rewrite da_iso, dist_sym. apply Hnext. } +cbn -[update] in Halive'. unfold rbg_fun in Halive'. +rewrite Hdie in Halive'. cbn in Halive'. discriminate. +Qed. + +Lemma connected_path_invariant : + connected_path config -> + killed_light_on config -> + killer_light_off config -> + exists_at_less_than_Dp config -> + connected_path config'. +Proof using D_pos Dmax_6D Hda scout_has_inv. +intros Hpath Hkilled Hkiller Hless g Halive'. +assert (Halive : alive (config (Good g)) = true) by now apply alive_before. +setoid_rewrite forever_same_id. unfold config'. simplify_round. setoid_rewrite Hround. +pose (real_target := new_frame â»Â¹ (choose_new_pos obs (tgt_loc (choose_target obs)))). +assert (Hlocal_state : loc local_state == origin). +{ unfold local_state, new_frame. cbn -[equiv]. rewrite <- da_iso. now rewrite center_prop. } +assert (Hobs : obs =/= {}%set). +{ rewrite connected_path_iff_obs_non_empty in Hpath. + specialize (Hpath g Halive). unfold complement in *. cbn -[equiv] in Hpath. + rewrite <- (map_empty_iff_empty (lift_target new_frame)) in Hpath; autoclass; []. + rewrite (obs_from_config_map _ I) in Hpath. apply Hpath. } +cbn -[equiv update location] in *. unfold rbg_fun in *. +destruct (should_die obs origin) eqn:Hdie; +[| destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs)))) eqn:Hmove]. +* cbn in Halive'. discriminate. +* cbn -[equiv location] in *. simplify_update. simpl loc in *. fold real_target. + (* Let r be the target of g in config *) + assert (Hrange := target_in_range Hobs). unfold obs in Hrange. + rewrite obs_from_config_spec in Hrange. fold obs in Hrange. + destruct Hrange as [r [Hr [Halive_r [Hdist_r Hid_r]]]]. + changeR2. fold config'. + destruct (alive (config' r)) eqn:Halive'_r. + + (* If r is still alive on config', we can use it *) + exists r. repeat split; trivial; []. + etransitivity; [apply (RealMetricSpace.triang_ineq _ (loc (config r))) |]. + replace Dmax with (D + Dp) by (unfold Dp; lra). + apply Rplus_le_compat; try apply move_at_most_D; []. + unfold real_target. rewrite dist_sym, <- dist_swap. + unfold local_state in Hdist_r. rewrite Hlocal_state in Hdist_r. + destruct (choose_new_pos_spec obs (new_frame (loc (config r))) Hdist_r) as [Hdist1 Hdist2]. + rewrite dist_sym, Hr. apply Hdist1. + + (* If r dies, we can use its killer, because it did not move *) + assert (Hlight_r : tgt_light (choose_target obs) = true). + { destruct (dead_means_Good (scout_invariant scout_has_inv) _ Halive'_r) as [g' ?]. subst r. + fold obs. rewrite Hr. cbn. now apply Hkilled, dies_as_it_should. } + destruct (find_killer Hkilled Hkiller r Halive_r Halive'_r) + as [r' [Hid_r' [Halive_r' [Hlight_r' Hdist_r']]]]. + exfalso. revert Hlight_r'. rewrite (Bool.not_false_iff_true (light (config r'))). + change (tgt_light (mk_tgt (local_config r')) = true). + eapply (target_light_on Hobs Hlight_r). unfold obs. + rewrite obs_from_config_spec. exists r'. repeat split. + - now apply alive_before. + - etransitivity; [apply (RealMetricSpace.triang_ineq _ (loc (local_config r))) |]. + replace Dmax with (D + Dp) by (unfold Dp; lra). apply Rplus_le_compat. + -- cbn. now rewrite dist_prop, dist_sym. + -- assert (Hless' : exists_at_less_than_Dp local_config). + { unfold local_config. now apply exists_at_less_than_Dp_map. } + apply tgt_loc_compat in Hr. rewrite <- Hr, dist_sym. revert Hobs Hlight_r. + now apply (target_light_on_at_most_Dp Hless'), Hlocal_state. + - cbn in *. lia. +* (* If g cannot move, it is because of a robot r too close to the target of g *) + assert (Hmove2 := Hmove). rewrite move_to_false in Hmove2. + destruct Hmove2 as [tgt [Hin_r Hdist_r]]. + unfold obs in Hin_r. rewrite obs_from_config_spec in Hin_r. + destruct Hin_r as [r [Hr [Halive_r [Hrange_r Hid_r]]]]. + destruct (alive (config' r)) eqn:Hr_alive'. + + (* As r is alive, we can use it *) + exists r. repeat split; trivial; []. + cbn. changeR2. rewrite simpl_inverse_l. + transitivity (D + 2 * D + D); try lra; []. fold (config' r). + transitivity ( + dist (loc (config' r)) (loc (config r)) + + dist (loc (config r)) (new_frame â»Â¹ (choose_new_pos obs (tgt_loc (choose_target obs)))) + + dist (new_frame â»Â¹ (choose_new_pos obs (tgt_loc (choose_target obs)))) + (loc (config (Good g)))); + [ etransitivity; [eapply RealMetricSpace.triang_ineq | + apply Rplus_le_compat; [eapply RealMetricSpace.triang_ineq |]]; reflexivity + | repeat apply Rplus_le_compat]. + - apply move_at_most_D. + - rewrite Hr in Hdist_r. simpl tgt_loc in Hdist_r. + now rewrite <- dist_swap, dist_sym in Hdist_r. + - assert (Hdist : dist (tgt_loc (choose_target obs)) 0 <= Dmax). + { rewrite <- Hlocal_state. + change (new_frame (loc (config (Good g)))) with (loc (local_config (Good g))). + now apply (in_range_dist local_config), target_in_range. } + destruct (choose_new_pos_spec obs _ Hdist) as [Hdist1 Hdist2]. + rewrite dist_swap, dist_sym, <- da_iso. unfold new_frame. + rewrite center_prop. fold obs in Hdist2. apply Hdist2. + + (* If r dies, its killer is alive and in range of g *) + destruct (find_killer Hkilled Hkiller _ Halive_r Hr_alive') + as [r' [Hid_r' [Halive_r' [Hlight_r' Hdist_r']]]]. + exists r'. cbn in Hid_r. repeat split; trivial; try lia; []. + cbn. rewrite simpl_inverse_l. changeR2. + transitivity (D + D + 2 * D + D); try lra; []. fold (config' r'). + transitivity (dist (loc (config' r')) (loc (config r')) + + dist (loc (config r')) (loc (config r)) + + dist (loc (config r)) real_target + + dist real_target (loc (config (Good g)))); + [ repeat (etransitivity; [eapply RealMetricSpace.triang_ineq + | apply Rplus_le_compat; [| reflexivity]]); reflexivity + | repeat apply Rplus_le_compat]. + - apply move_at_most_D. + - now rewrite dist_sym. + - unfold real_target. rewrite dist_sym, <- dist_swap. now rewrite Hr in Hdist_r. + - assert (Hdist : dist (tgt_loc (choose_target obs)) 0 <= Dmax). + { rewrite <- Hlocal_state. + change (new_frame (loc (config (Good g)))) with (loc (local_config (Good g))). + now apply (in_range_dist local_config), target_in_range. } + destruct (choose_new_pos_spec obs _ Hdist) as [Hdist1 Hdist2]. + unfold real_target. rewrite dist_swap, dist_sym, <- da_iso. unfold new_frame. + rewrite center_prop. fold obs in Hdist2. apply Hdist2. +Qed. + + +Lemma killed_light_on_invariant : + killed_light_on config -> + killer_light_off config -> + connected_path config -> + exists_at_less_than_Dp config -> + killed_light_on config'. +Proof using D_pos Dmax_6D Hda scout_has_inv. +intros Hkilled Hkiller Hpath HlessDp g obs' Halive' Hdie'. +assert (Hg_alive : alive (config (Good g)) = true) by apply alive_before, Halive'. +assert (Hshould_die := Hdie'). rewrite should_die_true in Hshould_die. +destruct Hshould_die as [Hobs' | Hobs']. +* (* Absurd case: there is no robot in range now *) + (* Since g is alive in config', it had a target r in config *) + unfold config'. simplify_round. cbn -[update equiv location] in *. unfold rbg_fun in *. + destruct (should_die obs 0%VS) eqn:Hdie; + [| destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs)))) eqn:Hmove]. + + cbn in Hdie'. discriminate. + + (* If g moved, its target r in config is still in range in config' *) + cbn -[equiv] in *. simplify_update. cbn -[equiv location] in *. exfalso. + assert (Hlocal_state : get_location local_state == origin). + { unfold local_state, new_frame. now rewrite <- da_iso, center_prop. } + assert (Hrange : dist (tgt_loc (choose_target obs)) 0 <= Dmax). + { rewrite <- (center_prop new_frame). unfold new_frame at 2. rewrite da_iso. + change (new_frame (get_location (config (Good g)))) with ((local_config (Good g)).(loc)). + apply (in_range_dist local_config). apply target_in_range; trivial; []. + change local_config with (map_config (lift (existT precondition new_frame I)) config). + unfold map_config at 2. unfold complement. + rewrite <- obs_from_config_map, map_empty_iff_empty; autoclass; []. + rewrite connected_path_iff_obs_non_empty in Hpath. now apply Hpath. } + destruct (choose_new_pos_spec obs _ Hrange) as [Hdist1 Hdist2]. fold obs in Hdist1, Hdist2. + rewrite <- (dist_prop (new_frame â»Â¹)), <- (loc_compat Hround) in Hdist1. + assert (Hstate_eq_0 : local_state == {{local_state with loc := origin}}). + { unfold local_state. rewrite <- da_iso. destruct (config (Good g)). split; [| repeat split]. + cbn -[equiv]. unfold new_frame. now rewrite center_prop. } + assert (Hobs : obs =/= {}%set). + { rewrite connected_path_iff_obs_non_empty in Hpath. specialize (Hpath g Halive'). + intro Habs. apply Hpath. cbn -[equiv]. + rewrite <- (map_empty_iff_empty (lift_target new_frame)); autoclass; []. + rewrite (obs_from_config_map _ I). apply Habs. } + assert (Htgt := target_in_range Hobs). + unfold obs in Htgt. rewrite obs_from_config_spec in Htgt. fold obs in Htgt. + destruct Htgt as [r [Hr [Halive_r [Hdist_r Hid_r]]]]. + (* is r still alive in config'? *) + destruct (alive (config' r)) eqn:Halive'_r. + - (* if so, r is the robot we are looking for *) + specialize (Hobs' (mk_tgt (config' r))). + rewrite empty_spec in Hobs'. rewrite <- Hobs'. unfold obs'. + rewrite obs_from_config_spec. exists r. split; [| repeat split]. + ++ reflexivity. + ++ assumption. + ++ replace Dmax with (D + Dp) by (unfold Dp; lra). + etransitivity; [apply (RealMetricSpace.triang_ineq _ (loc (config r))) |]. + apply Rplus_le_compat. + -- apply move_at_most_D. + -- rewrite Hr in Hdist1. cbn in Hdist1. rewrite simpl_inverse_l in Hdist1. + rewrite dist_sym. apply Hdist1. + ++ now rewrite 2 forever_same_id. + - (* otherwise, it was killed by another robot which must have been in range of g, + thus had its light on and could not kill r *) + destruct (find_killer Hkilled Hkiller r Halive_r Halive'_r) + as [r' [Hid_r' [Halive_r' [Hlight_r' Hdist_r']]]]. + destruct (dead_means_Good (scout_invariant scout_has_inv) r Halive'_r) as [g' ?]. subst r. + rewrite dies_iff_should_die in Halive'_r; auto; []. + destruct (Hkiller _ Halive_r Halive'_r) as [tgt [Hin [Hdist Hlight]]]. + revert Hlight. change (tgt_light tgt <> false). rewrite Bool.not_false_iff_true. + rewrite obs_from_config_spec in Hin. destruct Hin as [r [Htgt [Halive [Hd Hid]]]]. + apply Hkilled in Halive'_r; auto; []. + change (tgt_light (mk_tgt (config (Good g'))) = true) in Halive'_r. + assert (Hlight_on := target_light_on Hobs). + fold obs in Hlight_on. rewrite Hr in Hlight_on. + specialize (Hlight_on Halive'_r (mk_tgt (local_config r))). + rewrite Htgt. apply Hlight_on. unfold obs. + rewrite obs_from_config_spec. exists r. split; reflexivity || (repeat split). + ++ apply Halive. + ++ replace Dmax with (D + Dp) by (unfold Dp; lra). + etransitivity; [ apply (RealMetricSpace.triang_ineq _ (loc (local_config (Good g')))) |]. + apply Rplus_le_compat. + -- rewrite Htgt in Hdist. cbn. changeR2. now rewrite dist_prop, dist_sym. + -- apply (exists_at_less_than_Dp_map new_frame) in HlessDp. + assert (Hr_light := tgt_light_compat Hr). cbn in Hr_light, Halive'_r. + rewrite <- Hr_light in Halive'_r. + assert (Hless := target_light_on_at_most_Dp + HlessDp g Hg_alive Hlocal_state Hobs Halive'_r). + change (dist (loc (local_state)) (tgt_loc (choose_target obs)) <= Dp) in Hless. + rewrite Hr, dist_sym in Hless. apply Hless. + ++ now transitivity (id (config (Good g'))). + + reflexivity. +* (* If a robot r is close enough to kill g in config' *) + destruct Hobs' as [tgt [Hin' Hdist']]. + unfold config'. simplify_round. + cbn -[equiv update] in *. unfold rbg_fun in *. + destruct (should_die obs 0%VS) eqn:Hdie; + [| destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs)))) eqn:Hmove]. + + (* g cannot die in config' as it is assumed alive *) + cbn in Halive'. discriminate. + + (* Assume by contradiction that g moves so has its light off *) + cbn -[equiv] in *. simplify_update. exfalso. + simpl loc in Hdist', Hdie'. simpl alive in Halive'. + unfold obs' in Hin'. rewrite obs_from_config_spec in Hin'. + destruct Hin' as [r [Htgt [Halive [Hdist Hid]]]]. + rewrite move_to_true in Hmove. specialize (Hmove (lift_target new_frame (mk_tgt (config r)))). + cbn -[equiv] in Hround. + assert (Heq' : loc (config' (Good g)) + == (new_frame â»Â¹) (choose_new_pos obs (tgt_loc (choose_target obs)))) + by now rewrite Hround. + (* r was already in range of g in config *) + assert (Hin : (lift_target new_frame (mk_tgt (config r)) ∈ obs)%set). + { change (lift_target new_frame (mk_tgt (config r)) ∈ + obs_from_config (map_config (lift (existT precondition new_frame I)) config) + (lift (existT precondition new_frame I) (config (Good g))))%set. + rewrite <- obs_from_config_map, map_spec; autoclass; []. + exists (mk_tgt (config r)). split; try reflexivity; []. + rewrite obs_from_config_spec. exists r. split; [| repeat split]. + + reflexivity. + + now apply alive_before. + + transitivity (D + D + D); try lra; []. + transitivity (dist (loc (config r)) (loc (config' r)) + + dist (loc (config' r)) (loc (config' (Good g))) + + dist (loc (config' (Good g))) (loc (config (Good g)))); + [ etransitivity; [eapply RealMetricSpace.triang_ineq + | apply Rplus_le_compat; [eapply RealMetricSpace.triang_ineq |]] + | repeat apply Rplus_le_compat]; try reflexivity; [| |]. + - rewrite dist_sym. apply move_at_most_D. + - rewrite Htgt, <- Heq' in Hdist'. apply Hdist'. + - apply move_at_most_D. + + now setoid_rewrite <- forever_same_id. } + (* But r in config was not within 2 D of the location of g in config' since g moved, + so it could not be within range D of g in config' and kill it *) + apply Hmove in Hin. contradict Hin. apply Rle_not_gt. + replace (2 * D) with (D + D) by lra. + etransitivity; [apply (RealMetricSpace.triang_ineq _ (new_frame (tgt_loc tgt))) |]. + apply Rplus_le_compat. + - cbn. changeR2. rewrite dist_prop, Htgt, dist_sym. cbn. apply move_at_most_D. + - now rewrite <- dist_swap, dist_sym. + + cbn. reflexivity. +Qed. + +Lemma killer_light_off_invariant : + killed_light_on config -> + killer_light_off config -> + connected_path config -> + exists_at_less_than_Dp config -> + killer_light_off config'. +Proof using id_uniq D_pos Dmax_6D Hda scout_has_inv. +intros Hkilled Hkiller Hpath HlessDp g obs' Halive' Hdie'. +assert (Halive : alive (config (Good g)) = true) by now apply alive_before. +(* If g dies it has its light on *) +assert (Hkilled' : killed_light_on config') by now apply killed_light_on_invariant. +specialize (Hkilled' g Halive' Hdie'). +(* Thus g did not move *) +assert (Heq_loc : (loc (config' (Good g))) == (loc (config (Good g)))) + by now apply light_on_same_location. +assert (Hobs' : obs' =/= {}%set). +{ apply connected_path_invariant in Hpath; trivial; []. + rewrite connected_path_iff_obs_non_empty in Hpath. + now apply Hpath in Halive'. } +(* Let r be the robot killing g. *) +rewrite should_die_true in Hdie'. +destruct Hdie' as [Habs | [tgt [Hin Hdist]]]; try contradiction; []. +assert (Hr := Hin). unfold obs' in Hr. rewrite obs_from_config_spec in Hr. +destruct Hr as [r [Hr [Halive_r [Hrange_r Hid_r]]]]. +(* Assume by contradiction that the light of r is on *) +rewrite dist_sym in Hdist. +destruct (tgt_light tgt) eqn:Hlight; eauto; []. +exfalso. +(* Then r did not move *) +assert (Heq_loc_r : (loc (config' r)) == (loc (config r))). +{ rewrite Hr in Hlight. now apply light_on_same_location. } +(* So g should have died in config *) +revert Halive'. rewrite (Bool.not_true_iff_false (alive (config' (Good g)))). +rewrite dies_iff_should_die, should_die_true; trivial; []. +right. exists (tgt_loc tgt, light (config r)). split. ++ rewrite obs_from_config_spec. exists r. + repeat split. + - cbn -[equiv]. now rewrite Hr, Heq_loc_r. + - now apply alive_before. + - now rewrite Heq_loc, Heq_loc_r in Hrange_r. + - now rewrite 2 forever_same_id in Hid_r. ++ now rewrite dist_sym, <- Heq_loc. +Qed. + +Lemma exists_at_less_than_Dp_invariant : + killed_light_on config -> + killer_light_off config -> + connected_path config -> + exists_at_less_than_Dp config -> + exists_at_less_than_Dp config'. +Proof using D_pos Dmax_6D Hda scout_has_inv. +intros Hkilled Hkiller Hpath Hless g Halive' obs' Hall. +assert (Halive := alive_before _ Halive'). +revert Halive'. unfold config'. simplify_round. fold config'. intro Halive'. +pose (real_target := new_frame â»Â¹ (choose_new_pos obs (tgt_loc (choose_target obs)))). +assert (Hobs : obs =/= {}%set). +{ rewrite connected_path_iff_obs_non_empty in Hpath. + specialize (Hpath _ Halive). unfold complement in Hpath. cbn -[equiv] in Hpath. + rewrite <- (map_empty_iff_empty (lift_target new_frame)), + (obs_from_config_map _ I) in Hpath; autoclass. } +assert (Hlocal_less : exists_at_less_than_Dp local_config) + by now apply exists_at_less_than_Dp_map. +assert (Hlocal_state : loc (local_config (Good g)) == origin). +{ cbn -[equiv]. unfold new_frame. now rewrite <- da_iso, center_prop. } +destruct (light (config' (Good g))) eqn:Hlight'. +* (* If g has its light on in config', it is because of a robot r close to its target location + in config, thus at most 3D away from the current location of g. *) + unfold config' in Hlight'. rewrite Hround in Hlight'. + assert (Hmove : move_to obs (choose_new_pos obs (tgt_loc (choose_target obs))) = false). + { cbn -[equiv] in Halive', Hround, Hlight' |-*. unfold rbg_fun in *. + destruct (should_die obs 0%VS) eqn:Hdie; + [| destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs))))]. + + cbn in Halive'. discriminate. + + simplify_update. cbn in Hlight'. discriminate. + + reflexivity. } + rewrite move_to_false in Hmove. destruct Hmove as [tgt [Hin Hdist]]. unfold obs in Hin. + rewrite obs_from_config_spec in Hin. destruct Hin as [r [Htgt [Halive_r [Hdist_r Hid_r]]]]. + (* The main result depending on distances *) + assert (Hdist_main : dist (loc (config r)) (loc (config' (Good g))) <= 3 * D). + { replace (3 * D) with (2 * D + D) by lra. + etransitivity; [eapply (RealMetricSpace.triang_ineq _ real_target) | apply Rplus_le_compat]. + - unfold real_target. rewrite dist_sym, <- dist_swap. + rewrite Htgt in Hdist. apply Hdist. + - assert (Hrange : dist (tgt_loc (choose_target obs)) 0 <= Dmax). + { rewrite <- Hlocal_state. now apply (in_range_dist local_config), target_in_range. } + destruct (choose_new_pos_spec obs _ Hrange) as [Hdist1 Hdist2]. + fold obs in Hdist1, Hdist2. unfold real_target. + assert (Heq_loc : loc (config' (Good g)) == loc (config (Good g))). + { now apply light_on_same_location; rewrite Hround. } + rewrite dist_swap, dist_sym, Heq_loc. + change (new_frame (loc (config (Good g)))) with (loc (local_config (Good g))). + rewrite Hlocal_state. apply Hdist2. } + assert (Hprop : forall A B : Prop, (B -> A) -> B -> A ∧ B) by tauto. + destruct (alive (config' r)) eqn:Halive'_r. + + (* If r is still alive in config', we can use it *) + exists (mk_tgt (config' r)). simpl tgt_loc. unfold obs'. apply Hprop. + - intro Hdist'. rewrite obs_from_config_spec. exists r. split; [| repeat split]. + -- reflexivity. + -- assumption. + -- transitivity Dp; trivial; unfold Dp; lra. + -- rewrite 2 forever_same_id. apply Hid_r. + - transitivity (D + 3 * D); try (unfold Dp; lra); []. + etransitivity; [ eapply (RealMetricSpace.triang_ineq _ (loc (config r))) + | apply Rplus_le_compat]; trivial; apply move_at_most_D. + + (* r dies in config' so we use its killer *) + destruct (find_killer Hkilled Hkiller r Halive_r Halive'_r) + as [r' [Hid_r' [Halive_r' [Hlight_r' Hdist_r']]]]. + exists (mk_tgt (config' r')). apply Hprop. + - unfold obs'. rewrite obs_from_config_spec. + exists r'. split; [| repeat split]. + -- reflexivity. + -- assumption. + -- transitivity Dp; trivial; unfold Dp; lra. + -- rewrite 2 forever_same_id. cbn in *; lia. + - transitivity (D + D + 3 * D); try (unfold Dp; lra); []. + transitivity (dist (loc (config' r')) (loc (config r')) + + dist (loc (config r')) (loc (config r)) + + dist (loc (config r)) (loc (config' (Good g)))); + [ etransitivity; [eapply RealMetricSpace.triang_ineq + | apply Rplus_le_compat; [eapply RealMetricSpace.triang_ineq |]] + | repeat apply Rplus_le_compat]; try reflexivity; [| |]. + -- apply move_at_most_D. + -- now rewrite dist_sym. + -- apply Hdist_main. +* (* If g has its light off, we consider its target r in config. *) + (* The target r of g in config has its light on in config' and is at most Dp away in config'. *) + assert (Hrange := target_in_range Hobs). unfold obs in Hrange. + rewrite obs_from_config_spec in Hrange. fold obs in Hrange. + destruct Hrange as [r [Hr [Halive_r [Hdist_r Hid_r]]]]. + (* If r dies in config', it killer would have its light off and be in range of g. + Being in range of g, it should have had its light on, a contradiction *) + assert (Halive'_r : alive (config' r) = true). + { rewrite <- Bool.not_false_iff_true. intro Habs. + destruct (find_killer Hkilled Hkiller r Halive_r Habs) + as [r' [Hid_r' [Halive_r' [Hlight_r' Hdist_r']]]]. + assert (Hlight : tgt_light (choose_target obs) = true). + { rewrite Hr. cbn. + destruct (dead_means_Good (scout_invariant scout_has_inv) r Habs) as [g' ?]. subst r. + assert (Hdie := dies_as_it_should _ Halive_r Habs). + now apply Hkilled in Hdie. } + assert (Htarget := target_light_on_at_most_Dp Hlocal_less g Halive Hlocal_state Hobs Hlight). + change (local_config (Good g)) with local_state in Htarget. fold obs in Htarget. + assert (Hrange_r' : (mk_tgt (local_config r') ∈ obs)%set). + { unfold obs. rewrite obs_from_config_spec. exists r'. split; [| repeat split]. + + reflexivity. + + now apply alive_before. + + rewrite dist_sym. + etransitivity; [apply (RealMetricSpace.triang_ineq _ (loc (local_config r))) |]. + replace Dmax with (Dp + D) by (unfold Dp; lra). apply Rplus_le_compat. + - rewrite Hr in Htarget. apply Htarget. + - cbn. rewrite dist_prop. apply Hdist_r'. + + cbn in *; lia. } + apply (target_light_on Hobs Hlight) in Hrange_r'. + cbn in Hrange_r'. congruence. } + (* Since g is alive and has its light off in config', it moved *) + assert (Heq_loc : loc (config' (Good g)) == real_target). + { unfold config' in Hlight' |- *. rewrite Hround in Hlight' |- *. cbn -[equiv] in *. + unfold rbg_fun in *. + destruct (should_die obs 0%VS); + [| destruct (move_to obs (choose_new_pos obs (tgt_loc (choose_target obs))))]. + - cbn in Halive'. discriminate. + - simplify_update. reflexivity. + - cbn in Hlight'. discriminate. } + (* The distance to this new location and the location of r in config is at most Dp *) + assert (Hdist_main : dist (tgt_loc (choose_target obs)) + (choose_new_pos obs (tgt_loc (choose_target obs))) <= Dp). + { rewrite dist_sym. apply choose_new_pos_spec. rewrite Hr, <- Hlocal_state. apply Hdist_r. } + (* Thus, r is the robot we are looking for *) + assert (Hprop : forall A B : Prop, A -> (A -> B) -> A ∧ B) by tauto. + exists (mk_tgt (config' r)). simpl tgt_loc. apply Hprop. + + unfold obs'. rewrite obs_from_config_spec. + exists r. split; [| repeat split]. + - reflexivity. + - assumption. + - rewrite Heq_loc. unfold real_target. rewrite dist_sym, dist_swap. + replace Dmax with (D + Dp) by (unfold Dp; lra). + etransitivity; [apply (RealMetricSpace.triang_ineq _ (tgt_loc (choose_target obs))) |]. + apply Rplus_le_compat. + -- rewrite Hr. cbn. rewrite dist_prop. apply move_at_most_D. + -- apply Hdist_main. + - rewrite 2 forever_same_id. apply Hid_r. + + intro Hin. apply Hall, light_on_same_location in Hin; trivial; []. + rewrite Heq_loc, Hin. unfold real_target. + rewrite dist_sym, dist_swap. rewrite Hr in Hdist_main at 1. apply Hdist_main. +Qed. + +End InvariantProofs. + +Definition full_inv config := + scout_inv config + ∧ distinct_id config + ∧ no_collision config + ∧ killed_light_on config + ∧ killer_light_off config + ∧ exists_at_less_than_Dp config + ∧ connected_path config. + +Theorem global_invariant : forall d, Stream.forever (Stream.instant da_assumption) d -> + forall config, full_inv config -> + Stream.forever (Stream.instant full_inv) (execute rbg d config). +Proof using D_pos Dmax_6D. +coinduction Hcorec. ++ assumption. ++ now destruct H. ++ revert_all. intros [? [? [? [? [? [? ?]]]]]]. + destruct H. + repeat split; + apply scout_invariant || + apply distinct_id_invariant || + apply no_collision_invariant || + apply killed_light_on_invariant || + apply killer_light_off_invariant || + apply exists_at_less_than_Dp_invariant || + apply connected_path_invariant; auto. +Qed. + + + + +(** *** Existence of a solution by instantiating the robogram parameters **) + +(** We need a choose function on sets compatible with their equality. + Here, we pick the minimum. *) +Section Compatible_choose. + +Variable A : Type. +Variables (lt : relation A) (lt_dec : forall x y, {lt x y} + {~lt x y}). +Context `{EqDec A}. +Context (lt_compat : Proper (equiv ==> equiv ==> iff) lt). +Context (St : StrictOrder lt). +Variable (trichotomy : forall x y, {lt x y} + {x == y} + {lt y x}). + +Definition set_min (s : set A) : option A := + match choose s with + | None => None + | Some elt => Some (fold (fun e min => if lt_dec e min then e else min) s elt) + end. + +Lemma fold_min_spec : forall s e, + let min := fold (fun e min => if lt_dec e min then e else min) s e in + (min = e \/ InA equiv min (elements s)) + /\ ~lt e min /\ Forall (fun x => ~lt x min) (elements s). +Proof using A lt St lt_compat trichotomy. +intros s e min. subst min. rewrite fold_spec. revert e. +induction (elements s) as [| a l]; simpl; intro e. +* repeat split. + + now left. + + apply irreflexivity. + + constructor. +* destruct_match. + + specialize (IHl a). destruct IHl as [Hin [Hmin1 Hmin2]]. + repeat split. + - right. destruct Hin as [Hin | Hin]; (now rewrite Hin; left) || now right. + - intro Habs. apply Hmin1. now transitivity e. + - now constructor. + + specialize (IHl e). destruct IHl as [Hin [Hmin1 Hmin2]]. + repeat split. + - destruct Hin as [Hin | Hin]; (now left) || (now right; constructor 2). + - assumption. + - constructor; trivial; []. + intro Habs. destruct (trichotomy a e) as [[Hae | Hae] | Hae]. + -- contradiction. + -- rewrite Hae in *. contradiction. + -- apply Hmin1. now transitivity a. +Qed. + +Lemma fold_indep_elt : forall s e1 e2, (e1 ∈ s)%set -> (e2 ∈ s)%set -> + fold (fun e min => if lt_dec e min then e else min) s e1 + == fold (fun e min => if lt_dec e min then e else min) s e2. +Proof using trichotomy lt_compat St. +intros s e1 e2 He1 He2. +destruct (fold_min_spec s e1) as [Hin1 [Hmin1 Hall1]], + (fold_min_spec s e2) as [Hin2 [Hmin2 Hall2]]. +destruct (trichotomy (fold (λ e min : A, if lt_dec e min then e else min) s e1) + (fold (λ e min : A, if lt_dec e min then e else min) s e2)) + as [[Hlt | Heq] | Hlt]. ++ exfalso. rewrite Forall_forall in Hall2. + destruct Hin1 as [Hin1 | Hin1]. + - apply elements_1 in He1. rewrite InA_alt in He1. destruct He1 as [e1' [Heq1 He1]]. + apply Hall2 in He1. rewrite Hin1, Heq1 in *. contradiction. + - rewrite InA_alt in Hin1. destruct Hin1 as [e1' [Heq1 Hin1]]. + apply Hall2 in Hin1. rewrite Heq1 in *. contradiction. ++ assumption. ++ exfalso. rewrite Forall_forall in Hall1. + destruct Hin2 as [Hin2 | Hin2]. + - apply elements_1 in He2. rewrite InA_alt in He2. destruct He2 as [e2' [Heq2 He2]]. + apply Hall1 in He2. rewrite Hin2, Heq2 in *. contradiction. + - rewrite InA_alt in Hin2. destruct Hin2 as [e2' [Heq2 Hin2]]. + apply Hall1 in Hin2. rewrite Heq2 in *. contradiction. +Qed. + +Instance set_min_compat : Proper (equiv ==> equiv) set_min. +Proof using A lt St lt_compat lt_dec trichotomy. +intros s1 s2 Hs. unfold set_min. +destruct (choose s1) as [elt1 |] eqn:Hchoose1. +* assert (Hcase : exists elt2, choose s2 = Some elt2). + { destruct (choose s2) eqn:Habs; try eauto; []. apply choose_2 in Habs. + apply choose_1 in Hchoose1. rewrite Hs in Hchoose1. apply Habs in Hchoose1. tauto. } + destruct Hcase as [elt2 Hcase]. rewrite Hcase. cbn. + transitivity (fold (λ e min : A, if lt_dec e min then e else min) s1 elt2). + + apply fold_indep_elt. + - now apply choose_1. + - rewrite Hs. now apply choose_1. + + apply fold_compat. + - intros ? ? Heq1 ? ? Heq2. do 2 destruct_match; rewrite Heq1, Heq2 in *; intuition. + - intros ? ? ?. repeat destruct_match; auto. + -- exfalso. eapply irreflexivity. etransitivity; eauto. + -- exfalso. revert_one not. intro Habs. apply Habs. etransitivity; eauto. + -- destruct (trichotomy y x) as [[] |]; auto. + -- exfalso. revert_one not. intro Habs. apply Habs. etransitivity; eauto. + - assumption. +* assert (Hcase : choose s2 = None). + { destruct (choose s2) eqn:Habs; auto; []. apply choose_2 in Hchoose1. + apply choose_1 in Habs. rewrite <- Hs in Habs. now apply Hchoose1 in Habs. } + now rewrite Hcase. +Qed. + +Lemma set_min_None : forall s, set_min s = None <-> s == {}%set. +Proof using . +intro s. unfold set_min. destruct (choose s) eqn:Hcase. ++ split; intro Hs; try discriminate;[]. + apply choose_1 in Hcase. rewrite Hs in Hcase. now apply empty_spec in Hcase. ++ split; intros _; trivial; []. + apply choose_2 in Hcase. intro x. rewrite empty_spec. split; try tauto; []. apply Hcase. +Qed. + +Lemma set_min_Some : forall s e, + set_min s = Some e -> (e ∈ s)%set /\ forall x, (x ∈ s)%set -> ~lt x e. +Proof using St lt_compat trichotomy. +intros s e Hin. unfold set_min in Hin. +destruct (choose s) as [a |] eqn:Hcase; try discriminate; []. +inv Hin. destruct (fold_min_spec s a) as [Hin [Hmin Hall]]. split. ++ destruct Hin as [Hin | Hin]. + - rewrite Hin. now apply choose_1. + - now apply elements_2. ++ intros x Hx. apply elements_1 in Hx. rewrite InA_alt in Hx. + destruct Hx as [y [Hxy Hy]]. rewrite Forall_forall in Hall. + rewrite Hxy. now apply Hall. +Qed. + +Lemma set_min_spec : forall s e, + set_min s == Some e <-> (e ∈ s)%set /\ forall x, (x ∈ s)%set -> ~lt x e. +Proof using St lt lt_compat lt_dec trichotomy. +clear Dmax_6D. clear robogram_args. clear Dmax. +intros s e. destruct (set_min s) as [a |] eqn:Hcase. ++ apply set_min_Some in Hcase. cbn. split; intro Hin. + - rewrite <- Hin. setoid_rewrite <- Hin. apply Hcase. + - destruct Hcase, Hin, (trichotomy a e) as [[Hlt | Heq] | Hlt]; firstorder. ++ cbn. split; try tauto; []. + intros [Habs _]. rewrite set_min_None in Hcase. rewrite Hcase in Habs. + revert Habs. apply empty_spec. +Qed. + +End Compatible_choose. + +Definition tgt_lt x y := + Rlt (fst (tgt_loc x)) (fst (tgt_loc y)) + \/ fst (tgt_loc x) = fst (tgt_loc y) /\ Rlt (snd (tgt_loc x)) (snd (tgt_loc y)) + \/ tgt_loc x == tgt_loc y /\ tgt_light x = true /\ tgt_light y = false. + +Definition tgt_lt_dec : forall x y, {tgt_lt x y} + {~tgt_lt x y}. +Proof using . +intros x y. unfold tgt_lt. +destruct (Rlt_dec (fst (tgt_loc x)) (fst (tgt_loc y))). +* now do 2 left. +* destruct (Req_EM_T (fst (tgt_loc x)) (fst (tgt_loc y))); + [destruct (Rlt_dec (snd (tgt_loc x)) (snd (tgt_loc y))) |]. + + left. right. left. now split. + + destruct (tgt_loc x =?= tgt_loc y), (tgt_light x), (tgt_light y); + solve [ right; intro Habs; decompose [and or] Habs; clear Habs; congruence || lra + | left; do 2 right; now repeat split ]. + + right. intro Habs; decompose [and or] Habs; clear Habs; congruence || lra. +Defined. + +Instance tgt_lt_compat : Proper (equiv ==> equiv ==> iff) tgt_lt. +Proof using. intros [] [] [] [] [] []. simpl in *. now subst. Qed. + +Instance tgt_lt_SO : StrictOrder tgt_lt. +Proof using. split. ++ unfold tgt_lt. intros x Habs. decompose [and or] Habs; clear Habs; try congruence; [|]; + eapply (@irreflexivity _ Rlt); eauto; autoclass. ++ unfold tgt_lt. intros x y z Hlt1 Hlt2. + decompose [and or] Hlt1; decompose [and or] Hlt2; clear Hlt1 Hlt2; simpl in *; + solve [ lra | left; congruence | right; left; split; congruence ]. +Qed. + +Lemma tgt_trichotomy : forall x y, {tgt_lt x y} + {x == y} + {tgt_lt y x}. +Proof. +intros x y. destruct (tgt_lt_dec x y) as [| Hxy]; [| destruct (tgt_lt_dec y x) as [| Hyx]]. ++ now do 2left. ++ now right. ++ left; right. + assert (Hfst : fst (tgt_loc x) = fst (tgt_loc y)). + { destruct (total_order_T (fst (tgt_loc x)) (fst (tgt_loc y))) as [[Habs | Habs] | Habs]. + - exfalso. apply Hxy. now left. + - assumption. + - exfalso. apply Hyx. now left. } + assert (Hsnd : snd (tgt_loc x) = snd (tgt_loc y)). + { destruct (total_order_T (snd (tgt_loc x)) (snd (tgt_loc y))) as [[Habs | Habs] | Habs]. + - exfalso. apply Hxy. now right; left. + - assumption. + - exfalso. apply Hyx. now right; left. } + simpl. destruct x as [[] []], y as [[] []]; simpl in *; subst; repeat split; exfalso. + - apply Hxy. do 2 right. repeat split. + - apply Hyx. do 2 right. repeat split. +Defined. + +Local Hint Immediate tgt_trichotomy : core. + +Definition choose_min := set_min tgt_lt tgt_lt_dec. + +Instance choose_min_compat : Proper (equiv ==> equiv) choose_min. +Proof using. repeat intro. unfold choose_min. apply set_min_compat; autoclass. Qed. + +#[refine] +Instance concrete_params : Param := {| + choose_target obs := + match choose_min (filter (fun tgt => negb (tgt_light tgt)) obs) with + | Some tgt => tgt + | None => + match choose_min (filter (fun tgt => Rle_bool (dist (tgt_loc tgt) origin) Dp) obs) with + | Some tgt => tgt + | None => match choose_min obs with + | Some tgt => tgt + | None => (origin, false) (* dummy value *) + end + end + end; + choose_new_pos obs pt := + if Rle_bool (dist pt origin) Dp then origin else (D * unitary pt)%VS +|}. +Proof. +all: autoclass. +* (* choose_target_compat *) + intros obs1 obs2 Hobs. + assert (Hf1 : Proper (equiv ==> equiv) (fun tgt => negb (tgt_light tgt))). + { intros tgt1 tgt2 Htgt. now rewrite Htgt. } + assert (Hf2 : Proper (equiv ==> equiv) (fun tgt => Rle_bool (dist (tgt_loc tgt) 0) Dp)). + { intros tgt1 tgt2 Htgt. now rewrite Htgt. } + assert (Heq1 := filter_m). specialize (Heq1 _ Hf1 obs1 obs2 Hobs). + assert (Heq2 := filter_m). specialize (Heq2 _ Hf2 obs1 obs2 Hobs). + assert (Heq3 := choose_min_compat Hobs). + apply choose_min_compat in Heq1. apply choose_min_compat in Heq2. + repeat destruct_match; auto. +* (* choose_new_pos_compat *) + intros obs1 obs2 Hobs pt1 pt2 Hpt. rewrite Hpt. + destruct_match; try reflexivity; []. now rewrite Hpt. +* (* choose_target_spec *) + assert (Proper (equiv ==> eq) (fun tgt => negb (tgt_light tgt))). + { intros tgt1 tgt2 Htgt. now rewrite Htgt. } + assert (Proper (equiv ==> eq) (fun tgt => Rle_bool (dist (tgt_loc tgt) 0) Dp)). + { intros tgt1 tgt2 Htgt. now rewrite Htgt. } + intros obs Hobs. + destruct (choose_min (filter (fun tgt => negb (tgt_light tgt)) obs)) eqn:Hchoose1; + [| destruct (choose_min (filter (fun tgt => Rle_bool (dist (tgt_loc tgt) 0) Dp) obs)) + eqn:Hchoose2; + [| destruct (choose_min obs) eqn:Hchoose3]]. + + apply set_min_Some in Hchoose1; autoclass; []. + rewrite filter_spec in Hchoose1; trivial; []. destruct Hchoose1 as [Hin Hprop]. + repeat split. + - intuition. + - intro Habs. rewrite Habs in *. intuition discriminate. + - intro Habs. rewrite Habs in *. intuition discriminate. + + apply set_min_Some in Hchoose2; autoclass; []. + rewrite filter_spec in Hchoose2; trivial; []. destruct Hchoose2 as [Hin Hprop]. + assert (Hall : For_all (fun tgt => tgt_light tgt = true) obs). + { intros tgt Htgt. apply set_min_None in Hchoose1. specialize (Hchoose1 tgt). + rewrite <- Bool.not_false_iff_true. intro Habs. apply (empty_spec tgt). + rewrite <- Hchoose1. now rewrite filter_spec, Habs. } + repeat split. + - intuition. + - intros Hlight. apply Hall. + - intros Hlight Habs. exfalso. revert Habs. + apply Rle_not_lt. rewrite dist_sym. now rewrite <- Rle_bool_true_iff. + + apply set_min_Some in Hchoose3; autoclass; []. + assert (Hall1 : For_all (fun tgt => tgt_light tgt = true) obs). + { intros tgt Htgt. apply set_min_None in Hchoose1. specialize (Hchoose1 tgt). + rewrite <- Bool.not_false_iff_true. intro Habs. apply (empty_spec tgt). + now rewrite <- Hchoose1, filter_spec, Habs. } + assert (Hall2 : For_all (fun tgt => dist (tgt_loc tgt) origin > Dp) obs). + { intros tgt Htgt. apply set_min_None in Hchoose2. specialize (Hchoose2 tgt). + apply Rnot_le_gt. intro Habs. apply (empty_spec tgt). + rewrite <- Hchoose2, filter_spec; autoclass; []. + split; trivial; []. unfold Rle_bool. now destruct_match. } + repeat split. + - intuition. + - intros _. apply Hall1. + - intros _ _. setoid_rewrite dist_sym. apply Hall2. + + cbn zeta. exfalso. apply set_min_None in Hchoose3. contradiction. +* (* choose_new_pos_spec *) + intros obs pt Hle. cbn zeta. unfold Rle_bool. destruct_match. + + rewrite dist_same, dist_sym. split; trivial; lra. + + revert_one not. intro Hnot. + assert (Hnull : pt =/= origin). + { intro Habs. apply Hnot. rewrite Habs, dist_same. unfold Dp. lra. } + split. + - rewrite (unitary_id pt) at 2. rewrite dist_sym, norm_dist, <- minus_morph, add_morph. + rewrite norm_mul, norm_unitary, Rmult_1_r; trivial; []. + rewrite norm_dist, opp_origin, add_origin_r in Hnot. + rewrite Rabs_pos_eq. + -- unfold Dp. cut (norm pt <= Dmax); try lra; []. + now rewrite norm_dist, opp_origin, add_origin_r in Hle. + -- cut (D <= norm pt); try lra; []. apply Rnot_le_lt in Hnot. + unfold Dp in *. transitivity (Dmax - D); try lra; []. now apply Rlt_le. + - rewrite norm_dist, opp_origin, add_origin_r, norm_mul, norm_unitary; auto; []. + rewrite Rabs_pos_eq; lra. +Defined. + +End PursuitSetting. diff --git a/CaseStudies/Volumes/NoCollisionAndPath.v b/CaseStudies/Volumes/NoCollisionAndPath.v deleted file mode 100644 index 157996bfb5afa9250a28ea15df42f487980f101c..0000000000000000000000000000000000000000 --- a/CaseStudies/Volumes/NoCollisionAndPath.v +++ /dev/null @@ -1,11121 +0,0 @@ -(* This option was removed in 2019 https://github.com/coq/coq/pull/8094 *) -(*Set Automatic Coercions Import. *) (* coercions are available as soon as functor application *) -Require Import Bool. -Require Import Arith.Div2. -Require Import Lia Field. -Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def. -Require Import List. -Require Import SetoidList. -Require Import Relations. -Require Import RelationPairs. -Require Import Morphisms. -Require Import Psatz. -Require Import Inverse_Image. -Require Import Pactole.Spaces.R2. -Require Import Pactole.Core.Identifiers. -Require Import Pactole.Core.State. -Require Import Pactole.Models.Rigid. -Import Permutation. -Import Pactole.Spaces.Similarity. -Import Datatypes. (* to recover [id] *) -Set Implicit Arguments. -Close Scope R_scope. -Close Scope VectorSpace_scope. -Require Import Rbase. -Require Import SetoidList. -Require Import SetoidDec. -Require Import Decidable. -Require Import Pactole.Util.Preliminary. -Require Import Pactole.Util.Bijection. -Require Import Pactole.Setting. -Require Import Pactole.Util.MMultiset.MMultisetInterface. -Require Import Pactole.Observations.SetObservation. -Require Import Pactole.Observations.PointedObservation. -Set Implicit Arguments. -(* TODO: rename Dmax *) -(* Keeping connection from a fixed base to a free robot. All robots -(including the free one) share the same speed (max distance D in one -round) and the same detection radius (Dmax). The existence of a family -of solution is shown when Dmax > 7*D. No optimality is proved. *) - -(* We have a given finite number n of robot. The final lemma will require -that there are enough robots during the execution. *) -Parameter n: nat. - -(* All robots are non byzantine *) -Instance Identifiers : Names := Robots n 0. - -Definition identifiants := nat. -Definition indentifiants_setoid := eq_setoid nat. -(* Each robots have a light, which state is on or off*) -Definition light := bool. -(* A robot will be alive or not (the latter meaning it has withdrawn -from the task). No collision with dead robots are considered *) -Definition alive := bool. -(* A robot may be still at base, waiting to be needed by the task. No -collision implying non launched rbrobot is considered either. *) -Definition based := bool. -(* Max distance D traveled by a robot between two rounds. This defines - the "death zone": any two robots closer than that are immediately - considered colliding. *) -Parameter D: R. -Axiom D_0: (D>0)%R. - -(* Robots evolve on the plane *) -Instance Loc : Location := make_Location (R2). -Instance VS : RealVectorSpace location := R2_VS. -Instance ES : EuclideanSpace location := R2_ES. -Remove Hints R2_VS R2_ES : typeclass_instances. -(* Max detection radius: robot inside this radius are in the - observation of the robot, the other are not seen. *) -Parameter Dmax : R. -(* We fix the lower bound of Dmax to seven times the max travel - distance in one round. *) -Axiom Dmax_7D : (Dmax > 7*D)%R. - -(* Dp is the distance where a robot is still visible but from which it - becomes able to leave the visibility radius at next round. *) -Definition Dp(*oursuite*) := (Dmax - D)%R. - -(* The state of a robot contains its location and the four following -informations: its name, its light, its liveness and whether it has -been launched or not. *) -Definition ILA :Type := identifiants*light*alive*based. - - -Instance ILA_Setoid : Setoid ILA. -simple refine {| - equiv := fun ila1 ila2 => - match ila1, ila2 with - |(i1, l1, a1, b1), (i2, l2, a2, b2) => - i1 == i2 /\ l1 == l2 /\ a1 == a2 /\ b1 == b2 - end|} -. -Proof. - apply indentifiants_setoid. - apply bool_Setoid. - apply bool_Setoid. - apply bool_Setoid. - split; intro x; intuition; - try destruct x as (((i1, l1), a1), b1). - - destruct y as (((i2, l2), a2), b2); intuition. - - destruct z as (((i3, l3), a3), b3'); - intuition; unfold equiv in *; cbn -[equiv] in *. - now transitivity a0. - now transitivity b4. - now transitivity b3. - now transitivity b2. -Defined. - -Instance ILA_EqDec : EqDec ILA_Setoid. -Proof. - intros x y. - unfold equiv; cbn -[equiv]. - destruct x as (((ix, lx), ax), bx), y as (((iy, ly), ay), By); - unfold light,alive in *. - destruct (bool_eqdec lx ly), (bool_eqdec ax ay), (nat_eq_eqdec ix iy); intuition. - destruct (bool_eqdec bx By); intuition. -Defined. - - -Instance State_ILA : @State (make_Location R2) (R2*ILA) := - @AddInfo Loc R2 ILA _ _ (OnlyLocation (fun _ => location)). - - -(* Trying to avoid notation problem with implicit arguments *) - - -Notation "s [ x ]" := (multiplicity x s) (at level 2, no associativity, format "s [ x ]"). -Notation "!! config" := (@obs_from_config location _ _ _ set_observation config origin) (at level 10). -Notation support := (@support location _ _ _). - -Definition config:= @configuration Loc (R2*ILA) State_ILA Identifiers. - -Definition get_ident (elt : R2*ILA) : identifiants := fst(fst(fst(snd elt))). -Instance get_ident_compat : Proper (equiv ==> equiv) get_ident. -Proof. intros ? ? Heq. - destruct x as (px,(((ix,lx),ax), bx)), y as (py,(((iy,li),ay), By)), Heq as (_,(Hi,(_,_))). - simpl. - now unfold equiv, nat_Setoid, eq_setoid in Hi. -Qed. - -Axiom ident_unique : forall conf g g', g <> g' -> get_ident (conf (Good g)) <> get_ident (conf (Good g')). - - -Definition get_light (elt : R2*ILA) : light := snd(fst(fst(snd elt))). -Instance get_light_compat : Proper (equiv ==> equiv) get_light. -Proof. intros ? ? Heq. - destruct x as (px,(((ix,lx),ax), bx)), y as (py,(((iy,li),ay), By)), Heq as (_,(_,(Hl,_))). - simpl. - now unfold equiv, nat_Setoid, eq_setoid in Hl. -Qed. - -Definition get_alive (elt : R2*ILA) : alive := snd(fst(snd(elt))). -Instance get_alive_compat : Proper (equiv ==> equiv) get_alive. -Proof. intros ? ? Heq. - destruct x as (px,(((ix,lx),ax), bx)), y as (py,(((iy,li),ay), By)), Heq as (_,(_,(_,Ha))). - simpl. - now unfold equiv, nat_Setoid, eq_setoid in Ha. -Qed. - -Definition get_based (elt : R2*ILA) : alive := snd(snd(elt)). -Instance get_based_compat : Proper (equiv ==> equiv) get_based. -Proof. intros ? ? Heq. - destruct x as (px,(((ix,lx),ax), bx)), y as (py,(((iy,li),ay), By)), Heq as (_,(_,(_,Ha))). - simpl. - now unfold equiv, nat_Setoid, eq_setoid in Ha. -Qed. - - -(* In the problem of connection, the free robot (numbered 0) is - launched and alive. By definition of the problem. *) -Axiom ident_0_alive : forall conf g, get_ident (conf (Good g)) = 0 -> - get_alive (conf (Good g)) = true. - -Axiom ident_0_not_based : forall conf g, get_ident (conf (Good g)) = 0 -> - get_based (conf (Good g)) = false. - -Definition upt_robot_dies_b (config:config) (g:G) : bool := - existsb (fun elt : R2 * ILA => Rle_bool (dist (@get_location Loc _ _ elt) (@get_location Loc _ _ (config (Good g)))) D) - (List.filter (fun (elt : R2*ILA) => (get_ident(elt) <? get_ident (config (Good g))) && get_alive (elt) && (negb (get_based (elt)))) (config_list config)) - || negb - (existsb (fun elt : R2 * ILA => (Rle_bool (dist (@get_location Loc _ _ elt) (@get_location Loc _ _ (config (Good g)))) Dmax)) - (List.filter (fun (elt : R2*ILA) => (get_ident(elt) <? get_ident (config (Good g))) && get_alive (elt) && (negb (get_based (elt))))(config_list config))). - -Instance upt_robot_dies_b_compat : Proper (equiv ==> eq ==> equiv) upt_robot_dies_b. -Proof. - intros x y Hconf g g' Hg. - rewrite <- Hg. - assert (H_x := Hconf (Good g)). - unfold upt_robot_dies_b. - rewrite 2 orb_lazy_alt. - destruct (existsb - (fun elt : R2 * ILA => - Rle_bool (dist (get_location elt) (get_location (x (Good g)))) D) - (List.filter (fun elt : R2 * ILA => (get_ident elt <? get_ident (x (Good g))) && get_alive elt && negb (get_based elt)) - (config_list x))) eqn : Hex_x1; - destruct (existsb - (fun elt : R2 * ILA => - Rle_bool (dist (get_location elt) (get_location (y (Good g)))) D) - (List.filter (fun elt : R2 * ILA => (get_ident elt <? get_ident (y (Good g))) && get_alive elt && negb (get_based elt)) - (config_list y))) eqn : Hex_y1; try (now simpl in *); - destruct (negb - (existsb - (fun elt : R2 * ILA => - Rle_bool (dist (get_location elt) (get_location (x (Good g)))) Dmax) - (List.filter - (fun elt : R2 * ILA => (get_ident elt <? get_ident (x (Good g))) && get_alive elt && negb (get_based elt)) - (config_list x)))) eqn : Hex_x2; - - destruct (negb - (existsb - (fun elt : R2 * ILA => - Rle_bool (dist (get_location elt) (get_location (y (Good g)))) Dmax) - (List.filter - (fun elt : R2 * ILA => (get_ident elt <? get_ident (y (Good g))) && get_alive elt && negb (get_based elt)) - (config_list y)))) eqn : Hex_y2; - try ( - generalize (config_list_compat Hconf); intro Hcomp; try ( assert (true == false) by now (rewrite <- Hex_y1, <- Hex_x1; - apply (@existsb_compat _ equiv); - try (now intros w z Hwz; rewrite H_x, Hwz); - try (f_equiv; try (intros v w Hvw; - now rewrite H_x, Hvw); - try assumption))); - try discriminate); try auto. - - + rewrite negb_false_iff, negb_true_iff in *. - rewrite <- Hex_y2, <- Hex_x2. - rewrite (get_ident_compat (Hconf (Good g))). - rewrite (get_location_compat _ _ (Hconf (Good g))). - apply (@existsb_compat _ equiv). - auto. - intros u v Huv. - rewrite (get_location_compat _ _ Huv). - auto. - f_equiv. - intros a b Hab. - rewrite Hab. - reflexivity. - rewrite Hcomp. - reflexivity. - + rewrite negb_false_iff, negb_true_iff in *. - rewrite <- Hex_y2, <- Hex_x2. - apply (@existsb_compat _ equiv). - intros w z Hwz. - destruct H_x as (Hr,_). - rewrite Hwz, Hconf. - now simpl. - f_equiv. - intros v w Hvw; now rewrite Hvw, Hconf. now rewrite Hconf. -Qed. - -Definition too_far_aux_1 config g g' := ((if (negb (Nat.eqb (get_ident (config (Good g))) (get_ident (config (Good g'))))) then - ((Nat.ltb (get_ident (config (Good g))) (get_ident (config (Good g'))))) else true)). - - -Instance too_far_aux_1_compat : Proper (equiv ==> eq ==> eq ==> eq) too_far_aux_1. -Proof. - intros conf1 conf2 Hconf g1 g2 Hg g'1 g'2 Hg'. - rewrite Hg, Hg'. - unfold too_far_aux_1. - simpl. - - rewrite 2 (get_ident_compat (Hconf (Good g2))). - rewrite 2 (get_ident_compat (Hconf (Good g'2))). - reflexivity. -Qed. - - -Definition too_far_aux_2 config g g' := (negb (andb (negb (get_based (config (Good g')))) - (andb (get_alive (config (Good g'))) - (Rle_bool (dist (get_location (config (Good g'))) - (get_location (config (Good g)))) - (Dp-3*D)%R)))). - -Instance too_far_aux_2_compat : Proper (equiv ==> eq ==> eq ==> eq) too_far_aux_2. -Proof. - intros conf1 conf2 Hconf g1 g2 Hg g'1 g'2 Hg'. - rewrite Hg, Hg'. - unfold too_far_aux_2. - rewrite (get_based_compat (Hconf (Good g'2))). - rewrite (get_alive_compat (Hconf (Good g'2))). - rewrite (get_location_compat _ _ (Hconf (Good g2))). - rewrite (get_location_compat _ _ (Hconf (Good g'2))). - reflexivity. -Qed. - - -Definition upt_robot_too_far config g fchoice:= - if R2dec_bool (retraction fchoice (get_location (config (Good g)))) (0,0)%R - then if (forallb (too_far_aux_1 config g) (List.filter (fun g' => get_based (config (Good g')) && get_alive (config (Good g'))) Gnames)) - then if (forallb (too_far_aux_2 config g) Gnames) then true else false - else false - else false. - - -Instance upt_robot_too_far_compat : Proper (equiv ==> eq ==> equiv ==> eq) upt_robot_too_far. -Proof. - intros c1 c2 Hc g1 g2 Hg ch1 ch2 Hch. - unfold upt_robot_too_far. - rewrite (too_far_aux_1_compat Hc Hg). - rewrite (too_far_aux_2_compat Hc Hg). - rewrite Hg, (get_location_compat _ _ (Hc (Good g2))). - rewrite (retraction_compat Hch _ _ (reflexivity _)). - assert ((List.filter (fun g' : G => get_based (c1 (Good g')) && get_alive (c1 (Good g'))) Gnames) - = (List.filter (fun g' : G => get_based (c2 (Good g')) && get_alive (c2 (Good g'))) Gnames)). - apply filter_ext_in. - intros g Hin_g. - rewrite (get_based_compat (Hc (Good g))), (get_alive_compat (Hc (Good g))); auto. - rewrite H. - reflexivity. -Qed. - - -Instance RL_Setoid : Setoid (R2*light) := prod_Setoid R2_Setoid bool_Setoid. - - - -Instance robot_choice_RL : robot_choice (R2*light) := -{| robot_choice_Setoid := _ |}. - - -Instance Choice : update_choice bool := - {| update_choice_Setoid := _ - |}. - - -Definition bij_reflexion_f := - fun (point : R2) => - let (x,y) := point in - (x,(-y)%R). - -Corollary bij_reflexion_Inversion : forall (x y:R2), - bij_reflexion_f x == y <-> bij_reflexion_f y == x. -Proof. - intros (xa,xb) (ya,yb). - split; intros Heq; - rewrite <-Heq; unfold bij_reflexion_f; - intuition; - now rewrite Ropp_involutive. -Qed. - -Definition bij_reflexion : Bijection.bijection R2 := {| - Bijection.section := bij_reflexion_f; - Bijection.retraction := bij_reflexion_f; - Bijection.Inversion := bij_reflexion_Inversion |}. - -Lemma reflexion_zoom : forall x y, dist (bij_reflexion x) (bij_reflexion y) = (1 * dist x y)%R. -Proof. - intros (x1, x2) (y1, y2). - cbn in *. - rewrite Ropp_involutive. - assert (H: forall a b, (( a +- b) * (a +- b))%R = (a * a + b * b - 2 * a * b)%R). - intros. - replace (a +- b)%R with (a-b)%R by lra. - replace ((a - b) * (a - b))%R with (a-b)²%R. - rewrite R_sqr.Rsqr_minus. - now unfold Rsqr. - now unfold Rsqr. - replace (- x2 + y2)%R with (y2 +- x2)%R by lra. - repeat rewrite H. - rewrite Rmult_1_l. - f_equiv. - lra. -Qed. - - -Definition reflexion := - {| - sim_f := bij_reflexion; - zoom := 1; - dist_prop := reflexion_zoom|}. - -Instance Frame : @frame_choice Loc (R*(R2)*bool). -refine {| - frame_choice_bijection := - fun (elt:(R*(R2)*bool)) => - let (p,b) := elt in - let (r,t) := p in - compose (rotation r) - (compose (translation t) - (if b then reflexion else @Similarity.id R2 R2_Setoid _ _ _)) - - ; - frame_choice_Setoid := prod_Setoid (prod_Setoid R_Setoid R2_Setoid) bool_Setoid |} -. -Proof. - intros [[b1 v1] a1] [[b2 v2] a2] [[Hb Hv] Ha]. cbn -[equiv] in *. - repeat f_equiv; trivial; []. now rewrite Ha. -Defined. - -Lemma frame_dist : forall pos1 pos2 tframe, - let bij := frame_choice_bijection tframe in - dist pos1 pos2 == dist (bij pos1) (bij pos2). -Proof. - intros pos1 pos2 ((rot, (tra1, tra2)), b). - - unfold frame_choice_bijection, Frame. - assert (Hdist := dist_prop (rotation rot - ∘ (translation (tra1, tra2) ∘ (if b then reflexion else Similarity.id)))). - rewrite Hdist. - destruct b eqn : Hb; simpl; - lra. -Qed. - - -Definition upt_aux (config:config) (g: G) (rl : R2*light) fchoice: (R2*ILA) := - match config (Good g) with (r,(((i,l),a),b)) => - if b then - if upt_robot_too_far config g (frame_choice_bijection fchoice) then - (r, (((i,false),true),false)) - else - config (Good g) - else - if upt_robot_dies_b config g - then (r,(((i,l),false),b)) - else (fst rl,(((i,snd rl), a),b)) - end. - - - -Instance upt_aux_compat : Proper (equiv ==> eq ==> equiv ==> equiv ==> equiv) upt_aux. -Proof. - repeat intro. - unfold upt_aux. - rewrite (upt_robot_dies_b_compat H H0). - rewrite (@upt_robot_too_far_compat x y H x0 y0 H0 _ _ (@frame_choice_bijection_compat (@make_Location R2 R2_Setoid R2_EqDec) _ _ _ _ H2)). - - rewrite H0. - - specialize (H (Good y0)). - destruct (x (Good y0)) as (?,(((?,?),?),?)); - destruct (y (Good y0)) as (?,(((?,?),?),?)); - destruct b, b0; - try auto; - try now simpl in H. - - destruct (upt_robot_dies_b y y0); - destruct (upt_robot_too_far y y0); - simpl; - destruct (x (Good x0)) as (?,((?,?),?)); - destruct (y (Good x0)) as (?,((?,?),?)); - destruct H as (?,(?,(?,?))); - try now simpl in *. - destruct (upt_robot_dies_b y y0); - destruct (upt_robot_too_far y y0); - simpl; - destruct (x (Good x0)) as (?,((?,?),?)); - destruct (y (Good x0)) as (?,((?,?),?)); - destruct H as (?,(?,(?,?))); - try now simpl in *. - apply id. - apply id. -Qed. - - -Instance UpdFun : @update_function (R2*ILA) Loc State_ILA _ (R2*light)(* Trobot *) - (R*R2*bool)(* Tframe *) bool (* Tactive *) robot_choice_RL Frame Choice. -simple refine {| - update := fun config g r rl _ => upt_aux config g rl r|}. -Proof. - intros c1 c2 Hc g1 g2 Hg t1 t2 Ht ? ? ? ? ? ?. - apply upt_aux_compat; intuition. -Defined. - - - - -Instance Obs_ILA : Observation . - refine {| - observation := set (R2*ILA); - observation_Setoid := (Set_Setoid (R2*ILA)%type); - observation_EqDec := (@Set_EqDec (R2*ILA)%type _ _ (@FSetList.SetList (R2*ILA)%type _ _) - (@FSetList.SetListFacts (R2*ILA)%type _ _) ); - - obs_from_config (config:config) (pt:R2*ILA) := - @SetObservation.make_set (R2*ILA)%type _ (@state_EqDec _ _ State_ILA) - (List.filter - (fun elt => (Rle_bool (dist (fst elt) (fst pt)) Dmax) - && (get_alive elt) && get_alive pt && (get_ident elt <=? get_ident pt)) - (config_list config)); - obs_is_ok s config pt := - forall l, In l s <-> InA equiv l (config_list config) /\ (dist (fst l) (fst pt) <= Dmax)%R /\ get_alive l == true /\ get_alive pt = true /\ get_ident l <= get_ident pt |}. - Proof. - - intros config1 config2 Hconfig pt1 pt2 Hpt. - apply SetObservation.make_set_compat, eqlistA_PermutationA_subrelation. - f_equiv. - + intros ? ? Heq. rewrite (get_alive_compat Heq). - rewrite (dist_compat _ _ (fst_compat Heq) _ _ (fst_compat Hpt)). - rewrite Hpt, (get_ident_compat Heq). - auto. - + now rewrite (config_list_compat Hconfig). - - intros config pt x. - rewrite SetObservation.make_set_spec. - rewrite filter_InA. - split; intros Hlist. - + destruct Hlist as (rila, Htemp). - apply andb_prop in Htemp. - destruct Htemp as (Hdist, Hident). - apply andb_prop in Hdist. - destruct Hdist as (Hdist, Halive_pt). - apply andb_prop in Hdist. - destruct Hdist as (Hdist, Halive_x). - rewrite Rle_bool_true_iff in Hdist. - intuition. - intuition. - apply leb_complete in Hident. - intuition. - + repeat split; try repeat apply andb_true_intro; try split; destruct Hlist as (?,(?,?)); - try rewrite Rle_bool_true_iff in *; intuition. - apply andb_true_intro. - try split; - try rewrite Rle_bool_true_iff in *; intuition. - try rewrite Rle_bool_true_iff in *; intuition. - rewrite andb_true_iff. - split. - rewrite Rle_bool_true_iff. - auto. - auto. - simpl in H2. - unfold equiv in H2. - auto. - simpl. - now rewrite Nat.leb_le. - + intros (?,(((?,?),?),?)) (?,(((?,?),?),?)) (?,(?,(?,(?,?)))). simpl in *. - now rewrite H, H0, H1, H2. -Defined. - -Definition obs_ILA := @observation (R2*ILA) Loc State_ILA _ Obs_ILA. - -(* ******************************************** *) -(* The following axioms (introduced by Context or Axiom) are the - constraints a protocol should verify to be proved correct in the - following. These constraints define the family of protocols we - prove correct. *) -(* ******************************************** *) - -(* ********* Hypothesis on chose_target ******* *) -(* The protocole should provide a function that chose a robot to - "follow" among it observable (visible) ones. *) -Context {choose_target : obs_ILA -> (R2*ILA)}. -Context {choose_target_compat : Proper (equiv ==> equiv) choose_target}. - -(* This function should always chose a robot alive. *) -Axiom choose_target_alive : forall obs target, - choose_target obs == target -> get_alive target == true. - -(* It should always prefer unlit robots. *) -Axiom light_off_first: forall obs target, - choose_target obs == target -> - (* if the chosen rrobot is lit *) - get_light (target) == true -> - (* Then there was no unlitten ones *) - For_all (fun (elt: R2*ILA) => get_light elt == true) obs . - -(* it should always return a robot whose id is strictly smaller than -the observing robot itself (i.e. the one that is at (0,0)). *) -Axiom target_lower : forall obs target (zero: R2*ILA), - In zero obs -> get_location zero == (0,0)%R -> - choose_target obs == target -> - get_ident target < get_ident zero. - - -(* If it choses a lit robot despite light_off_first, it should prefer - robot closer than Dp if possible. *) -Axiom light_close_first : forall obs target, - choose_target obs == target -> - get_light target == true -> - ((dist (0,0)%R (get_location target)) > Dp)%R -> - For_all (fun (elt : R2*ILA) => - ((dist (0,0)%R (get_location elt)) > Dp \/ get_location elt = (0,0))%R) obs. - -(* It should always return a visible robot. *) -Axiom choose_target_in_obs : forall obs target, - choose_target obs = target -> In target obs. - -(* ********* Hypothesis on move_to ********* *) -(* move_to returns true if it is safe to go to a target.. *) - -Context {move_to: obs_ILA -> location -> bool }. - -(* it should return true only if there are no *other* robot near the - target (distance > 2*D) *) -Axiom move_to_Some_zone : forall obs choice, - move_to obs choice = true -> - (forall x, In x obs -> let (pos_x, light_x) := x in - dist pos_x (0,0) > 0 -> - dist choice pos_x > 2*D)%R. - - -(* it should return false if there is at least another observable - robots and one robot smaller than an observable one is too close to - target. *) -Axiom move_to_None : - forall (obs : obs_ILA) (choice : location), - move_to obs choice = false -> - exists other : R2 * ILA, - (other ∈ obs)%set /\ - (exists inf, In inf obs /\ get_ident other < get_ident inf) /\ - (dist (get_location other) choice <= 2*D)%R. - -Context {move_to_compat : Proper (equiv ==> equiv ==> equiv) move_to}. - -(* ********* Hypothesis on choose_new_pos ********* *) -(* choose_new_pos returns the position to aim, givent he robot to follow *) - -Context {choose_new_pos: obs_ILA -> location -> location}. -Context {choose_new_pos_compat : Proper (equiv ==> equiv ==> equiv) choose_new_pos}. - -(* It should always return a target reachable in 1 round, and within - reasonable range of the robot folloed. *) -Axiom choose_new_pos_correct : forall obs target new, - new == choose_new_pos obs target -> - (dist new target <= Dp /\ dist new (0,0) <= D)%R. - -(* The protocol is defined as follows, it makes use of the three - generic auxiliary functions defined above. *) -Definition rbg_fnc (s:obs_ILA) : R2*light - := - (* on choisit la target *) - let target := choose_target s in - let new_pos := choose_new_pos s (fst target) in - match move_to s new_pos with - (* Si on a pu bouger, on change la position, et on est pas allumé *) - | true => (new_pos,false) - (* si on a pas pu bouger, on change de couleur *) - | false => ((0,0)%R, true) - end. - - -Instance rbg_compat : Proper (equiv ==> equiv) rbg_fnc. -Proof. - intros s1 s2 Hs. - unfold rbg_fnc. - assert (Hmcompat := move_to_compat Hs (choose_new_pos_compat Hs (fst_compat (choose_target_compat Hs )))). - rewrite Hmcompat. - rewrite (choose_new_pos_compat Hs (fst_compat (choose_target_compat Hs ))). - reflexivity. -Qed. - -Definition rbg_ila : robogram := - {| pgm := rbg_fnc |}. - -(* This needs to instantiated for our generic model to fully - operational, but as there are no byzantine robot it not used - anywhere. *) -Context {inactive_choice_ila : inactive_choice bool}. - -(* A demon in this context. *) -Definition da_ila := @demonic_action (R2*ILA) Loc State_ILA _ (R2*light) (R*R2*bool) bool bool robot_choice_RL Frame Choice _. - -(* This says that a demonic choice always centers the observation on -the ibserving robot. *) -Definition change_frame_origin (da:da_ila):= forall (config:config) g (tframe:R*R2*bool), - da.(change_frame) config g = tframe -> - let bij := frame_choice_bijection tframe in - bij (get_location (config (Good g))) == (0,0)%R. - -(* da_predicat d means that d is Fully synchronous and follows the above property. *) -Definition da_predicat (da:da_ila) := (change_frame_origin da) (*/\ (choose_update_close da)*) /\ FSYNC_da da. - -Definition demon_ila_type := @demon (R2*ILA) Loc State_ILA _ (R2*light) (R*R2*bool) bool bool robot_choice_RL Frame Choice _. - -Definition demon_ILA (demon : demon_ila_type) := Stream.forever (Stream.instant da_predicat) demon. - -(* This needs to instantiated for our generic model to fully - operational, but as there are no byzantine robot it not used - anywhere. *) -Context {inactive_ila : @inactive_function (R2 * ILA) Loc State_ILA Identifiers bool - inactive_choice_ila}. - -(* Initial configuration -===============================================================================*) - -(* All this properties state that the original configuration verifies - the global invariant + a few others specific properties.. *) - -(* TODO: Same thig for DA/le démon, have a predicate ""global"" *) - -Definition config_init_not_close (config_init:config) := forall id, - match (config_init id) with - (pos, (((ident, light), alive),based)) => - based = false -> - forall id', - match config_init id' with - (pos', (((ident', light'), alive'), based')) => - based' = false -> (dist pos pos' > D)%R end end. - -(* tout le monde est eteint au début *) -Definition config_init_off (config_init:config) := forall id, - match (config_init id) with - (_,((( _, light), _),_)) => light = false - end. - -(* tout robot voit un robot inférieur, ou alors on est le robot 1. *) -Definition config_init_lower (config_init:config) := forall id, - match (config_init id) with - (pos,(((ident,light),alive), based)) - => alive = true -> - ident = 0 \/ - exists id', match(config_init id') with - (pos',(((ident',_),alive'), _)) => - alive' = true /\ - ident' < ident /\ - (dist pos pos' < Dmax)%R - end - end. - - -Definition config_init_base_linked config := - exists g, get_based (config (Good g)) = false - /\ get_alive (config (Good g)) = true - /\ (dist (get_location (config (Good g))) (0,0)%R <= Dp-D)%R. - -Definition config_init_based_higher_id config := - forall g g', get_based (config (Good g)) = false -> get_based (config (Good g')) = true -> - get_ident (config (Good g)) < get_ident (config (Good g')). - -Definition config_init_based_0 config := - forall g, get_based (config (Good g)) = true -> get_alive (config (Good g)) = true /\ - get_location (config (Good g)) = (0,0)%R. - -Definition config_init_pred config := - config_init_lower config - /\ config_init_off config - /\ config_init_not_close config - /\ config_init_base_linked config - /\ config_init_based_higher_id config -/\ config_init_based_0 config. - - -Lemma let_split : forall {A B C E:Type} (x w:A) (y t:B) (a b:C) (u v : E) (z: A*B*C*E), - let (q,u) := z in - let (p,a) := q in - let (x,y):= p in - let (q0, v) := z in - let (p0,b) := q0 in - let (w,t) := p0 in - x=w /\ y=t /\ a = b /\ u = v. -Proof. intros. destruct z, p, p. now simpl. Qed. - -Lemma round_good_g : - forall g config (da:da_ila), - da_predicat da -> - round rbg_ila da config (Good g) - == - let frame_choice := change_frame da config g in - let new_frame := frame_choice_bijection frame_choice in - let local_config := - map_config - (lift (existT precondition new_frame (precondition_satisfied da config g))) - config in -(* let local_pos := get_location (local_config (Good g)) in*) - let obs := obs_from_config local_config (local_config (Good g)) in - let local_robot_decision := rbg_ila obs in - let choice := choose_update da local_config g local_robot_decision in - let new_local_state := - update local_config g frame_choice local_robot_decision choice in - lift - (existT precondition (new_frame â»Â¹) (precondition_satisfied_inv da config g)) - new_local_state -. -Proof. - intros. - destruct H as (Hchange, HFSYNC). - specialize (HFSYNC (Good g)). - unfold round. - destruct (activate da (Good g)) eqn : Hact; try discriminate. - - simpl. - repeat split. - apply let_split. - apply 0. apply 0. - apply true. - apply true. - apply true. - apply true. - apply true. - apply true. -Qed. - -Ltac changeR2 := - change R2 with location in *; - change R2_Setoid with location_Setoid in *; - change R2_EqDec with location_EqDec in *; - change R2_VS with VS in *; - change R2_ES with ES in *. - -Lemma robot_dead_means : - forall config g (da:da_ila), - da_predicat da -> - get_alive (round rbg_ila da config (Good g)) == false - -> get_alive (config (Good g)) == false - \/ (exists g', - get_ident (config (Good g')) < get_ident (config (Good g)) /\ - get_alive (config (Good g')) = true /\ - get_based (config (Good g')) = false /\ - Rle_bool (dist (get_location (config (Good g))) - (get_location (config (Good g')))) D = true) - \/ (forall g', - get_ident (config (Good g')) < get_ident (config (Good g)) -> - get_alive (config (Good g')) = true -> - get_based (config (Good g')) = false -> - negb (Rle_bool - (dist (get_location (config (Good g))) - (get_location (config (Good g')))) Dmax) = true). -Proof. - intros config0 g da Hda_pred Hround. - rewrite (round_good_g g config0 Hda_pred) in Hround. - simpl in Hround. - changeR2. - unfold Datatypes.id in *. - unfold upt_aux in *. - destruct (config0 (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconfig. - simpl in Hround. - destruct (upt_robot_dies_b _) eqn : Hrobot_dies; - destruct (upt_robot_too_far _ _) eqn : Hrobot_far; - simpl in Hround; - unfold get_alive in *; - simpl in Hround. - - - - unfold upt_robot_dies_b in Hrobot_dies. - rewrite orb_true_iff in Hrobot_dies; destruct Hrobot_dies as [Hex|Hneg]. - + rewrite existsb_exists in Hex. - destruct Hex as (rila, (Hin,Hex)). - right; left. - rewrite filter_In in Hin. - rewrite config_list_spec, in_map_iff in Hin. - destruct Hin as ((id, HIn),Hid_low). - destruct (change_frame da config0 g) as ((r,t),b). - rewrite Hconfig in *. - destruct id as [g'|byz]. - simpl in Hid_low. - exists (g'). - split. - * destruct HIn as (Heq, _). - rewrite <- Heq in Hid_low. - unfold get_ident in *; simpl in *. - rewrite 2 andb_true_iff in *. - destruct Hid_low as (?,?). - rewrite Hconfig in *. - now rewrite Nat.ltb_lt in *. - * destruct HIn as (Heq, Hin). - repeat simpl in Heq. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location in *. - unfold Datatypes.id in *. - simpl (fst _) in *. - assert (Heq' : (R2.bij_rotation_f r - ((fst - ((if b then reflexion else Similarity.id) (fst (config0 (Good g')))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (config0 (Good g')))) + - snd t)%R), snd (config0 (Good g'))) == rila) by now rewrite <- Heq. - destruct Heq' as (Heq', Helse). - rewrite (frame_dist pos (fst (config0 (Good g'))) (r,t,b)). - rewrite <- Heq' in Hex. - unfold frame_choice_bijection, Frame. - rewrite 2 andb_true_iff in *. - destruct Hid_low as ((Hid_low, Halive), Hbased). - repeat split. - -- unfold get_alive in Halive. - destruct (config0 (Good g')) as (?,(((?,?),?),?)) eqn : ?. - simpl in Helse. - destruct rila. - simpl in *. - destruct i0, p,p; simpl in *; intuition. - now rewrite H0. - -- rewrite <- Heq in Hbased. - unfold get_based in *; simpl in *. - now rewrite negb_true_iff in *. - -- destruct b; rewrite dist_sym; rewrite Hconfig in *; - simpl in *; - now rewrite Hex. - * assert (Hbyz := @In_Bnames _ byz). - unfold Bnames, Identifiers in *. - now simpl in *. - + rewrite forallb_existsb, forallb_forall in Hneg. - right; right. - destruct (change_frame da config0 g) as ((r,t),b) eqn : Hchange. - intros g' Hid Halive Hbased'; specialize (Hneg ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (config0 (Good g'))), - snd (config0 (Good g')))). - assert (Hin :@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst (@location Loc) ILA (config0 (@Good Identifiers g')))) - (@snd (@location Loc) ILA (config0 (@Good Identifiers g')))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@map_config Loc (prod R2 ILA) (prod R2 ILA) State_ILA State_ILA - Identifiers - (fun x : prod R2 ILA => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES)) - end))) (@fst R2 ILA x)) (@snd R2 ILA x)) config0 - (@Good Identifiers g)))) (get_alive elt)) - (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (@map_config Loc (prod R2 ILA) (prod R2 ILA) State_ILA State_ILA Identifiers - (fun x : prod R2 ILA => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA x)) (@snd R2 ILA x)) config0)))). - { rewrite filter_In. - split. - * rewrite config_list_spec. - rewrite in_map_iff. - exists (Good g'). - split; intuition. - destruct b; reflexivity. - apply In_names. - * rewrite 2 andb_true_iff. - unfold get_ident in *; simpl in *. - split. - rewrite Nat.ltb_lt in *. - now rewrite Hconfig; simpl. - unfold get_based in *; simpl in *. - now rewrite negb_true_iff. - } - specialize (Hneg Hin). - rewrite <- Hneg. - do 2 f_equiv. - assert (Hdist := frame_dist pos (fst (config0 (Good g'))) ((r, t), b)). - unfold frame_choice_bijection, Frame in Hdist. - rewrite Hdist. - unfold map_config. - rewrite Hconfig in *. - destruct b; - rewrite dist_sym; simpl; intuition. - - - unfold upt_robot_dies_b in Hrobot_dies. - rewrite orb_true_iff in Hrobot_dies; destruct Hrobot_dies as [Hex|Hneg]. - + rewrite existsb_exists in Hex. - destruct Hex as (rila, (Hin,Hex)). - right; left. - rewrite filter_In in Hin. - rewrite config_list_spec, in_map_iff in Hin. - destruct Hin as ((id, HIn),Hid_low). - destruct (change_frame da config0 g) as ((r,t),b). - rewrite Hconfig in *. - destruct id as [g'|byz]. - simpl in Hid_low. - exists (g'). - split. - * destruct HIn as (Heq, _). - rewrite <- Heq in Hid_low. - unfold get_ident in *; simpl in *. - rewrite 2 andb_true_iff in *. - destruct Hid_low as (?,?). - rewrite Hconfig in *. - now rewrite Nat.ltb_lt in *. - * destruct HIn as (Heq, Hin). - repeat simpl in Heq. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location in *. - unfold Datatypes.id in *. - simpl (fst _) in *. - assert (Heq' : (R2.bij_rotation_f r - ((fst - ((if b then reflexion else Similarity.id) (fst (config0 (Good g')))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (config0 (Good g')))) + - snd t)%R), snd (config0 (Good g'))) == rila) by now rewrite <- Heq. - destruct Heq' as (Heq', Helse). - rewrite (frame_dist pos (fst (config0 (Good g'))) (r,t,b)). - rewrite <- Heq' in Hex. - unfold frame_choice_bijection, Frame. - rewrite 2 andb_true_iff in *. - destruct Hid_low as ((Hid_low, Halive), Hbased). - repeat split. - -- unfold get_alive in Halive. - destruct (config0 (Good g')) as (?,(((?,?),?),?)) eqn : ?. - simpl in Helse. - destruct rila. - simpl in *. - destruct i0, p,p; simpl in *; intuition. - now rewrite H0. - -- rewrite <- Heq, negb_true_iff in Hbased. - unfold get_based in *; simpl in *; auto. - -- destruct b; rewrite dist_sym; rewrite Hconfig in *; - simpl in *; - now rewrite Hex. - * assert (Hbyz := @In_Bnames _ byz). - unfold Bnames, Identifiers in *. - now simpl in *. - + rewrite forallb_existsb, forallb_forall in Hneg. - right; right. - destruct (change_frame da config0 g) as ((r,t),b) eqn : Hchange. - intros g' Hid Halive Hbased; specialize (Hneg ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (config0 (Good g'))), - snd (config0 (Good g')))). - assert (Hin :@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst (@location Loc) ILA (config0 (@Good Identifiers g')))) - (@snd (@location Loc) ILA (config0 (@Good Identifiers g')))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@map_config Loc (prod R2 ILA) (prod R2 ILA) State_ILA State_ILA - Identifiers - (fun x : prod R2 ILA => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES)) - end))) (@fst R2 ILA x)) (@snd R2 ILA x)) config0 - (@Good Identifiers g)))) (get_alive elt)) - (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (@map_config Loc (prod R2 ILA) (prod R2 ILA) State_ILA State_ILA Identifiers - (fun x : prod R2 ILA => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA x)) (@snd R2 ILA x)) config0)))). - { rewrite filter_In. - split. - * rewrite config_list_spec. - rewrite in_map_iff. - exists (Good g'). - split; intuition. - destruct b; reflexivity. - apply In_names. - * rewrite 2 andb_true_iff. - unfold get_ident in *; simpl in *. - split. - rewrite Nat.ltb_lt in *. - now rewrite Hconfig; simpl. - now unfold get_based; rewrite negb_true_iff; simpl. - } - specialize (Hneg Hin). - rewrite <- Hneg. - do 2 f_equiv. - assert (Hdist := frame_dist pos (fst (config0 (Good g'))) ((r, t), b)). - unfold frame_choice_bijection, Frame in Hdist. - rewrite Hdist. - unfold map_config. - rewrite Hconfig in *. - destruct b; - rewrite dist_sym; simpl; intuition. - - changeR2. - left. - rewrite Hconfig in Hround. - unfold get_alive; simpl in *. - simpl in Hround. - destruct based; now simpl in *. - - changeR2. - left. - rewrite Hconfig in Hround. - unfold get_alive; simpl in *. - simpl in Hround. - destruct based; now simpl in *. -Qed. - - - - -Lemma moving_means_light_off : forall conf g (da:da_ila), - da_predicat da -> - get_based (conf (Good g)) = false -> - get_alive (round rbg_ila da conf (Good g)) = true -> - get_location (conf (Good g)) <> get_location (round rbg_ila da conf (Good g)) -> - get_light (round rbg_ila da conf (Good g)) = false. -Proof. - intros conf g da Hpred Hbased Halive Hmoving. - rewrite (round_good_g g _ Hpred) in *. - unfold rbg_ila, rbg_fnc in *. - - Set Printing Implicit. - changeR2. - set (s :=(obs_from_config _ _ )). - destruct ( move_to s (choose_new_pos s (fst (choose_target s)))) eqn : Hmove. - Unset Printing Implicit. - simpl in *. - fold s. - fold s in Hmoving, Halive. - rewrite Hmove in *. - simpl in *. - destruct (conf (Good g)) as (pos,(((ident, light), alive), based))eqn : Hconf. - simpl in *. - changeR2. - repeat cbn in *. - rewrite Hconf in *; simpl in *. - unfold get_alive in Halive; simpl in Halive. - assert (ifsplit_al : forall {A B C D E:Type} (a:bool) (b c:(D*(((B*C)*A)*E))) (d: A), snd (fst (snd(if a then b else c))) = d -> ((a= true -> snd (fst (snd b)) = d) /\ (a = false -> snd (fst (snd c)) = d))). - intros; destruct H, a; intuition. - rewrite Hbased in *. - apply ifsplit_al in Halive. - destruct Halive as (Hfalse, Halive). - unfold map_config in *. - destruct ((upt_robot_dies_b - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (let (p, b) := - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g in - let (r, t) := p in - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (if b - then reflexion - else - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)))))) - (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id))) g)) eqn : Heq. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - unfold upt_robot_too_far in *. - unfold get_light; simpl. - specialize (Hfalse (reflexivity true)). - simpl in Hfalse. - discriminate. - specialize (Halive (reflexivity false)). - destruct (change_frame da conf g) as ((r,t),b); simpl. - now unfold get_light; simpl. - destruct Hmoving. - fold s. - fold s in Halive. - simpl in *. - rewrite Hmove in *. - destruct Hpred as (Horigin,_). - unfold change_frame_origin in *. - unfold get_alive in Halive. - simpl in Halive. - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - Set Printing Implicit. - destruct ((@change_frame (R2 * ILA) Loc State_ILA Identifiers - (R2 * light) (R * R2 * bool) bool bool robot_choice_RL Frame - Choice inactive_choice_ila da conf g)) as ((r,t),b) eqn : Hchange. - unfold frame_choice_bijection, Frame in *. - Unset Printing Implicit. - unfold upt_aux in *. - destruct (conf (Good g)) as (pos,(((ident, light), alive), based)) eqn : Hconf. - replace (fst (pos, (ident, light, alive, based))) with pos by auto. - unfold Datatypes.id. - Set Printing All. - destruct (upt_robot_dies_b - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (R2.bij_rotation_f r - (@pair R R - (Rplus - (@fst R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end) (@fst R2 ILA (conf id)))) - (@fst R R t)) - (Rplus - (@snd R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end) (@fst R2 ILA (conf id)))) - (@snd R R t)))) (@snd R2 ILA (conf id))) g) eqn : Heq. - Unset Printing All. - destruct b. - simpl in *. - unfold map_config in *. - rewrite Heq, Hconf in Halive. - rewrite Hconf in *. - simpl in *. - unfold get_based in *; simpl in Hbased; rewrite Hbased in *. - simpl in *. - discriminate. - unfold map_config in *. - rewrite Hconf. - unfold get_based in *; simpl in Hbased; rewrite Hbased in *. - assert (Hsimplify : (fst - (let (p, b) := snd (pos, (ident, light, alive, false)) in - let (p0, a) := p in - let (i, l) := p0 in - if b - then - if - upt_robot_too_far - (fun id : Identifiers.ident => - (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) (fst (conf id)), - snd (conf id))) g (frame_choice_bijection (r,t,false)) - then - (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) - (fst (pos, (ident, light, alive, false))), (i, false, true, false)) - else - (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) - (fst (pos, (ident, light, alive, false))), snd (pos, (ident, light, alive, false))) - else - if - upt_robot_dies_b - (fun id : Identifiers.ident => - (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) (fst (conf id)), - snd (conf id))) g - then - (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) - (fst (pos, (ident, light, alive, false))), (i, l, false, b)) - else (fst (0%R, 0%R, true), (i, snd (0%R, 0%R, true), a, b)))) = - fst ( (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) - (fst (pos, (ident, light, alive, false))), (ident, light, false, false)))). - { simpl in *. rewrite Heq. - now simpl. - } - unfold ILA in * . - unfold NoCollisionAndPath.based, NoCollisionAndPath.light, NoCollisionAndPath.alive in *. - unfold R2 in *. - rewrite Hsimplify. - replace ((fst - ((comp (bij_rotation r) (comp (bij_translation t) Similarity.id)) - (fst (pos, (ident, light, alive, false))), (ident, light, false, false)))) with - ((comp (bij_rotation r) (comp (bij_translation t) Similarity.id)) - (fst (pos, (ident, light, alive, false)))) by auto. - now rewrite retraction_section. - rewrite <- Horigin in *. - assert (upt_robot_dies_b - (fun id : Identifiers.ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + - snd t)%R), snd (conf id))) g == upt_robot_dies_b - (fun id : Identifiers.ident => - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) g). - f_equiv. - intros id; split; simpl; auto. - destruct b; auto. - destruct (conf id) as ((?,?),?); simpl; auto. - now destruct i, p, p. - changeR2. - unfold map_config. - rewrite Hconf. - unfold get_based in Hbased; simpl in Hbased. - rewrite Hbased. - unfold ILA in * . - unfold NoCollisionAndPath.based, NoCollisionAndPath.light, NoCollisionAndPath.alive in *. - unfold R2 in *. - unfold Loc, State_ILA, AddInfo, OnlyLocation in *. simpl in *. - unfold R2 in *. - rewrite Heq. - simpl. - assert (Hsec_ret := retraction_section (frame_choice_bijection (r,t,b))). - specialize (Hsec_ret pos). - rewrite <- Hsec_ret at 1. - now simpl. -Qed. - - - -Definition no_collision_conf (conf : config) := - forall g g', - g <> g' -> - get_based (conf (Good g)) = false -> get_based (conf (Good g')) = false -> - get_alive (conf (Good g)) = true -> get_alive (conf (Good g')) = true -> - dist (get_location (conf (Good g))) (get_location (conf (Good g'))) <> 0%R. - -Definition no_collision (e:(@execution (R2*ILA) Loc State_ILA _ )):= - forall (config:config) (demon:demon_ila_type), - e = @execute (R2*ILA) Loc State_ILA _ _ _ _ _ _ robot_choice_RL Frame _ _ UpdFun _ rbg_ila demon config -> - @Stream.forever NoCollisionAndPath.config (Stream.instant no_collision_conf) e. - - -Definition path_conf (conf:config) := - forall g, get_alive (conf (Good g)) = true -> - (get_ident (conf (Good g)) = 0 \/ - (exists g', get_alive (conf (Good g')) = true /\ - (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) <= Dmax)%R /\ - get_ident (conf (Good g')) < get_ident (conf (Good g)) /\ - get_based (conf (Good g')) = false)). - - -Definition visibility_at_base (conf: config) := - exists g, get_based (conf (Good g)) = false /\ - get_alive (conf (Good g)) = true /\ - (dist (get_location (conf (Good g))) (0,0)%R <= Dp-D)%R. - - - -Definition visibility_based (e:execution) := - forall (config:config) (demon:demon_ila_type), - e = @execute (R2*ILA) Loc State_ILA _ _ _ _ _ _ robot_choice_RL Frame _ _ UpdFun _ rbg_ila demon config -> - Stream.forever (Stream.instant visibility_at_base) e. - - -Definition exists_path (e:execution) := - forall (config:config) (demon:demon_ila_type), - e = @execute (R2*ILA) Loc State_ILA _ _ _ _ _ _ robot_choice_RL Frame _ _ UpdFun _ rbg_ila demon config -> - Stream.forever (Stream.instant path_conf) e. - - -Lemma base_visibility_init : forall config_init, config_init_pred config_init -> - visibility_at_base config_init. -Proof. - intros. - unfold visibility_at_base. - unfold config_init_pred, config_init_base_linked in H. intuition. -Qed. - - -Lemma no_collision_init : forall config_init, config_init_pred config_init -> - no_collision_conf config_init. -Proof. - intros config_init Hinit; destruct Hinit as (lower, (off, (close, visib))). - unfold config_init_not_close in close. - unfold no_collision_conf. - intros g g' Hg Hbased Hbased' Halive Halive'. - specialize (close (Good g)); destruct (config_init (Good g)) as (?,(((?,?),?),?)). - unfold get_based in *; destruct b eqn : Hb; simpl in Hbased; try discriminate. - specialize (close (reflexivity false)). - specialize (close (Good g')). destruct (config_init (Good g')) as (?,(((?,?),?),?)). - unfold get_based in *; destruct b0 eqn : Hb0; simpl in Hbased'; try discriminate. - simpl in *. - unfold Datatypes.id. - set (dist := sqrt - ((@fst R R r + - @fst R R r0) * (@fst R R r + - @fst R R r0) + - (@snd R R r + - @snd R R r0) * (@snd R R r + - @snd R R r0))). - fold dist in close. - assert (D_0 := D_0). - specialize (close (reflexivity _)). - lra. -Qed. - -Lemma paht_conf_init : forall config_init, config_init_pred config_init -> - path_conf config_init. -Proof. - intros conf_init Hconf_pred. - unfold path_conf. - intros g Halive. - unfold config_init_pred in Hconf_pred. - unfold config_init_not_close, config_init_lower, config_init_off in Hconf_pred. - destruct Hconf_pred as (Hconf_lower, (Hconf_off, (Hconf_not_close, (Hvisib, (Hbased_higher, Hbased_0))))). - specialize (Hconf_lower (Good g)). - destruct (conf_init (Good g)) as (pos,(((ident,light), alive), based)) eqn : Hconf_init. - unfold get_alive in *; simpl in Halive. - specialize (Hconf_lower Halive). - destruct Hconf_lower as [Hconf_lower_0|Hconf_lower_other]. - left. - unfold get_ident; simpl in *; auto. - right. - - (* si on est based, un non based a toujours un id plus petit *) - destruct Hconf_lower_other as ([g_other|byz], Hother). - destruct (conf_init (Good g_other)) as (p_other,(((i_other, l_other), a_other), b_other)) eqn : Hconf_other. - unfold config_init_based_higher_id in Hbased_higher. - destruct (b_other) eqn : Hb_other. - - destruct based eqn : Hbased. - + unfold config_init_base_linked in *. - destruct Hvisib as (g', (Hbf, (Ha, Hd))). - exists g'. - repeat split. - unfold get_alive in *; simpl in *; auto. - specialize (Hbased_0 g). - rewrite Hconf_init in Hbased_0; unfold get_based in Hbased_0; simpl in Hbased_0; specialize (Hbased_0 (reflexivity _)). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location in *. - simpl (fst _). - unfold Datatypes.id in *. - destruct Hbased_0 as (Halive_0, Hbased_0). - rewrite Hbased_0. - unfold Dp in Hd. - generalize Dmax_7D, D_0. - rewrite dist_sym; lra. - specialize (Hbased_higher g' g). - unfold get_based in *; rewrite Hconf_init in *; simpl (snd _) in *. - specialize (Hbased_higher Hbf (reflexivity _)). - auto. - auto. - + specialize (Hbased_higher g g_other). - unfold get_based in *; rewrite Hconf_init, Hconf_other in *; simpl (snd _) in *. - specialize (Hbased_higher (reflexivity _) (reflexivity _)). - unfold get_ident in *; simpl in *. - intuition. - - exists g_other. - rewrite Hconf_other. - unfold get_ident. - unfold get_location, State_ILA, AddInfo, OnlyLocation , get_location. - repeat split; try (now simpl in *). - simpl (fst _). - unfold Datatypes.id. - intuition. - - assert (Hfalse := In_Bnames byz). - now simpl in *. -Qed. - - -Definition based_higher conf := - (forall g, get_based (conf (Good g)) = true -> get_alive (conf (Good g)) = true /\ get_location (conf (Good g)) = (0,0))%R /\ - - (exists g, get_based (conf (Good g)) = false - /\ get_alive (conf (Good g)) = true - /\ (dist (get_location (conf (Good g))) (0,0)%R <= Dp - D)%R -) /\ (forall g g', get_based (conf (Good g)) = true -> - get_based (conf (Good g')) = false -> - get_ident (conf (Good g')) < get_ident (conf (Good g))). - - -Lemma still_alive_means_alive : - forall conf g (da:da_ila), - da_predicat da -> - based_higher conf -> - get_alive (round rbg_ila da conf (Good g)) = true -> - get_alive(conf (Good g)) = true. -Proof. - intros conf g da Hpred Hbased Halive_r. - rewrite round_good_g in Halive_r. - simpl in *. - unfold get_alive in *; simpl in *. - unfold upt_aux in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - simpl in *. - unfold map_config in *. - destruct (upt_robot_dies_b - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool - bool robot_choice_RL Frame Choice - inactive_choice_ila da conf g - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid - (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec - VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end)) - end - end (@fst R2 ILA (conf id))) - (@snd R2 ILA (conf id))) g); - rewrite Hconf in *; - simpl in *; destruct based; - try (destruct (upt_robot_too_far _ _)); - try discriminate; auto. - destruct Hbased as (Hap0, (Hex, Hhi)). - specialize (Hap0 g). - unfold get_based, get_alive in Hap0; rewrite Hconf in Hap0; simpl in Hap0; specialize (Hap0 (reflexivity _)). - intuition. - destruct Hbased as (Hap0, (Hex, Hhi)). - specialize (Hap0 g). - unfold get_based, get_alive in Hap0; rewrite Hconf in Hap0; simpl in Hap0; specialize (Hap0 (reflexivity _)). - intuition. - apply Hpred. -Qed. - - - -Lemma dist_round_max_d : forall g conf da, - da_predicat da -> - path_conf conf -> - get_alive (conf (Good g)) == true -> - Rle_bool (dist (get_location (conf (Good g))) - (get_location (round rbg_ila da conf (Good g)))) - D == true. -Proof. - intros g conf da Hpred Hpath Halive. - rewrite (round_good_g g conf Hpred). - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - simpl (get_location _). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - unfold upt_aux. - rewrite Hconf; simpl (upt_robot_dies_b _). - destruct (upt_robot_dies_b _) eqn : Hupt. - - unfold Datatypes.id. - unfold map_config. - rewrite Hconf. - simpl (fst _). - destruct b; - destruct based; try (destruct (upt_robot_too_far _ g ); - replace ((R2.bij_rotation_f r - ((fst (reflexion pos) + fst t)%R, (snd (reflexion pos) + snd t)%R))) - with ((comp (bij_rotation r) (comp (bij_translation t) reflexion)) pos) by (now simpl); - simpl (fst _); - replace (R2.bij_rotation_f r - ((fst (bij_reflexion_f pos) + fst t)%R, (snd (bij_reflexion_f pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) reflexion) pos) by (now simpl); - rewrite retraction_section; - rewrite R2_dist_defined_2; - generalize D_0; simpl; rewrite Rle_bool_true_iff; - lra); try destruct (upt_robot_too_far _ g); - try (replace ((R2.bij_rotation_f r - ((fst (reflexion pos) + fst t)%R, (snd (reflexion pos) + snd t)%R))) - with ((comp (bij_rotation r) (comp (bij_translation t) reflexion)) pos) by (now simpl); - simpl (fst _); - replace (R2.bij_rotation_f r - ((fst (bij_reflexion_f pos) + fst t)%R, (snd (bij_reflexion_f pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) reflexion) pos) by (now simpl); - rewrite retraction_section; - rewrite R2_dist_defined_2; - generalize D_0; simpl; rewrite Rle_bool_true_iff; - lra); try ( - replace ((R2.bij_rotation_f r - ((fst (reflexion pos) + fst t)%R, (snd (reflexion pos) + snd t)%R))) - with ((comp (bij_rotation r) (comp (bij_translation t) reflexion)) pos) by (now simpl); - simpl (fst _); - replace (R2.bij_rotation_f r - ((fst ( pos) + fst t)%R, (snd (pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) Similarity.id) pos) by (now simpl); - rewrite retraction_section; - rewrite R2_dist_defined_2; - generalize D_0; simpl; rewrite Rle_bool_true_iff; - lra). - - unfold rbg_fnc. - destruct (move_to _) eqn : Hmove. - set (new_pos := choose_new_pos _ _ ). - set (obs := obs_from_config _ _ ) in *. - assert (Hcorrect := @choose_new_pos_correct obs (fst - (choose_target obs)) new_pos). - destruct Hcorrect. - unfold new_pos. reflexivity. - fold new_pos. - simpl (fst _). - rewrite <- Rle_bool_true_iff in H0. - destruct based; rewrite Hconf; - try (destruct (upt_robot_too_far _ _)); destruct b. - + simpl (fst _). - replace (R2.bij_rotation_f r - ((fst (bij_reflexion_f pos) + fst t)%R, (snd (bij_reflexion_f pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) reflexion) pos) by (now simpl); rewrite retraction_section, dist_same. - simpl; rewrite Rle_bool_true_iff. - generalize D_0; lra. - + simpl (fst _). - replace (R2.bij_rotation_f r - ((fst (pos) + fst t)%R, (snd ( pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) Similarity.id ) pos) by (now simpl); rewrite retraction_section, dist_same. - simpl; rewrite Rle_bool_true_iff. - generalize D_0; lra. - + simpl (fst _). - replace (R2.bij_rotation_f r - ((fst (bij_reflexion_f pos) + fst t)%R, (snd (bij_reflexion_f pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) reflexion) pos) by (now simpl); rewrite retraction_section, dist_same. - simpl; rewrite Rle_bool_true_iff. - generalize D_0; lra. - + simpl (fst _). - replace (R2.bij_rotation_f r - ((fst (pos) + fst t)%R, (snd ( pos) + snd t)%R)) with (comp (bij_rotation r) (comp (bij_translation t) Similarity.id ) pos) by (now simpl); rewrite retraction_section, dist_same. - simpl; rewrite Rle_bool_true_iff. - generalize D_0; lra. - + simpl (fst _). - rewrite <- H0. - f_equiv. - rewrite (frame_dist _ _ (r,t,true)), dist_sym. - destruct Hpred as (Hori, _). - revert Hori; intro Hori. - unfold change_frame_origin in *. - specialize (Hori conf g (r,t,true) Hchange). - unfold frame_choice_bijection, Frame in *. - simpl (_ ∘ _) in *. - unfold Datatypes.id. - rewrite section_retraction. - rewrite Hconf in Hori. - simpl (get_location _) in Hori. - unfold Datatypes.id in *. - rewrite Hori. - unfold new_pos. - reflexivity. - + simpl (fst _). - rewrite <- H0. - f_equiv. - rewrite (frame_dist _ _ (r,t,false)), dist_sym. - destruct Hpred as (Hori, _). - revert Hori; intro Hori. - unfold change_frame_origin in *. - specialize (Hori conf g (r,t,false) Hchange). - unfold frame_choice_bijection, Frame in *. - simpl (_ ∘ _) in *. - unfold Datatypes.id. - rewrite section_retraction. - rewrite Hconf in Hori. - simpl (get_location _) in Hori. - unfold Datatypes.id in *. - rewrite Hori. - unfold new_pos. - reflexivity. - + rewrite <- H0. - f_equiv. - rewrite (frame_dist _ _ (r,t,true)), dist_sym. - destruct Hpred as (Hori, _). - revert Hori; intro Hori. - unfold change_frame_origin in *. - specialize (Hori conf g (r,t,true) Hchange). - unfold frame_choice_bijection, Frame in *. - simpl (_ ∘ _) in *. - unfold Datatypes.id. - rewrite section_retraction. - rewrite Hconf in Hori. - simpl (get_location _) in Hori. - unfold Datatypes.id in *. - rewrite Hori. - unfold new_pos. - reflexivity. - + destruct Hpred as (Hori, _). - revert Hori; intro Hori. - rewrite <- H0. - unfold change_frame_origin in *. - specialize (Hori conf g (r,t, false) Hchange). - unfold frame_choice_bijection, Frame in *. - simpl (_ ∘ _) in *. - unfold map_config. - simpl (fst _). - rewrite (frame_dist pos _ (r,t,false)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - simpl (_ ∘ _) in *. - - rewrite <- Hori, Hconf. - rewrite section_retraction, dist_sym. - simpl; reflexivity. - + unfold map_config. - simpl (get_location _); unfold Datatypes.id. - destruct based; rewrite Hconf; - try (destruct (upt_robot_too_far _ _)); simpl (fst _); try rewrite dist_same; - replace (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) pos) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) pos) + snd t)%R)) - with (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos) by (now destruct b; simpl); try (destruct b; rewrite retraction_section, dist_same); - try (generalize D_0; simpl equiv; rewrite Rle_bool_true_iff; lra); - destruct Hpred as (Hori, _); - revert Hori; intro Hori; - unfold change_frame_origin in *; - specialize (Hori conf g (r,t, b) Hchange); - unfold frame_choice_bijection, Frame in *; - simpl (_ ∘ _) in *; - unfold map_config; - simpl (fst _); - rewrite (frame_dist pos _ (r,t,b)); - unfold frame_choice_bijection, Frame, Datatypes.id in *; - simpl (_ ∘ _) in *; - - rewrite <- Hori, Hconf; - rewrite section_retraction, dist_sym, dist_same; generalize (D_0); simpl equiv; rewrite Rle_bool_true_iff; lra. -Qed. - - -Section DecidableProperty. - -Context (Robots : Names). - -Variable P : @ident Identifiers -> Prop. -Hypothesis Pdec : forall id, {P id} + {~P id}. -Hypothesis exists_one : exists id, P id. - -Definition Pbool x := if Pdec x then true else false. - -Lemma Pbool_spec : forall x, Pbool x = true <-> P x. -Proof using. clear. intro. unfold Pbool. destruct_match; firstorder. discriminate. Qed. - -Corollary Pbool_false : forall x, Pbool x = false <-> ~P x. -Proof using . intro. rewrite <- Pbool_spec. now destruct (Pbool x). Qed. - -Definition ident_ltb (config:config) (id1 id2: @ident Identifiers) := Nat.ltb (get_ident (config id1)) (get_ident (config id2)). -Definition ident_leb (config:config) id1 id2 := if names_eq_dec id1 id2 then true else ident_ltb config id1 id2. - -(* Orders in Prop *) -Definition ident_lt config id1 id2 := ident_ltb config id1 id2 = true. -Definition ident_le config id1 id2 := ident_leb config id1 id2 = true. - -Lemma ident_ltb_trans : forall config id1 id2 id3, - ident_ltb config id1 id2 = true -> ident_ltb config id2 id3 = true -> ident_ltb config id1 id3 = true. -Proof using Robots. -unfold ident_ltb in *. -induction names; intros config id1 id2 id3; simpl. -+ unfold get_ident in *. rewrite 3 Nat.ltb_lt. lia. -+ repeat destruct_match; tauto || discriminate || eauto. -Qed. - -Lemma ident_ltb_irrefl : forall config id, ident_ltb config id id = false. -Proof using . intros config id. unfold ident_ltb. - rewrite Nat.ltb_nlt. lia. Qed. - -Instance ident_lt_strict config : StrictOrder (ident_lt config). -Proof using Robots. unfold ident_lt. split. -+ intro. simpl. now rewrite ident_ltb_irrefl. -+ repeat intro. eauto using ident_ltb_trans. -Qed. - -Instance ident_le_preorder config : PreOrder (ident_le config). -Proof using Robots. unfold ident_le, ident_leb. split. -+ intro. now destruct_match. -+ intros x y z. repeat destruct_match; subst; auto; []. apply ident_ltb_trans. -Qed. - -Lemma ident_le_linear : forall conf id id', {ident_le conf id id'} + {ident_le conf id' id}. -Proof using . - intros. - unfold ident_le, ident_leb. - do 2 destruct (names_eq_dec _ _); auto. - unfold ident_ltb. - simpl. - destruct (get_ident (conf id) <? get_ident (conf id')) eqn : H1; auto. - rewrite Nat.ltb_nlt in H1. - destruct (get_ident (conf id') <? get_ident (conf id)) eqn : H2; auto. - rewrite Nat.ltb_ge in H2. - destruct H1. - destruct (Nat.eq_dec (get_ident (conf id)) (get_ident (conf id'))). - destruct id as [g|byz], id' as [g'|byz']; - try (assert (b_In := In_Bnames byz); - now simpl in b_In); - try (assert (b_In := In_Bnames byz'); - now simpl in b_In). - destruct (Geq_dec g g'). - now rewrite e0 in *. - now assert (Hfalse := ident_unique conf n2). - lia. -Qed. - -Lemma ident_le_dec : forall conf id id', {ident_le conf id id'} + {~ident_le conf id id'}. -Proof using . - intros. - unfold ident_le, ident_leb. - destruct (names_eq_dec _ _); try auto. - unfold ident_ltb. - destruct (get_ident (conf id) <? get_ident (conf id')) eqn : H1; auto. -Qed. - - -Theorem exists_min : forall conf, exists id, P id /\ forall id', P id' -> ident_le conf id id'. -Proof using Pdec Robots exists_one. -pose (P' := fun id => List.In id names /\ P id). -assert (Hequiv : forall id, P id <-> P' id). -{ intros id. unfold P'. assert (Hin := In_names id). tauto. } -revert exists_one. setoid_rewrite Hequiv. -unfold P'. clear P' Hequiv. -generalize (@names_NoDup Identifiers). -induction names as [| id l]; simpl; [firstorder |]. -intros Hnodup Hex conf. inversion_clear Hnodup. -destruct (Pdec id) as [Hid | Hid]. -* (* the first element id satisfies P *) - destruct (List.existsb Pbool l) eqn:HPl. - + (* there exists other elements satisfying P in l, so a minimum id_min *) - rewrite List.existsb_exists in HPl. setoid_rewrite Pbool_spec in HPl. - specialize (IHl H0 HPl conf). destruct IHl as [id_min [[Hinl HPmin] Hmin]]. - - destruct (ident_le_dec conf id id_min). - - (* id is smaller than id_min *) - exists id. split; auto; []. - intros id' [[| Hid'] HPid']; subst. - -- reflexivity. - -- transitivity id_min; trivial; []. now apply Hmin. - - (* id_min is smaller than id *) - exists id_min. repeat split; auto; []. - intros id' [[Hid' | Hid'] HPid']; subst. - -- now destruct (ident_le_linear conf id_min id'). - -- now apply Hmin. - + (* there is no other element satisfying P *) - exists id. split; auto; []. - intros id' [[Hid' | Hid'] HPid']; subst; try reflexivity; []. - rewrite <- Bool.negb_true_iff, ListComplements.forallb_existsb, List.forallb_forall in HPl. - apply HPl in Hid'. now rewrite Bool.negb_true_iff, Pbool_false in Hid'. -* (* the first element id does not satisfy P *) - destruct Hex as [? [[] ?]]; subst; try tauto; []. - assert (exists id, List.In id l /\ P id). - exists x. repeat split; auto. - destruct (IHl H0 H3 conf) as [id_min [[Hinl HPmin] Hmin]]. - exists id_min. repeat split; auto; []. - intros. apply Hmin. intuition. now subst. -Qed. - - -Corollary robin : forall conf, exists id, P(id) /\ forall id', id <> id' -> P(id') -> get_ident (conf id) < get_ident (conf id'). -Proof using Pdec Robots exists_one. - intro conf. - destruct (exists_min conf) as [id_min [? Hmin]]. -exists id_min. split; trivial; []. -intros id' Hneq Hid'. apply Hmin in Hid'. clear Hmin. -unfold ident_le, ident_lt, ident_leb in *. -revert Hid'. destruct_match; try now auto. -unfold ident_ltb. -rewrite Nat.ltb_lt; auto. -Qed. - -End DecidableProperty. - -Lemma moving_means_not_near : forall conf (da:da_ila) g, - da_predicat da -> - no_collision_conf conf -> - path_conf conf -> - based_higher conf -> - get_location (conf (Good g)) <> get_location (round rbg_ila da conf (Good g)) - -> get_alive (round rbg_ila da conf (Good g)) = true - -> ((dist (get_location (conf (Good g))) - (get_location (round rbg_ila da conf (Good g))) - <= D)%R - /\ (forall g', get_ident (conf (Good g')) < get_ident (conf (Good g)) - -> get_alive (conf (Good g')) = true -> get_based (conf (Good g')) = false -> - (Rle_bool (dist (get_location (conf (Good g'))) (get_location ((round rbg_ila da conf (Good g))))) (2*D) == false)%R) - /\ (exists g', get_ident (conf (Good g')) < get_ident (conf (Good g)) /\ - get_alive (conf (Good g')) = true /\ get_based (conf (Good g')) = false /\ - ((Rle_bool (dist (get_location (conf (Good g'))) (get_location (round rbg_ila da conf (Good g)))) Dp) == true)%R)). -Proof. - intros conf da g Hpred Hcoll Hpath Hhigher Hmove Halive. - destruct (round rbg_ila da conf (Good g)) as (p_round, (((i_round, l_round), a_round), b_round)) eqn : Hconf_round. - assert (Hconf_round' : round rbg_ila da conf (Good g) == (p_round, (((i_round, l_round), a_round), b_round))) by now rewrite Hconf_round. - rewrite <- Hconf_round in *. - rewrite round_good_g in *; try apply Hpred. - simpl (get_location _) in *. - simpl in Halive, Hconf_round'. - unfold upt_aux, rbg_fnc, Datatypes.id in *. - unfold map_config in *. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - destruct (conf (Good g)) as (pos, (((i,l),a), b')) eqn : Hconf. - simpl (fst _) in *. - set (Hdies := upt_robot_dies_b - (fun id : @ident Identifiers => - @pair R2 ILA - (R2.bij_rotation_f r - (@pair R R - (Rplus - (@fst R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec - VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end) (@fst R2 ILA (conf id)))) - (@fst R R t)) - (Rplus - (@snd R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec - VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end) (@fst R2 ILA (conf id)))) - (@snd R R t)))) (@snd R2 ILA (conf id))) g). - repeat fold Hdies in *. - unfold get_alive in Halive. - simpl (snd (let (p,b0) := snd (pos, (i,l,a,b')) in _)) in Halive, Hconf_round'. - destruct b' eqn : Hbased. - - simpl in Halive. - destruct (upt_robot_too_far _ _) eqn : Htoo_far; simpl in Halive; try discriminate. - unfold upt_robot_too_far in Htoo_far. - destruct (R2dec_bool _ _) eqn : Hpos_conf; try (simpl in Htoo_far; discriminate). - rewrite R2dec_bool_true_iff in Hpos_conf. - destruct (forallb (too_far_aux_1 _ _)) eqn : Htoo_far_aux1; try (simpl in Htoo_far; discriminate). - destruct (forallb (too_far_aux_2 _ _)) eqn : Htoo_far_aux2; try (simpl in Htoo_far; discriminate). - rewrite Hconf in *. - simpl (fst (pos, _)) in *. - replace (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) pos) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) pos) + snd t)%R)) - with (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos) - in * by now destruct b; simpl. - simpl in Hpos_conf. - simpl (fst _). - repeat split. - + assert (Hsec_ret := retraction_section (frame_choice_bijection (r,t,b))). - specialize (Hsec_ret pos). - rewrite <- Hsec_ret at 1. - unfold frame_choice_bijection, Frame. simpl (retraction _ _). - destruct b; rewrite dist_same; generalize D_0; lra. - + intros g' Hident' Halive' Hbased'. - unfold too_far_aux_2 in *. - rewrite forallb_forall in *. - specialize (Htoo_far_aux2 g' (In_Gnames g')). - rewrite negb_true_iff, 2 andb_false_iff in Htoo_far_aux2. - destruct Htoo_far_aux2. - unfold get_based in *; simpl in *; rewrite negb_false_iff, H in *; discriminate. - destruct H. - unfold get_alive in *; simpl in *. - now rewrite H in *. - unfold equiv, bool_Setoid, eq_setoid. - rewrite Rle_bool_false_iff in *. - intro. - destruct H. - assert (Hframe_dist := frame_dist (get_location (conf (Good g'))) (get_location (conf (Good g))) ((r,t,b))). - replace (dist - (get_location - (R2.bij_rotation_f r - (fst ((if b then reflexion else Similarity.id) (fst (conf (Good g')))) + fst t, - snd ((if b then reflexion else Similarity.id) (fst (conf (Good g')))) + snd t), - snd (conf (Good g')))) - (get_location - (R2.bij_rotation_f r - (fst ((if b then reflexion else Similarity.id) (fst (conf (Good g)))) + fst t, - snd ((if b then reflexion else Similarity.id) (fst (conf (Good g)))) + snd t), - snd (conf (Good g)))))%R - with (dist (get_location (conf (Good g'))) (get_location (conf (Good g)))). - assert (Hdist_r := @dist_round_max_d g conf da Hpred Hpath). - destruct Hhigher as (Ha0, (Hex, Hhi)). - specialize (Ha0 g). - rewrite Hconf in Ha0. - unfold get_based, get_alive in Ha0; simpl in Ha0; specialize (Ha0 (reflexivity _)). - destruct Ha0 as (Ha, Hp0). - unfold get_alive in Hdist_r; rewrite Hconf in Hdist_r; simpl (snd _) in Hdist_r. - specialize (Hdist_r Ha). - assert (Htri := RealMetricSpace.triang_ineq (get_location (conf (Good g'))) (fst (round rbg_ila da conf (Good g))) (get_location (conf (Good g)))) . - transitivity (dist (get_location (conf (Good g'))) (fst (round rbg_ila da conf (Good g))) + - dist (fst (round rbg_ila da conf (Good g))) (get_location (conf (Good g))))%R; try auto. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - transitivity (2*D + dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g))))%R. - apply Rplus_le_compat_r. - lra. - transitivity (2*D + D)%R. - apply Rplus_le_compat_l. - unfold equiv, bool_Setoid, eq_setoid in Hdist_r. - rewrite Rle_bool_true_iff, dist_sym, <- Hconf in Hdist_r. - apply Hdist_r. - unfold Dp; generalize Dmax_7D, D_0. - lra. - + replace ((fst - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos, - (i, l, a, false)))) - with ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos)) - in Hmove by auto. - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g)))). - unfold frame_choice_bijection, Frame in Hret_sec. - rewrite Hconf in Hret_sec. - simpl (_ ∘ _) in Hret_sec. - simpl (fst _) in Hret_sec. - rewrite <- Hret_sec in Hmove at 1. - destruct Hmove. - destruct b; simpl;auto. - + replace ((R2.bij_rotation_f r - (fst ((if b then reflexion else Similarity.id) pos) + fst t, - snd ((if b then reflexion else Similarity.id) pos) + snd t)))%R - with ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos)) in Hmove. - replace ((fst - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos, - (i, l, a, false)))) - with ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)) pos)) - in Hmove by auto. - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g)))). - unfold frame_choice_bijection, Frame in Hret_sec. - rewrite Hconf in Hret_sec. - simpl (_ ∘ _) in Hret_sec. - simpl (fst _) in Hret_sec. - rewrite <- Hret_sec in Hmove at 1. - destruct Hmove. - destruct b; simpl;auto. - destruct b; simpl;auto. - - destruct Hdies eqn : Hd. - simpl in Halive. - unfold Hdies in Hd. - rewrite Hd in Halive. - simpl in Halive. - discriminate. - simpl in Halive; unfold Hdies in Hd; rewrite Hd in Halive. - simpl in Halive. - destruct (move_to _) eqn : Hmove_to. - + set (new_pos := choose_new_pos _ _). - assert (Hchoose_new_pos_dist := @choose_new_pos_correct _ _ (new_pos) (reflexivity _)). - repeat split. - * - - rewrite dist_sym; - destruct Hchoose_new_pos_dist as (_,Hchoose); - - rewrite (frame_dist _ pos ((r,t),b)) in *. - destruct Hpred. - specialize (H conf g ((r,t),b) Hchange). - rewrite Hconf in H. - simpl (let bij := frame_choice_bijection (r, t, b) in _ ) in H. - unfold frame_choice_bijection, Frame. - simpl (rotation _ ∘ _). - rewrite section_retraction. - simpl (((Similarity.comp (rotation r) - (Similarity.comp (translation t) (if b then reflexion else Similarity.id))) - pos)). - revert H. - destruct b; - unfold Datatypes.id; - intros H; rewrite H; - simpl (fst _); - apply Hchoose. - * assert (Hmove_some := move_to_Some_zone Hmove_to ). - fold new_pos in Hmove_some. - unfold equiv, bool_Setoid, bool_eqdec, eq_setoid. - intros g' Hident' Halive_g' Hbased'. - rewrite Rle_bool_false_iff. - apply Rgt_not_le. - destruct (Rle_lt_dec (dist (fst (conf (Good g'))) pos) Dmax). - assert ((((fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + snd t)%R), - snd (conf id))) (Good g')) - ∈ obs_from_config - (fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + snd t)%R), - snd (conf id))) - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) pos) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) pos) + snd t)%R), - (i, l, a, b')))%set). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA. - split. - rewrite config_list_InA. - exists (Good g'). - reflexivity. - rewrite andb_true_iff. - split. - -- simpl (fst _). - assert (Hframe_dist := frame_dist (fst (conf (Good g'))) pos ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe_dist. - simpl (rotation _ ∘ _) in Hframe_dist. - simpl. - simpl in *; rewrite <- Hframe_dist. - rewrite 2 andb_true_iff. repeat split; try rewrite Rle_bool_true_iff; auto. - -- unfold get_ident in *; simpl in *. - rewrite Nat.leb_le; lia. - -- intros x y Hxy. - destruct x as (rx,(((ix,lx),ax),bx)), y as (ry,(((iy,ly),ay),By)). - simpl in Hxy. - destruct Hxy as (Hr,(Hi,(Hl,(Ha, Hb)))). - simpl (fst _); simpl (snd _). - rewrite Hr,Hl,Hi,Ha, Hb. - reflexivity. - -- specialize (Hmove_some ((fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + snd t)%R), - snd (conf id))) (Good g'))). - specialize (Hmove_some H). - rewrite Hd in Hconf_round'. - simpl in Hconf_round'. - fold new_pos in Hconf_round'. - destruct Hconf_round' as (Hpos_round, _). - revert Hpos_round; intros Hpos_round. - clear H. - rewrite Hconf_round. - simpl (fst (_,_)). - rewrite <- Hpos_round. - rewrite (frame_dist _ _ ((r,t),b)). - unfold frame_choice_bijection, Frame. - simpl (_ ∘ _). - assert (Hsec_ret := section_retraction ((Similarity.comp (rotation r) - (Similarity.comp (translation t) (if b then reflexion else Similarity.id)))) new_pos). - simpl (retraction _) in Hsec_ret. - assert (((Similarity.comp (rotation r) - (Similarity.comp (translation t) (if b then reflexion else Similarity.id))) - (retraction (if b then reflexion else Similarity.id) - (Rgeom.xr (fst new_pos) (snd new_pos) (- r) + - fst t, - Rgeom.yr (fst new_pos) (snd new_pos) (- r) + - snd t))) == new_pos)%R. - destruct b; simpl in *; auto. - destruct (Geq_dec g g'). - +++ - rewrite <- e, Hconf in Hident'. lia. - +++ - destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - simpl in Horigin. - unfold Datatypes.id in Horigin. - - specialize (Hcoll g g' n0). rewrite Hconf in Hcoll; unfold get_based at 1 in Hcoll. - assert (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) <> 0)%R. - rewrite Hconf; apply Hcoll; - unfold get_alive; simpl in *; auto. - rewrite <- Horigin in Hmove_some. - destruct b; rewrite H; - rewrite dist_sym; apply Hmove_some. - rewrite (frame_dist _ _ (r,t,true)) in H0. - - assert (forall (x y : R2), dist x y <> 0%R -> dist x y > 0%R)%R. - intros. - generalize (dist_nonneg x y). lra. - apply H1. - rewrite dist_sym in H0. - simpl in *; apply H0. - rewrite (frame_dist _ _ (r,t,false)) in H0. - - assert (forall (x y : R2), dist x y <> 0%R -> dist x y > 0%R)%R. - intros. - generalize (dist_nonneg x y). lra. - apply H1. - rewrite dist_sym in H0. - simpl in *; apply H0. - -- assert (Hd_round' := dist_round_max_d g' Hpred Hpath Halive_g'). - unfold equiv, bool_Setoid, eq_setoid in Hd_round'. - rewrite Rle_bool_true_iff in Hd_round'. - unfold get_location, State_ILA, AddInfo, OnlyLocation in Hd_round'. - unfold get_location, State_ILA, AddInfo, OnlyLocation in Hd_round'. - unfold Datatypes.id in *. - - assert (Dmax - D < dist (fst (round rbg_ila da conf (Good g'))) pos)%R. - generalize Dmax_7D. - intros Dmax_7D. - assert (Htri := RealMetricSpace.triang_ineq (fst (conf (Good g'))) (fst (round rbg_ila da conf (Good g'))) pos). - assert (Htri' : (dist (fst (conf (Good g'))) (fst (round rbg_ila da conf (Good g'))) + - dist (fst (round rbg_ila da conf (Good g'))) pos <= D + dist (fst (round rbg_ila da conf (Good g'))) pos )%R). - apply Rplus_le_compat_r. - assumption. - lra. - assert (Hd_round := dist_round_max_d g Hpred Hpath). - unfold get_alive in Hd_round. - rewrite Hconf in *. - simpl (snd _) in Hd_round. - specialize (Hd_round Halive). - simpl (get_location _) in Hd_round. - unfold Datatypes.id in *. - unfold equiv, bool_Setoid, eq_setoid in Hd_round. - rewrite Rle_bool_true_iff in Hd_round. - set (d2_r := (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila - UpdFun inactive_ila rbg_ila da conf - (@Good Identifiers g)))) in *. - set (d1 := (fst (conf (Good g')))) in *. - set (d1_r := (@fst R2 ILA - (@round (R2 * ILA) Loc State_ILA Identifiers Obs_ILA - (R2 * light) (R * R2 * bool) bool bool robot_choice_RL Frame Choice - inactive_choice_ila UpdFun inactive_ila rbg_ila da conf - (@Good Identifiers g)))) in *. - unfold equiv, bool_Setoid, eq_setoid. - apply Rnot_le_gt. - intro. - generalize Dmax_7D. - intros Dmax_7D. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id, equiv, bool_Setoid, eq_setoid in Hd_round; simpl (fst _) in *. - assert (7*D < dist d1 pos)%R by lra. - assert (2*D < dist d1 d2_r)%R. - { assert (Htri := RealMetricSpace.triang_ineq d1 d2_r pos). - assert (Htri' : (3*D < dist d1 d2_r + - dist d2_r pos)%R). - generalize D_0. - lra. - apply (Rplus_lt_reg_l D). - rewrite dist_sym in Hd_round. - assert (3*D - dist d2_r pos < dist d1 d2_r)%R by lra. - assert (Hrlelt := Rplus_lt_le_compat (3 * D - dist d2_r pos) (dist d1 d2_r) (dist d2_r pos) D H2 Hd_round). - replace (3 * D - dist d2_r pos + dist d2_r pos)%R with (dist d2_r pos + (3*D - dist d2_r pos))%R in Hrlelt. - rewrite Rplus_minus in Hrlelt. - lra. - replace ((dist d2_r pos + (3 * D - dist d2_r pos))%R) with - (dist d2_r pos + 3 * D - dist d2_r pos)%R by lra. - assert (Hr : forall (x y z:R), (x + y - z = y - z + x)%R). - intros. - replace (x + y - z)%R with (x + (y - z))%R. - rewrite Rplus_comm. lra. - lra. - apply Hr. - } - apply (Rle_not_lt _ _ H0 H2). - * set ( obs := ( obs_from_config _ _ )) in *. - assert (Hin_obs := choose_target_in_obs (reflexivity (choose_target obs))). - assert (Hg: exists g0, (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf (Good g0)))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf (Good g0)))) + - snd t)%R), snd (conf (Good g0))) == choose_target obs). - { unfold obs, obs_from_config, Obs_ILA in Hin_obs. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff in Hin_obs. - destruct Hin_obs as ((id,Hin), (Hdist, Halive_g')). - destruct id as [g0|byz]. - unfold obs, obs_from_config, Obs_ILA. - exists g0. - symmetry. - apply Hin. - assert (b_In := In_Bnames byz). - now simpl in b_In. - intros x y Hxy. - simpl in Hxy. - destruct x as (rx,(((ix,lx),ax), bx)), - y as (ry,(((iy,ly),ay), By)), - Hxy as (Hr, (Hi, (Hl, (Ha, Hb)))). - simpl in Hr; - rewrite Hr, Hi, Hl, Ha, Hb. - reflexivity. - } - destruct Hg as (g', Hg). - exists g'. - assert (Hlower := @target_lower - obs (choose_target obs) - ((fun id : ident => - (R2.bij_rotation_f - r - ((fst ((if b then reflexion - else Similarity.id) (fst (conf id))) + - fst t)%R, - (snd ((if b then reflexion - else Similarity.id) (fst (conf id))) + - snd t)%R), snd (conf id))) - (Good g))). - assert (Heq : (fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + - snd t)%R), snd (conf id))) (Good g) == ((0,0)%R, snd (conf (Good g)))). - destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - simpl in Horigin. - unfold Datatypes.id in Horigin. - now rewrite Horigin. - assert (Hident_g'_g : get_ident (conf (Good g')) < get_ident (conf (Good g))). - {destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - simpl in Horigin. - unfold Datatypes.id in Horigin. - assert (Hin_l : ((fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + - snd t)%R), snd (conf id))) (Good g) ∈ obs)%set). - { rewrite Heq. - unfold obs, obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - - split. - -- exists (Good g). - now rewrite Horigin. - -- rewrite Hconf in Horigin. - rewrite <- Horigin. - simpl (fst _ ) in *. - rewrite dist_same. - repeat rewrite andb_true_iff; - split; intuition. - generalize Dmax_7D, D_0. - rewrite Rle_bool_true_iff. - lra. - unfold get_alive. - now rewrite Hconf; simpl. - rewrite Hconf in *. - unfold get_ident; simpl. - rewrite Nat.leb_le. - lia. - -- intros (?,(((?,?),?),?)) (?,(((?,?),?),?)) (Hr,(Hi,(Hl, (Ha, Hb)))). - simpl in *. - now rewrite Ha, Hl, Hr, Hi. } - specialize (Hlower Hin_l). - clear Hin_l. - assert ( get_location - ((fun id : ident => - (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf id))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf id))) + - snd t)%R), snd (conf id))) (Good g)) = - (0%R, 0%R)). - { - rewrite Horigin. - now simpl in *. - } - specialize (Hlower H (reflexivity _)); clear H. - simpl in Hlower. - assert (Hident_g : forall g0, get_ident (R2.bij_rotation_f r - ((fst ((if b then reflexion else Similarity.id) (fst (conf (Good g0)))) + - fst t)%R, - (snd ((if b then reflexion else Similarity.id) (fst (conf (Good g0)))) + - snd t)%R), snd (conf (Good g0))) == get_ident (conf (Good g0))). - unfold get_ident; now simpl. - rewrite <- 2 Hident_g. - rewrite Hg. - apply Hlower. } - repeat split. - ++ rewrite Hconf in Hident_g'_g. - apply Hident_g'_g. - ++ assert (Ht_alive := choose_target_alive (symmetry Hg)). - now unfold get_alive in *; simpl in *. - ++ destruct (get_based (conf (Good g'))) eqn : Hbased'. - destruct Hhigher as (Hap0,(Hexists_based, Hhigher)). - specialize (Hhigher g' g). unfold get_based in *; rewrite Hconf in *; simpl in Hhigher. - specialize (Hhigher Hbased' (reflexivity _)). - lia. - auto. - ++ assert (Hgood := @round_good_g g conf da Hpred). - rewrite Hconf_round. - simpl (fst (_,_)). - destruct Hconf_round' as (Hpos_round, _). - rewrite <- Hpos_round. - rewrite Hd. - rewrite (frame_dist _ _ ((r,t),b)). - unfold frame_choice_bijection, Frame. - simpl (_ ∘ _). - assert (Hsec_ret := section_retraction ((Similarity.comp (rotation r) - (Similarity.comp (translation t) (if b then reflexion else Similarity.id)))) new_pos). - simpl (retraction _) in Hsec_ret. - assert (((Similarity.comp (rotation r) - (Similarity.comp (translation t) (if b then reflexion else Similarity.id))) - (retraction (if b then reflexion else Similarity.id) - (Rgeom.xr (fst new_pos) (snd new_pos) (- r) + - fst t, - Rgeom.yr (fst new_pos) (snd new_pos) (- r) + - snd t))) == new_pos)%R. - destruct b; simpl in *; auto. - fold obs in new_pos. - fold new_pos. - destruct b; rewrite H; - rewrite dist_sym; - revert Hg; intro Hg; - destruct Hchoose_new_pos_dist as (Htarg, _); - revert Htarg; intro Htarg; - rewrite <- (fst_compat (Hg)) in Htarg; - unfold equiv, bool_Setoid, eq_setoid; - rewrite Rle_bool_true_iff; - simpl in *; lra. - + rewrite Hd in Hmove. - simpl ((fst (fst (0%R, 0%R, true), (i, snd (0%R, 0%R, true), a)))) in Hmove. - destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - destruct Hmove. - rewrite <- Horigin. - unfold frame_choice_bijection, Frame. - assert ((@retraction R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) - (@section (@location Loc) (@location_Setoid Loc) - (@sim_f R2 R2_Setoid R2_EqDec R2_VS - (@Normed2Metric R2 R2_Setoid R2_EqDec R2_VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec R2_VS R2_ES)) - (@compose - (@similarity R2 R2_Setoid R2_EqDec R2_VS - (@Normed2Metric R2 R2_Setoid R2_EqDec R2_VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec R2_VS R2_ES))) - (@similarity_Setoid R2 R2_Setoid R2_EqDec R2_VS - (@Normed2Metric R2 R2_Setoid R2_EqDec R2_VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec R2_VS R2_ES))) - (@SimilarityComposition R2 R2_Setoid R2_EqDec R2_VS - (@Normed2Metric R2 R2_Setoid R2_EqDec R2_VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec R2_VS R2_ES))) - (rotation r) - (@compose - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - (@similarity_Setoid R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - (@SimilarityComposition R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - (@translation R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES) t) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) - (@get_location Loc (prod R2 ILA) State_ILA (conf (@Good Identifiers g))))) - == - ((retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id)))) - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (get_location (conf (Good g)))))). - destruct b; unfold Similarity.id; simpl in *; auto. - rewrite H. - rewrite retraction_section. - rewrite Hconf; simpl. - auto. -Qed. - - - -(* d'autre propriété sur les conf global en utilisant les axioms de move_to mais les énoncé ne mentionnent pas move_to. *) - - -Lemma still_based_means_based : forall conf da g, - da_predicat da -> - get_based (round rbg_ila da conf (Good g)) = true-> - get_based (conf (Good g)) = true. -Proof. - intros conf da g Hpred Hbased_round. - rewrite round_good_g in Hbased_round; try apply Hpred. - simpl in *. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - unfold get_based in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - unfold upt_aux, map_config in *. - rewrite Hconf in *. - simpl in *. - destruct based; try auto. - destruct (upt_robot_dies_b _); try auto. -Qed. - - - -Lemma no_collision_cont : forall (conf:config) (da:da_ila), - no_collision_conf conf -> da_predicat da -> - path_conf conf -> based_higher conf -> - no_collision_conf (round rbg_ila da conf). -Proof. - intros conf da no_col Hpred Hpath Hbased_higher g g' Hg Hbased_r Hbased'_r Halive Halive'. - assert (Hcoll := no_col). - specialize (no_col g g' Hg). - destruct (get_based (conf (Good g))) eqn : Hbased, (get_based (conf (Good g'))) eqn : Hbased'. - + rewrite round_good_g in Hbased_r, Hbased'_r; try apply Hpred. - unfold get_based in *; simpl in Hbased_r, Hbased'_r. - unfold upt_aux, map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf; - destruct (conf (Good g')) as (pos', (((ident', light'), alive'), based')) eqn : Hconf'. - simpl in *. - rewrite Hbased, Hbased' in *. - destruct (upt_robot_too_far _ g) eqn : Hfar; - destruct (upt_robot_too_far _ g') eqn : Hfar'; try discriminate. - unfold upt_robot_too_far in *. - destruct (R2dec_bool _ _) eqn : Hpos ; try discriminate. - destruct (R2dec_bool (@retraction (@location (@make_Location R2 R2_Setoid R2_EqDec)) - (@location_Setoid (@make_Location R2 R2_Setoid R2_EqDec)) - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g' - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@get_location (@make_Location R2 R2_Setoid R2_EqDec) (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g' - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g')))))) (0%R, 0%R)) eqn : Hpos' ; try discriminate. - destruct (forallb (too_far_aux_1 _ g ) _) eqn : Htoo_far_1; - destruct (forallb (too_far_aux_1 _ g' ) _) eqn : Htoo_far_1'; try discriminate. - rewrite forallb_forall in *. - - unfold too_far_aux_1 in *. - specialize (Htoo_far_1 g'); specialize (Htoo_far_1' g). - changeR2. - assert ( - (@get_location - (@make_Location (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc)) - (prod (@location Loc) ILA) State_ILA - (@pair (@location Loc) ILA - (@section (@location Loc) (@location_Setoid Loc) - match - @change_frame (prod (@location Loc) ILA) Loc State_ILA Identifiers - (prod (@location Loc) NoCollisionAndPath.light) - (prod (prod R (@location Loc)) bool) bool bool robot_choice_RL Frame - Choice inactive_choice_ila da conf g' - return (@bijection (@location Loc) (@location_Setoid Loc)) - with - | pair p b => - match p return (@bijection (@location Loc) (@location_Setoid Loc)) with - | pair r t => - @comp (@location Loc) (@location_Setoid Loc) - (bij_rotation r) - (@comp (@location Loc) (@location_Setoid Loc) - (@bij_translation (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS t) - (@sim_f (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS ES)) - match - b - return - (@similarity (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) - (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) - (@location_EqDec Loc) VS ES))) - with - | true => reflexion - | false => - @Similarity.id (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) - (@location_EqDec Loc) VS ES)) - end)) - end - end (@fst (@location Loc) ILA (conf (@Good Identifiers g')))) - (@snd (@location Loc) ILA (conf (@Good Identifiers g'))))) - == (((let (p, b) := change_frame da conf g' in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g')))))). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - simpl; auto. - destruct (change_frame _ _ g') as ((?,?),?) eqn : Hchange'. - destruct b; auto. - rewrite H in Hpos'. - destruct (change_frame _ _ g') as ((?,?),?) eqn : Hchange'. - assert (Hdist_simp: R2dec_bool (fst (conf (Good g'))) (0,0)%R = true). - destruct b; rewrite retraction_section in Hpos'; - assumption. - unfold get_based, get_alive in Htoo_far_1. - simpl (snd (_)) in Htoo_far_1. - assert (Hin_g' : List.In g' (List.filter (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) - Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g'). - rewrite Hconf'. - simpl. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g'). - rewrite Hconf' in Hap0; unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - assert (Hin_g : List.In g (List.filter (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) - Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g). - rewrite Hconf. - simpl. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in Hap0; unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - specialize (Htoo_far_1 Hin_g'). - unfold get_based, get_alive in Htoo_far_1'; simpl (snd _) in Htoo_far_1'. - specialize (Htoo_far_1' Hin_g). - destruct (negb _) eqn : Hident_far_1 in Htoo_far_1; simpl in Htoo_far_1; try discriminate. - unfold get_ident in Htoo_far_1; simpl in Htoo_far_1. - destruct (negb _) eqn : Hident_far_1' in Htoo_far_1'; simpl in Htoo_far_1'; try discriminate. - unfold get_ident in Htoo_far_1'; simpl in Htoo_far_1'. - rewrite Nat.ltb_lt in *. - lia. - rewrite negb_true_iff , negb_false_iff in *. - rewrite Nat.eqb_eq, Nat.eqb_neq in *. - destruct Hident_far_1; unfold get_ident in *; simpl in *; auto. - destruct (ident_unique conf Hg). - rewrite negb_false_iff, Nat.eqb_eq in *; unfold get_ident in *; simpl in *; auto. - - - + rewrite round_good_g in Hbased_r, Hbased'_r; try apply Hpred. - unfold get_based in *; simpl in Hbased_r, Hbased'_r. - unfold upt_aux, map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf; - destruct (conf (Good g')) as (pos', (((ident', light'), alive'), based')) eqn : Hconf'. - simpl in *. - rewrite Hbased in *. - destruct (upt_robot_too_far _ _) eqn : Htoo_far; try discriminate. - unfold upt_robot_too_far in *. - destruct (R2dec_bool _ _) eqn : Hpos, (forallb (too_far_aux_1 _ _) _) eqn : Htoo_far_1; - try discriminate. - destruct (forallb _ _) eqn : Htoo_far_2 in Htoo_far; try discriminate. - unfold too_far_aux_2 in *. - rewrite forallb_forall in Htoo_far_2. - specialize (Htoo_far_2 g' (In_Gnames g')). - rewrite negb_true_iff, 2 andb_false_iff, negb_false_iff in *. - destruct (Htoo_far_2) as [Hbased_false| [Halive_false|Hdist_lower_Dp]]. - unfold get_based in *; simpl in Hbased_false. - rewrite Hconf' in *; simpl in *. - now rewrite Hbased_false in *. - apply still_alive_means_alive in Halive'. - unfold get_alive in *; simpl in *; rewrite Halive' in *; discriminate. - apply Hpred. - apply Hbased_higher. - rewrite Rle_bool_false_iff in *. - intro Hdist. - destruct Hdist_lower_Dp. - assert (Hdist_same := frame_dist (fst (conf (Good g'))) (fst (conf (Good g))) (change_frame da conf g)). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - simpl in *. - rewrite <- Hdist_same. - assert ((dist (fst (conf (Good g'))) (fst (conf (Good g))) <= Dp - 3*D)%R -> (sqrt - ((fst (fst (conf (Good g'))) + - fst (fst (conf (Good g)))) * - (fst (fst (conf (Good g'))) + - fst (fst (conf (Good g)))) + - (snd (fst (conf (Good g'))) + - snd (fst (conf (Good g)))) * - (snd (fst (conf (Good g'))) + - snd (fst (conf (Good g))))) <= Dp-3*D)%R) by now simpl. - apply H. - assert (dist (fst (round rbg_ila da conf (Good g))) (fst (round rbg_ila da conf (Good g'))) = 0%R) by now simpl. - assert (dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))) <= D)%R. - - { assert (Hdist_r:= @dist_round_max_d g' conf da Hpred Hpath (@still_alive_means_alive conf g' da Hpred Hbased_higher Halive')). - unfold equiv, bool_Setoid, eq_setoid in Hdist_r. - rewrite Rle_bool_true_iff in Hdist_r. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g))) (fst (round rbg_ila da conf (Good g'))) (fst (conf (Good g')))). - rewrite H0 in Htri. - transitivity (dist (fst (round rbg_ila da conf (Good g'))) (fst (conf (Good g')))); try lra. - rewrite dist_sym. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - apply Hdist_r. - } - assert (dist (fst ( conf (Good g))) (fst (conf (Good g'))) <= 2*D)%R. - { - assert (Hdist_r:= @dist_round_max_d g conf da Hpred Hpath (still_alive_means_alive g Hpred Hbased_higher Halive)). - unfold equiv, bool_Setoid, eq_setoid in Hdist_r. - rewrite Rle_bool_true_iff in Hdist_r. - assert (Htri := RealMetricSpace.triang_ineq (fst (conf (Good g))) (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g')))). - transitivity (dist (fst (conf (Good g))) (fst (round rbg_ila da conf (Good g))) + - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))))%R; try auto. - transitivity (D+ dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))))%R. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - apply Rplus_le_compat_r. - apply Hdist_r. - replace (2*D)%R with (D + D)%R. - apply Rplus_le_compat_l, H1. - generalize D_0; lra. - } - rewrite dist_sym. unfold Dp; generalize Dmax_7D, D_0; lra. - + rewrite round_good_g in Hbased_r, Hbased'_r; try apply Hpred. - unfold get_based in *; simpl in Hbased_r, Hbased'_r. - unfold upt_aux, map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf; - destruct (conf (Good g')) as (pos', (((ident', light'), alive'), based')) eqn : Hconf'. - simpl in *. - rewrite Hbased' in *. - clear Hbased_r. - destruct (upt_robot_too_far _ _) eqn : Htoo_far; try discriminate. - unfold upt_robot_too_far in *. - destruct (R2dec_bool _ _) eqn : Hpos; try discriminate. - destruct (forallb _ _) eqn : Htoo_far_1 in Htoo_far; - try discriminate. - destruct (forallb _ _) eqn : Htoo_far_2 in Htoo_far; try discriminate. - unfold too_far_aux_2 in *. - rewrite forallb_forall in Htoo_far_2. - specialize (Htoo_far_2 g (In_Gnames g)). - rewrite negb_true_iff, 2 andb_false_iff, negb_false_iff in *. - destruct (Htoo_far_2) as [Hbased_false| [Halive_false|Hdist_lower_Dp]]. - unfold get_based in *; simpl in Hbased_false. - rewrite Hconf in *; simpl in *. - now rewrite Hbased_false in *. - apply still_alive_means_alive in Halive. - unfold get_alive in *; simpl in *; rewrite Halive in *; discriminate. - apply Hpred. - apply Hbased_higher. - rewrite Rle_bool_false_iff in *. - intro Hdist. - destruct Hdist_lower_Dp. - assert (Hdist_same := frame_dist (fst (conf (Good g))) (fst (conf (Good g'))) (change_frame da conf g')). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - simpl in *. - rewrite <- Hdist_same. - assert ((dist (fst (conf (Good g))) (fst (conf (Good g'))) <= Dp-3*D)%R -> (sqrt - ((fst (fst (conf (Good g))) + - fst (fst (conf (Good g')))) * - (fst (fst (conf (Good g))) + - fst (fst (conf (Good g')))) + - (snd (fst (conf (Good g))) + - snd (fst (conf (Good g')))) * - (snd (fst (conf (Good g))) + - snd (fst (conf (Good g'))))) <= Dp-3*D)%R) by now simpl. - apply H. - assert (dist (fst (round rbg_ila da conf (Good g))) (fst (round rbg_ila da conf (Good g'))) = 0%R) by now simpl. - assert (dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))) <= D)%R. - { assert (Hdist_r:= @dist_round_max_d g' conf da Hpred Hpath (still_alive_means_alive g' Hpred Hbased_higher Halive')). - unfold equiv, bool_Setoid, eq_setoid in Hdist_r. - rewrite Rle_bool_true_iff in Hdist_r. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g))) (fst (round rbg_ila da conf (Good g'))) (fst (conf (Good g')))). - rewrite H0 in Htri. - transitivity (dist (fst (round rbg_ila da conf (Good g'))) (fst (conf (Good g')))); try lra. - rewrite dist_sym. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - apply Hdist_r. - } - assert (dist (fst ( conf (Good g))) (fst (conf (Good g'))) <= 2*D)%R. - { assert (Hdist_r:= @dist_round_max_d g conf da Hpred Hpath (still_alive_means_alive g Hpred Hbased_higher Halive)). - unfold equiv, bool_Setoid, eq_setoid in Hdist_r. - rewrite Rle_bool_true_iff in Hdist_r. - assert (Htri := RealMetricSpace.triang_ineq (fst (conf (Good g))) (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g')))). - transitivity (dist (fst (conf (Good g))) (fst (round rbg_ila da conf (Good g))) + - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))))%R; try auto. - transitivity (D+ dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g'))))%R. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - apply Rplus_le_compat_r. - apply Hdist_r. - replace (2*D)%R with (D + D)%R. - apply Rplus_le_compat_l, H1. - generalize D_0; lra. - } - unfold Dp; generalize Dmax_7D, D_0; lra. - + specialize (no_col (reflexivity _) (reflexivity _) (still_alive_means_alive g Hpred Hbased_higher Halive) (still_alive_means_alive g' Hpred Hbased_higher Halive')). - destruct (R2_EqDec (@get_location Loc _ _ (conf (Good g))) (@get_location Loc _ _ (round rbg_ila da conf (Good g)))%R); - destruct (R2_EqDec (@get_location Loc _ _ (conf (Good g'))) (@get_location Loc _ _ (round rbg_ila da conf (Good g')))%R). - - now rewrite <- e, <- e0. - - unfold get_based in *; simpl in Hbased, Hbased'. - assert (Hmove := moving_means_not_near g' Hpred Hcoll Hpath Hbased_higher c Halive'). - destruct Hmove as (HD, (H2d, Hdp)). - specialize (H2d g). - destruct ((leb (get_ident (conf (Good g'))) (get_ident (conf (Good g))))) - eqn : Hident. - ++ rewrite Nat.leb_le in Hident. - destruct (Rle_bool (dist (get_location (conf (Good g'))) (get_location (conf (Good g)))) D) eqn : Hdist_D. - * rewrite round_good_g in Halive; try apply Hpred. - unfold get_alive in Halive; - simpl in Halive. - destruct (conf (Good g)) as (pos, (((id, li),al), ba)) eqn : Hconf. - simpl in Halive. - unfold upt_aux in Halive. - simpl in Hbased. - rewrite Hbased in Halive. - destruct (upt_robot_dies_b _) eqn : Htrue. - unfold map_config in *; rewrite Hconf in Halive. - simpl in Halive. - rewrite Hbased in Halive. - now simpl in Halive. - unfold map_config in *. - rewrite Hconf in Halive; simpl in Halive. - unfold upt_robot_dies_b in Htrue. - rewrite orb_false_iff in Htrue. - destruct Htrue as (Htrue,_). - rewrite <- negb_true_iff in Htrue. - rewrite forallb_existsb, forallb_forall in Htrue. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - specialize (Htrue (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g')))), snd (conf (Good g')))). - assert (@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g')))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g)))))) - (get_alive elt)) (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (fun id : @ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id)))))). - { rewrite filter_In. - rewrite config_list_spec. - rewrite in_map_iff. - repeat split; simpl. - exists (Good g'). - split; auto. - destruct b; simpl; auto. - apply In_names. - generalize (ident_unique conf Hg). - rewrite Hconf; - unfold get_ident in *; simpl in *; auto. - apply le_lt_or_eq in Hident. - rewrite 2 andb_true_iff. - destruct Hident; auto. - repeat split; try now rewrite Nat.ltb_lt. - apply (still_alive_means_alive _ Hpred Hbased_higher) in Halive'. - now unfold get_alive in *; simpl. - rewrite negb_true_iff. - unfold get_based in *; auto. - } - specialize (Htrue H). - clear H. - assert (Hframe := frame_dist (fst (conf (Good g'))) (fst (conf (Good g))) (r,t,b)). - unfold frame_choice_bijection, Frame in Hframe. - - assert (dist (fst (conf (Good g'))) (fst (conf (Good g))) == - (@dist (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS ES)) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g'))))) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec - VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g))))))). - rewrite Hframe. - destruct b; simpl in *; auto. - rewrite <- H in Htrue; clear H. - rewrite Hconf in *. clear Hframe. - simpl in *; unfold Datatypes.id in *; - rewrite Hdist_D in Htrue. - exfalso; apply no_fixpoint_negb in Htrue; auto. - * rewrite Rle_bool_false_iff in Hdist_D. - apply Rnot_le_lt in Hdist_D. - assert (Hdist' :( 0 < dist (get_location (round rbg_ila da conf (Good g'))) (get_location (conf (Good g))))%R). - { set (x := (get_location (round rbg_ila da conf (Good g')))) in *; - set (y := (get_location (conf (Good g')))) in *; - set (z := (get_location (conf (Good g)))) in *. - apply Rlt_Rminus in Hdist_D. - assert (dist y z - D <= dist y z - dist y x)%R. - lra. - assert (Htri := RealMetricSpace.triang_ineq y x z). - assert (dist y z - dist y x <= dist x z)%R by lra. - lra. - } - rewrite <- e. rewrite dist_sym. lra. - ++ rewrite Nat.leb_gt in Hident. - specialize (H2d Hident ((still_alive_means_alive g Hpred Hbased_higher Halive))). - unfold equiv, bool_Setoid, eq_setoid in H2d. - rewrite Rle_bool_false_iff in H2d. - intros Hd; destruct H2d. - unfold get_based; assumption. - rewrite e, Hd. - generalize D_0; - lra. - - unfold get_based in *. - assert (Hmove := moving_means_not_near g Hpred Hcoll Hpath Hbased_higher c Halive). - destruct Hmove as (HD, (H2d, Hdp)). - specialize (H2d g'). - destruct ((leb (get_ident (conf (Good g))) (get_ident (conf (Good g'))))) - eqn : Hident. - ++ rewrite Nat.leb_le in Hident. - destruct (Rle_bool (dist (get_location (conf (Good g))) (get_location (conf (Good g')))) D) eqn : Hdist_D. - * rewrite round_good_g in Halive'; try apply Hpred. - unfold get_alive in Halive'; - simpl in Halive'. - destruct (conf (Good g')) as (pos, (((id, li),al), ba)) eqn : Hconf. - simpl in Halive'. - unfold upt_aux in Halive'. - destruct (upt_robot_dies_b _) eqn : Htrue. - unfold map_config in *; - rewrite Hconf in Halive'; - simpl in Halive'; simpl in Hbased'; rewrite Hbased' in *. - now simpl in Halive'. - unfold map_config in *; - rewrite Hconf in Halive'; simpl in Halive'. - unfold upt_robot_dies_b in Htrue. - rewrite orb_false_iff in Htrue. - destruct Htrue as (Htrue,_). - rewrite <- negb_true_iff in Htrue. - rewrite forallb_existsb, forallb_forall in Htrue. - destruct (change_frame da conf g') as ((r,t),b) eqn : Hchange. - specialize (Htrue (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g)))), snd (conf (Good g)))). - assert (@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g)))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g')))))) - (get_alive elt)) (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (fun id : @ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id)))))). - { rewrite filter_In. - rewrite config_list_spec. - rewrite in_map_iff. - repeat split; simpl. - exists (Good g). - split; auto. - destruct b; simpl; auto. - apply In_names. - generalize (ident_unique conf Hg). - rewrite Hconf; - unfold get_ident in *; simpl in *; auto. - rewrite 2 andb_true_iff. - repeat split. - rewrite Nat.ltb_lt. - apply le_lt_or_eq in Hident. - destruct Hident. - auto. - now destruct H. - apply (still_alive_means_alive _ Hpred Hbased_higher) in Halive. - unfold get_alive in *; simpl in *; auto. - rewrite negb_true_iff; unfold get_based in *; simpl in *; auto. - } - specialize (Htrue H). - clear H. - assert (Hframe := frame_dist (fst (conf (Good g))) (fst (conf (Good g'))) (r,t,b)). - unfold frame_choice_bijection, Frame in Hframe. - - assert (dist (fst (conf (Good g))) (fst (conf (Good g'))) == - (@dist (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS ES)) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g))))) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec - VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g'))))))). - rewrite Hframe. - destruct b; simpl in *; auto. - rewrite <- H in Htrue; clear H. - rewrite Hconf in *. clear Hframe. - simpl in *; unfold Datatypes.id in *; - rewrite Hdist_D in Htrue. - exfalso; apply no_fixpoint_negb in Htrue; auto. - * rewrite Rle_bool_false_iff in Hdist_D. - intros H_0. - apply Hdist_D. - rewrite e. - assert (Htri := RealMetricSpace.triang_ineq (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g))) (get_location (round rbg_ila da conf (Good g')))). - rewrite H_0 in Htri. - lra. - ++ rewrite Nat.leb_gt in Hident. - specialize (H2d Hident ((still_alive_means_alive g' Hpred Hbased_higher Halive'))). - unfold equiv, bool_Setoid, eq_setoid in H2d. - rewrite Rle_bool_false_iff in H2d. - intros Hd; destruct H2d. - unfold get_based; assumption. - rewrite dist_sym. - rewrite e, Hd. - generalize D_0; - lra. - - unfold get_based in *; - assert (Hmove := moving_means_not_near g Hpred Hcoll Hpath Hbased_higher c Halive). - destruct Hmove as (HD, (H2d, Hdp)). - specialize (H2d g'). - destruct ((leb (get_ident (conf (Good g))) (get_ident (conf (Good g'))))) - eqn : Hident. - ++ rewrite Nat.leb_le in Hident. - destruct (Rle_bool (dist (get_location (conf (Good g))) (get_location (conf (Good g')))) D) eqn : Hdist_D. - * rewrite round_good_g in Halive'; try apply Hpred. - unfold get_alive in Halive'; - simpl in Halive'. - destruct (conf (Good g')) as (pos, (((id, li),al), ba)) eqn : Hconf. - simpl in Halive'. - unfold upt_aux in Halive'. - destruct (upt_robot_dies_b _) eqn : Htrue. - unfold map_config in *; - rewrite Hconf in Halive'; - simpl in Halive; simpl in Hbased'; rewrite Hbased' in *; - now simpl in Halive'. - unfold map_config in *; - rewrite Hconf in Halive'; simpl in Halive'. - unfold upt_robot_dies_b in Htrue. - rewrite orb_false_iff in Htrue. - destruct Htrue as (Htrue,_). - rewrite <- negb_true_iff in Htrue. - rewrite forallb_existsb, forallb_forall in Htrue. - destruct (change_frame da conf g') as ((r,t),b) eqn : Hchange. - specialize (Htrue (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g)))), snd (conf (Good g)))). - assert (@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g)))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g')))))) - (get_alive elt)) (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (fun id : @ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id)))))). - { rewrite filter_In. - rewrite config_list_spec. - rewrite in_map_iff. - repeat split; simpl. - exists (Good g). - split; auto. - destruct b; simpl; auto. - apply In_names. - generalize (ident_unique conf Hg). - rewrite Hconf; - unfold get_ident in *; simpl in *; auto. - apply le_lt_or_eq in Hident. - destruct Hident; auto. - rewrite 2 andb_true_iff. - repeat split; try now rewrite Nat.ltb_lt. - apply (still_alive_means_alive _ Hpred Hbased_higher) in Halive; - unfold get_alive in *; simpl in *; auto. - rewrite negb_true_iff; unfold get_based in *; simpl in *; auto. - } - specialize (Htrue H). - clear H. - assert (Hframe := frame_dist (fst (conf (Good g))) (fst (conf (Good g'))) (r,t,b)). - unfold frame_choice_bijection, Frame in Hframe. - - assert (dist (fst (conf (Good g))) (fst (conf (Good g'))) == - (@dist (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS ES)) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g))))) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec - VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec - VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf (@Good Identifiers g')))) - (@snd R2 ILA (conf (@Good Identifiers g'))))))). - rewrite Hframe. - destruct b; simpl in *; auto. - rewrite <- H in Htrue; clear H. - rewrite Hconf in *. clear Hframe. - simpl in *; unfold Datatypes.id in *; - rewrite Hdist_D in Htrue. - exfalso; apply no_fixpoint_negb in Htrue; auto. - * assert (Hmove' := moving_means_not_near g' Hpred Hcoll Hpath Hbased_higher c0 Halive'). - destruct Hmove' as (HD', (H2d', Hdp')). - specialize (H2d' g). - assert ( get_ident (conf (Good g)) < get_ident (conf (Good g'))). - { - generalize (ident_unique conf Hg). - apply le_lt_or_eq in Hident. - destruct Hident; now intro. - } - specialize (H2d' H (still_alive_means_alive g Hpred Hbased_higher Halive)). - clear H. - unfold equiv, bool_Setoid, eq_setoid in H2d'. - rewrite Rle_bool_false_iff in H2d'. - intros Heq. - destruct H2d'. - unfold get_based; assumption. - set (x := (get_location (round rbg_ila da conf (Good g)))) in *; - set (y := (get_location (conf (Good g)))) in *; - set (w := (get_location (round rbg_ila da conf (Good g')))) in *; - set (z := (get_location (conf (Good g')))) in *. - assert (Htri := RealMetricSpace.triang_ineq y x w). - assert (dist y w <= D)%R by lra. - generalize D_0; lra. - ++ rewrite Nat.leb_gt in Hident. - specialize (H2d Hident ((still_alive_means_alive g' Hpred Hbased_higher Halive'))). - unfold equiv, bool_Setoid, eq_setoid in H2d. - rewrite Rle_bool_false_iff in H2d. - intros Hd; destruct H2d. - unfold get_based ; assumption. - rewrite dist_sym. - assert (Hmove' := moving_means_not_near g' Hpred Hcoll Hpath Hbased_higher c0 Halive'). - destruct Hmove' as (HD', (H2d', Hdp')). - specialize (H2d' g). - set (x := (get_location (round rbg_ila da conf (Good g)))) in *; - set (y := (get_location (conf (Good g)))) in *; - set (w := (get_location (round rbg_ila da conf (Good g')))) in *; - set (z := (get_location (conf (Good g')))) in *. - assert (Htri := RealMetricSpace.triang_ineq x w z). - rewrite dist_sym in HD'. - assert (dist x z <= D)%R by lra. - generalize D_0; lra. -Qed. - - - -Definition target_alive conf := - forall g, - get_ident (conf (Good g)) <> 0 -> - get_alive (conf (Good g)) = true -> - exists g', get_alive (conf (Good g')) = true /\ - get_ident (conf (Good g')) < get_ident (conf (Good g)) /\ - (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) - <= Dmax)%R. - -Lemma ident_preserved : forall conf g da, - da_predicat da -> - get_ident (conf (Good g)) = get_ident (round rbg_ila da conf (Good g)). -Proof. - intros conf g da Hpred. - rewrite (round_good_g g conf Hpred). - unfold get_ident; simpl. - unfold upt_aux, rbg_fnc; simpl. - destruct (conf (Good g)) as (?,(((?,?),?),?)). - destruct b; - destruct (upt_robot_dies_b _); simpl; auto; - destruct (upt_robot_too_far _ _); simpl; auto. -Qed. - - - -Lemma executed_means_not_moving : forall conf da g, - da_predicat da -> - get_alive (conf (Good g)) = true -> - get_alive (round rbg_ila da conf (Good g)) = false -> - get_location (conf (Good g)) == (get_location (round rbg_ila da conf (Good g))). -Proof. - intros conf da g Hpred. - intros Hal' Hdead'. - rewrite (@round_good_g g) in *; try apply Hpred. - simpl in *. - unfold Datatypes.id in *. - destruct (change_frame da ( conf) g) as ((r,t),b). - unfold upt_aux in *. - unfold map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - destruct based; simpl (fst _) in *; simpl (snd _) in *. - destruct (upt_robot_too_far _ _) eqn : Htoo_far; - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g)))); - unfold frame_choice_bijection, Frame in Hret_sec; - rewrite Hconf in Hret_sec; - simpl (_ ∘ _) in Hret_sec; - simpl (fst _) in Hret_sec; - rewrite <- Hret_sec at 1; - simpl in *; auto. - destruct (upt_robot_dies_b _) eqn : Hbool. - - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g)))); - unfold frame_choice_bijection, Frame in Hret_sec; - rewrite Hconf in Hret_sec; - simpl (_ ∘ _) in Hret_sec; - simpl (fst _) in Hret_sec; - rewrite <- Hret_sec at 1; - simpl in *; auto. - - unfold get_alive in *; simpl in *. - now rewrite Hal' in Hdead'. -Qed. - - -Lemma executed_means_alive_near : forall conf g da, - da_predicat da -> - get_alive (conf (Good g)) = true -> - get_ident (conf (Good g)) <> 0 -> - get_alive (round rbg_ila da conf (Good g)) = false -> - path_conf conf -> - (exists g0, - get_alive (conf (Good g0)) = true /\ - get_based (conf (Good g0)) = false /\ - get_ident (conf (Good g0)) < get_ident (conf (Good g)) /\ - (dist (get_location (conf (Good g0))) (get_location (conf (Good g))) - <= Dmax)%R) -> - exists g', get_ident (conf (Good g')) < get_ident (conf (Good g)) /\ - (dist (get_location (conf (Good g))) - (get_location (conf (Good g'))) <= D)%R - /\ get_based (conf (Good g')) = false - /\ get_alive ((conf (Good g'))) = true. -Proof. - intros conf g da Hpred Halive_g Hident_0 Halive_gr Hpath Hexists. - assert (Hnot_moving := executed_means_not_moving conf g Hpred). - specialize (Hnot_moving Halive_g Halive_gr). - destruct Hexists as (g0, (Halive_g0, (Hident, Hdist))). - unfold path_conf in Hpath. - rewrite round_good_g in Halive_gr. - unfold get_alive in Halive_gr; simpl in Halive_gr. - unfold upt_aux in *. - unfold map_config in *. - destruct (conf (Good g)) as (p_g, (((i_g, l_g), a_g), b_g)) eqn : Hconf. - destruct b_g. - simpl in Halive_gr. - unfold get_alive in Halive_g; simpl in Halive_g. - destruct (upt_robot_too_far _ _); simpl in Halive_gr; rewrite Halive_g in *; discriminate. - destruct (upt_robot_dies_b _) eqn : Hupt. - simpl in Halive_gr. - unfold upt_robot_dies_b in Hupt. - apply orb_prop in Hupt. - destruct Hupt as [Hnear|Halone]. - - rewrite existsb_exists in Hnear. - destruct Hnear as (exec, (Hin, Hnear)). - rewrite filter_In in Hin. - destruct Hin as (HIn, HIn_ident). - rewrite Hconf in *. - unfold get_ident in *; - simpl in HIn_ident. - rewrite config_list_spec in HIn. - rewrite in_map_iff in HIn. - destruct HIn as (id_exec, (Hid_exec, HIn)). - destruct id_exec as [g_exec |byz]. - + exists g_exec. - split. - * simpl. - rewrite 2 andb_true_iff in HIn_ident. - destruct HIn_ident as ((Hin_exec,Halive_exec), Hbased_exec). - rewrite Nat.ltb_lt, <- Hid_exec in Hin_exec. - now simpl in *. - * repeat split. - -- rewrite <- Hid_exec in Hnear. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - simpl (fst _) in Hnear. - assert (Hframe_dist := frame_dist (fst (conf (Good g_exec))) p_g (change_frame da conf g)). - rewrite <- Hframe_dist in Hnear. - rewrite dist_sym. now rewrite Rle_bool_true_iff in *; simpl in *. - -- rewrite andb_true_iff, negb_true_iff in HIn_ident. - rewrite <- Hid_exec in HIn_ident. - destruct HIn_ident. - unfold get_based in *; simpl in *; auto. - -- rewrite 2 andb_true_iff in HIn_ident. - destruct HIn_ident. - rewrite <- Hid_exec in *. - now unfold get_alive in *; simpl in *. - + simpl in HIn. - generalize (In_Bnames byz), (Bnames_length). - now simpl. - - rewrite forallb_existsb, forallb_forall in Halone. - specialize (Halone (((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) (fst (conf (Good g0)))), snd (conf(Good g0)))). - assert (Hin : @List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 light) (prod (prod R R2) bool) bool bool robot_choice_RL Frame - Choice inactive_choice_ila da conf g return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end) - end - end (@fst R2 ILA (conf (@Good Identifiers g0)))) - (@snd R2 ILA (conf (@Good Identifiers g0)))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end)) - end - end (@fst R2 ILA (conf (@Good Identifiers g)))) - (@snd R2 ILA (conf (@Good Identifiers g)))))) - (get_alive elt)) (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (fun id : @ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 light) (prod (prod R R2) bool) bool bool robot_choice_RL - Frame Choice inactive_choice_ila da conf g - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id)))))). - { - rewrite filter_In. - rewrite 2 andb_true_iff. - repeat split. - * rewrite config_list_spec, in_map_iff. - exists (Good g0). - split; auto; try apply In_names. - destruct (change_frame da conf g) as (?,b); - now destruct b. - * unfold get_ident in *; rewrite Hconf in *; simpl in *; auto. - now rewrite Nat.ltb_lt. - * now unfold get_alive in *; simpl in *. - * rewrite negb_true_iff; unfold get_based in *. auto. - } - specialize (Halone Hin). - clear Hin. - rewrite negb_true_iff in *. - rewrite Hconf in *. - assert (Hframe_dist := frame_dist (fst (conf (Good g0))) p_g (change_frame da conf g)). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - unfold frame_choice_bijection, Frame in Hframe_dist. - destruct (change_frame _) as ((r,t),b) eqn : Hchange. - simpl (fst _) in *. - assert (((@dist (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) - (@location_Setoid Loc) (@location_EqDec Loc) VS ES)) - (R2.bij_rotation_f r - (@pair R R - (Rplus - (@fst R R - (@section R2 R2_Setoid - match b return (@bijection R2 R2_Setoid) with - | true => bij_reflexion - | false => @id R2 R2_Setoid - end (@fst R2 ILA (conf (@Good Identifiers g0))))) - (@fst R R t)) - (Rplus - (@snd R R - (@section R2 R2_Setoid - match b return (@bijection R2 R2_Setoid) with - | true => bij_reflexion - | false => @id R2 R2_Setoid - end (@fst R2 ILA (conf (@Good Identifiers g0))))) - (@snd R R t)))) - (R2.bij_rotation_f r - (@pair R R - (Rplus - (@fst R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end) p_g)) (@fst R R t)) - (Rplus - (@snd R R - (@section R2 R2_Setoid - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS - ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - end) p_g)) (@snd R R t)))))) == dist - ((rotation r - ∘ (translation t ∘ (if b then reflexion else Similarity.id))) - (fst (conf (Good g0)))) - ((rotation r - ∘ (translation t ∘ (if b then reflexion else Similarity.id))) - p_g)). - now destruct b; simpl. - rewrite <- H in Hframe_dist. - rewrite <- Hframe_dist in Halone. - clear H. - rewrite Rle_bool_false_iff in Halone; now destruct Halone. - - unfold get_alive in *; simpl in Halive_gr. - simpl in *. - now rewrite Halive_g in *. - - apply Hpred. -Qed. - - - - - - - -Definition executioner_means_light_off conf := - forall g da, - da_predicat da -> - get_alive (conf (Good g)) = true -> - (exists g', get_alive (conf (Good g')) = true /\ get_based (conf (Good g')) = false /\ - get_ident (conf (Good g)) < get_ident (conf (Good g')) /\ - (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) - <= D)%R) -> - get_light (conf (Good g)) = false. - - -Definition executed_means_light_on conf := forall g da, - da_predicat da -> - get_alive (conf (Good g)) = true -> - get_alive (round rbg_ila da conf (Good g)) = false -> - get_light (conf (Good g)) = true. - - - -Lemma light_on_means_not_moving_before : forall conf g da, - da_predicat da -> - path_conf conf -> - get_alive (round rbg_ila da conf (Good g)) = true -> - get_light (round rbg_ila da conf (Good g)) = true -> - get_location (conf (Good g)) == get_location (round rbg_ila da conf (Good g)). -Proof. - intros conf g da Hpred Hpath Halive Hlight. - rewrite round_good_g in *; try apply Hpred. - destruct (conf (Good g)) as (pos, (((ide, lig), ali), bas)) eqn : Hconf. - simpl in *. - unfold upt_aux in *; - unfold map_config in *. - rewrite Hconf in *. - destruct bas. - unfold get_alive in *; simpl in *. - destruct (upt_robot_too_far _ _); - assert (Hret_sec := retraction_section (frame_choice_bijection (change_frame da conf g)) (fst (conf (Good g)))); - unfold frame_choice_bijection, Frame in Hret_sec; - rewrite Hconf in Hret_sec; - simpl (_ ∘ _) in Hret_sec; - simpl (fst _) in Hret_sec; - rewrite <- Hret_sec at 1; - simpl in *; auto. - destruct (upt_robot_dies_b _) eqn : Hbool. - unfold get_alive in *; - simpl in *. - discriminate. - unfold get_light in *; simpl in *. - unfold rbg_fnc in *. - destruct (move_to _) eqn : Hmove. - simpl in *; discriminate. - destruct (change_frame _) as ((r,t),b) eqn : Hchange. - destruct Hpred as (Horigin,_). - specialize (Horigin conf g ((r,t),b) Hchange). - symmetry. - fold equiv. - assert (Datatypes.id - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (0%R, 0%R, true))) == Datatypes.id pos). - unfold Datatypes.id in *. - destruct b; rewrite <- Inversion; - simpl in *; - unfold Datatypes.id in *; rewrite Hconf in *; simpl in *; - apply Horigin. - simpl in *. - unfold Datatypes.id in *. - now destruct b; simpl in *. -Qed. - - - - -Definition exists_at_less_that_Dp conf := - forall g, - get_alive (conf (Good g)) = true -> - get_ident (conf (Good g)) > 0 -> - (forall g_near, get_alive (conf (Good g_near)) = true -> - (dist (get_location (conf (Good g_near))) - (get_location (conf (Good g))) <= Dmax)%R -> - get_ident (conf (Good g_near)) < get_ident (conf (Good g)) -> - get_light (conf (Good g_near)) = true) -> - exists g', - get_alive (conf (Good g')) = true /\ - get_ident (conf (Good g')) < get_ident (conf (Good g)) /\ - (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) <= Dp)%R -. - - -Lemma based_higher_round : forall conf da, - da_predicat da -> - path_conf conf -> - based_higher conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - (exists g, get_based (conf (Good g)) = true) -> - based_higher (round rbg_ila da conf). -Proof. - intros conf da Hpred Hpath Hbased_higher Hexecutioner Hexecuted Hat_based. - repeat split. - - rewrite round_good_g in *; try apply Hpred. - simpl in *. - unfold upt_aux, map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - unfold get_based, get_alive in *. - simpl in *. - destruct based eqn : Hbased. - + destruct (upt_robot_too_far _ _); simpl in *; try discriminate. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in *. - unfold get_based, get_alive in Hap0; simpl in Hap0; specialize (Hap0 (reflexivity _)). - intuition. - + destruct (upt_robot_dies_b _ ) eqn : Hbool; simpl in *; try discriminate. - - rewrite round_good_g in *; try apply Hpred. - simpl in *. - unfold upt_aux, map_config in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - unfold get_based, get_alive in *. - simpl in *. - destruct based eqn : Hbased. - + destruct (upt_robot_too_far _ _); simpl in *; try discriminate. - rewrite retraction_section. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in *. - unfold get_based, get_alive in Hap0; simpl in Hap0; specialize (Hap0 (reflexivity _)). - intuition. - + destruct (upt_robot_dies_b _ ) eqn : Hbool; simpl in *; try discriminate. - - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (Hap0, (Hex, Hhigher)). - destruct (existsb (fun g => get_alive (conf (Good g)) && - negb (get_based (conf (Good g))) && - Rle_bool (dist (@get_location - Loc _ - _ (conf (Good g)) - ) (0,0)) (Dp - 3*D))%R Gnames) eqn : Hdist_DP_D. - + rewrite existsb_exists in Hdist_DP_D. - destruct Hdist_DP_D as (g_ex, (_, Haliveandall)). - rewrite 2 andb_true_iff, Rle_bool_true_iff in Haliveandall. - destruct Haliveandall as ((Hex_alive, Hex_based), Hdist_DP_D). - rewrite negb_true_iff in Hex_based. - destruct (get_alive (round rbg_ila da conf (Good g_ex))) eqn : Halive_ex_round. - - exists g_ex. - assert (dist (get_location (round rbg_ila da conf (Good g_ex))) (0,0) <= Dp - D)%R. - { assert (Hdist_d := dist_round_max_d g_ex Hpred Hpath). - destruct (conf (Good g_ex)) as (pos_ex,(((ident_ex, light_ex), alive_ex),based_ex)) eqn : Hconf_ex. - specialize (Hdist_d Hex_alive). - unfold equiv, bool_Setoid, eq_setoid in Hdist_d. - rewrite Rle_bool_true_iff in *. - rewrite dist_sym in Hdist_d. - assert (Htri := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g_ex))) (get_location (pos_ex, (ident_ex, light_ex, alive_ex, false))) (0,0)%R). - rewrite <- (Rplus_minus D Dp). - transitivity (D + (Dp - 3*D))%R. - transitivity (dist (get_location (round rbg_ila da conf (Good g_ex))) - (get_location (pos_ex, (ident_ex, light_ex, alive_ex, false))) + - dist (get_location (pos_ex, (ident_ex, light_ex, alive_ex, false))) (0, 0))%R; try auto. - transitivity (dist (get_location (round rbg_ila da conf (Good g_ex))) - (get_location (pos_ex, (ident_ex, light_ex, alive_ex, false))) + (Dp - 3*D))%R; try lra. - apply Rplus_le_compat_l. - apply Hdist_DP_D. - apply Rplus_le_compat_r. - apply Hdist_d. - generalize D_0; lra. - } - - rewrite round_good_g in *; try apply Hpred. - simpl (let _ := _ in _ ) in *. - unfold upt_aux, map_config in *. - destruct (conf (Good g_ex)) as (pos_ex,(((ident_ex, light_ex), alive_ex),based_ex)) eqn : Hconf_ex. - destruct based_ex; try (unfold get_based in *; now simpl in *). - destruct (upt_robot_dies_b ) eqn : Hdies. unfold get_alive in *; simpl in *. - discriminate. - repeat split. - now unfold get_alive in *; simpl in *. - auto. - (* comme on a dist_dp_d, le bourreau de g_ex est au moins à Dp, donc ça marche *) - rewrite round_good_g in Halive_ex_round; try apply Hpred. - simpl in Halive_ex_round. - unfold upt_aux, map_config in Halive_ex_round. - destruct (conf (Good g_ex)) as (pos_ex,(((ident_ex, light_ex), alive_ex),based_ex)) eqn : Hconf_ex. - unfold get_alive in Halive_ex_round; simpl in Halive_ex_round. - unfold get_based in Hex_based; simpl in Hex_based; rewrite Hex_based in Halive_ex_round. - unfold get_alive in *; simpl in *. - destruct (upt_robot_dies_b _) eqn : Hbool_round in Halive_ex_round; simpl in Halive_ex_round; - try (rewrite Hex_alive in *; discriminate). - assert (Hbool_aux := Hbool_round). - unfold upt_robot_dies_b in Hbool_round. - rewrite orb_true_iff in *. - destruct Hbool_round as [Hnear|Hfar]. - rewrite existsb_exists in *. - destruct Hnear as (near, (Hin_near, Hdist_near)). - rewrite filter_In, config_list_spec, in_map_iff in Hin_near. - rewrite 2 andb_true_iff, Nat.ltb_lt in Hin_near. - destruct Hin_near as (([g_near|byz], (Hnear_spec, _)), ((Hident_near, Halive_near), Hbased_near)); try ( - assert (Hfalse := In_Bnames byz); - now simpl in *). - exists g_near. - - assert (get_alive (round rbg_ila da conf (Good g_near)) = true). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux. - destruct (conf (Good g_near)) as (p_n, (((i_n, l_n), a_n), b_n)) eqn : Hconf_near. - rewrite negb_true_iff in *. - rewrite <- Hnear_spec in *; unfold get_based in *; simpl in Hbased_near. - rewrite Hbased_near. - - destruct (upt_robot_dies_b _ g_near) eqn : Hbool_near. - - assert (Hfalse : get_alive (round rbg_ila da conf (Good g_near)) = false). - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config. - unfold get_alive ; simpl. - rewrite Hconf_near. - simpl. - rewrite Hbased_near. - destruct (upt_robot_dies_b _ g_near). - now destruct (conf (Good g_near)) as (?, ((?,?),?)); simpl. - discriminate. - unfold get_alive in *; simpl in *. - assert (Hlight_false : get_light (conf (Good g_near)) = true). - apply (Hexecuted g_near da Hpred). - unfold get_alive; rewrite Hconf_near; simpl; auto. - apply Hfalse. - assert (Hlight_true : get_light (conf (Good g_near)) = false). - apply (Hexecutioner g_near da Hpred). - unfold get_alive; rewrite Hconf_near; simpl; auto. - exists g_ex. - repeat split; try auto. - rewrite Hconf_ex; unfold get_alive; simpl; auto. - rewrite Hconf_ex; unfold get_based; simpl; auto. - rewrite Hconf_ex, Hconf_near in *; unfold get_ident in *; simpl in *; auto. - rewrite Rle_bool_true_iff in *. - rewrite (frame_dist _ _ (change_frame da conf g_ex)). - simpl in *; rewrite Hconf_ex, Hconf_near in *; simpl in *; auto. - rewrite Hlight_true in *. - discriminate. - - unfold get_alive in *; - simpl; auto. - rewrite Hconf_near. - simpl; auto. - rewrite Hbased_near. - simpl; auto. - } - repeat split. - * rewrite round_good_g; try apply Hpred. - simpl. - rewrite negb_true_iff in *. - rewrite <- Hnear_spec in Hbased_near. - unfold get_based in Hbased_near; simpl in Hbased_near. - unfold upt_aux, map_config. - simpl. - destruct (conf (Good g_near)) as (pos_n,(((ident_n, light_n), alive_n), based_n)) eqn : Hconf_near. - simpl in *. - rewrite Hbased_near. - unfold get_based; simpl. - destruct (upt_robot_dies_b _ g_near) eqn : Hbool_near; simpl; try auto. - * apply H. - * assert (dist (fst (round rbg_ila da conf (Good g_near))) (0,0)%R <= Dp-D -> (sqrt - ((fst (Datatypes.id (fst (round rbg_ila da conf (Good g_near)))) + - 0) * - (fst (Datatypes.id (fst (round rbg_ila da conf (Good g_near)))) + - 0) + - (snd (Datatypes.id (fst (round rbg_ila da conf (Good g_near)))) + - 0) * - (snd (Datatypes.id (fst (round rbg_ila da conf (Good g_near)))) + - 0)) <= Dp-D)%R)%R. - simpl ; auto. - apply H0. - clear H0. - assert (Hdist_DP_D_aux : (dist (pos_ex) (0,0)%R <= Dp - 3*D)%R) by now simpl. - destruct (conf (Good g_near)) as (pos_n,(((ident_n, light_n), alive_n), based_n)) eqn : Hconf_near. - assert (Hdist_near_aux : (dist (pos_n) pos_ex <= D)%R). - rewrite <- Hnear_spec, Rle_bool_true_iff in Hdist_near. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite Hconf_ex in *. - rewrite (frame_dist _ _ (change_frame da conf g_ex)). - now simpl in *. - assert (get_alive (conf( Good g_near)) = true). - unfold get_alive in *; rewrite <- Hnear_spec, Hconf_near in *; simpl in *; auto. - assert (Hdist_r_n := @dist_round_max_d g_near conf da Hpred Hpath H0). - unfold equiv, bool_Setoid, eq_setoid in Hdist_r_n. - rewrite Rle_bool_true_iff in *. - rewrite <- (Rplus_minus D Dp). - rewrite <- (Rplus_minus D (D + (Dp - D))%R). - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_near))) (fst (conf( Good g_near))) (0,0)%R). - transitivity ( dist (fst (round rbg_ila da conf (Good g_near))) (fst (conf (Good g_near))) + - dist (fst (conf (Good g_near))) (0, 0))%R; try auto. - rewrite dist_sym in Hdist_r_n. - transitivity (D + dist (fst (conf (Good g_near))) (0,0))%R. - apply Rplus_le_compat_r. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - apply Hdist_r_n. - replace (D + (D + (Dp - D) - D) - D)%R with ((D) + ((D + (Dp - D) - D) - D))%R by lra. - apply Rplus_le_compat_l. - assert (Htri2 := RealMetricSpace.triang_ineq (fst (conf (Good g_near))) (fst (conf (Good g_ex))) (0,0)%R). - transitivity (dist (fst (conf (Good g_near))) (fst (conf (Good g_ex))) + - dist (fst (conf (Good g_ex))) (0, 0))%R; try auto. - rewrite Hconf_near, Hconf_ex in *; simpl (fst _) in *. - transitivity (D + dist pos_ex (0,0))%R. - apply Rplus_le_compat_r. - lra. - replace (D + (Dp - D) - D - D)%R with (D + (Dp - 3*D) )%R by lra. - apply Rplus_le_compat_l. - apply Hdist_DP_D_aux. - * rewrite forallb_existsb in Hfar. - unfold path_conf in Hpath. - assert (Hpath_backup := Hpath). - specialize (Hpath g_ex). - rewrite Hconf_ex in Hpath; unfold get_alive in Hpath; - simpl (snd _) in Hpath; specialize (Hpath Hex_alive). - destruct Hpath. - rewrite <- Hconf_ex in H. - assert (Halive_0 := ident_0_alive (round rbg_ila da conf) g_ex). - rewrite <- ident_preserved in Halive_0. - specialize (Halive_0 H). - rewrite round_good_g in Halive_0; try apply Hpred. - simpl in Halive_0. - unfold upt_aux, map_config in *. - rewrite Hconf_ex, Hex_based in Halive_0; simpl in Halive_0. - rewrite Hbool_aux in Halive_0. - unfold get_alive in Halive_0; simpl in Halive_0; discriminate. - apply Hpred. - destruct H as (g_close, (Halive_close - , (Hdist_close, (Hident_close, Hbased_close)))). - rewrite forallb_forall in *. - specialize (Hfar ((let (p, b) := change_frame da conf g_ex in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_close)) - ), snd (conf (Good g_close)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hfar. - destruct Hfar. - rewrite filter_In, config_list_spec, in_map_iff. - rewrite 2 andb_true_iff, Nat.ltb_lt. - repeat split. - exists (Good g_close). - split; auto. - destruct (change_frame da conf g_ex) as ((?,?),[|]); - reflexivity. - apply In_names. - rewrite Hconf_ex; unfold get_ident in *; simpl in *; lia. - unfold get_alive in *; simpl in *; auto. - rewrite negb_true_iff in *; unfold get_based in *; simpl in *;auto. - unfold get_location, State_ILA, AddInfo, OnlyLocation, - get_location, Datatypes.id in *; - rewrite (frame_dist _ _ (change_frame da conf g_ex)) - in Hdist_close. - rewrite dist_sym, Hconf_ex. - destruct (change_frame _ _ g_ex) - as ((?,?),[|]) eqn : Hchange_ex; simpl in *; - apply Hdist_close. - + rewrite <- negb_true_iff, forallb_existsb in Hdist_DP_D. - assert (Hexists : exists g, get_based (conf (Good g)) = true /\ - forall g', Good g' <> Good g -> get_based (conf (Good g')) = true -> - (get_ident (conf (Good g))) < (get_ident (conf (Good g')))). - (* il existe un minimum aux identifiants des robots placé a la bases *) - assert (Hrobin := robin Identifiers (fun id => get_based (conf (id)) = true)). - assert (Hrobin_aux : forall id : ident, {get_based (conf id) = true} + {get_based (conf id) <> true}). - intros [g_ro|byz]; - try ( - assert (Hfalse := In_Bnames byz); - now simpl in *). - destruct (get_based (conf (Good g_ro))); intuition. - specialize (Hrobin Hrobin_aux). - clear Hrobin_aux. - assert (Hrobin_aux : exists id : ident, get_based (conf id) = true). - destruct Hat_based as (g_at, Hat_based); exists (Good g_at). - apply Hat_based. - specialize (Hrobin Hrobin_aux conf). - clear Hrobin_aux. - destruct (Hrobin) as ([g_ro|byz], (?,?)); - try ( - assert (Hfalse := In_Bnames byz); - now simpl in *). - exists (g_ro). - split; auto. - rewrite forallb_forall in *. - destruct Hexists as (g_e, (Hbased_e, Hsmall)). - assert (Hbased_e_r: (get_based (round rbg_ila da conf (Good g_e))) = false). - rewrite round_good_g; try apply Hpred. - simpl. - unfold upt_aux, map_config in *. - destruct (conf (Good g_e)) as (p_e, (((i_e, l_e), a_e), b_e)) eqn : Hconf_e. - simpl. - unfold get_based in *; simpl in Hbased_e; rewrite Hbased_e. - destruct (upt_robot_too_far _ _) eqn : Htoo_far; try auto. - * unfold upt_robot_too_far in *. - destruct (R2dec_bool _ _) eqn : Hpos; try discriminate. - destruct (forallb _ _) eqn : Htoo_far_1 in Htoo_far; - try discriminate. - destruct (forallb _ _) eqn : Htoo_far_2 in Htoo_far; try discriminate. - unfold too_far_aux_2 in *. - rewrite <- forallb_existsb in Htoo_far_2. - rewrite negb_false_iff, existsb_exists in Htoo_far_2. - destruct (Htoo_far_2) as (g_n, (_, Hdist)). - rewrite 2 andb_true_iff in *. - rewrite negb_true_iff in *. - destruct Hdist as (Hbased_n, (Halive_n, Hdist_n)). - specialize (Hdist_DP_D g_n (In_Gnames g_n)). - rewrite negb_true_iff, 2 andb_false_iff in Hdist_DP_D. - destruct Hdist_DP_D as [[Halive_false|Hbased_false]|Hdist_false]. - unfold get_alive in *; simpl in *; rewrite Halive_false in *; auto. - unfold get_based in *; simpl in *; rewrite negb_false_iff, Hbased_false in *; auto. - rewrite Rle_bool_false_iff in *; destruct Hdist_false. - assert ((0,0)%R = ((get_location (conf (Good g_e))))). - rewrite R2dec_bool_true_iff in Hpos. - symmetry. - rewrite <- (retraction_section (frame_choice_bijection (change_frame da conf g_e))). - simpl in *; auto. - rewrite H. - rewrite Rle_bool_true_iff in *. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite (frame_dist _ _ (change_frame da conf g_e)); simpl in *; auto. - unfold too_far_aux_1 in *. - rewrite <- negb_true_iff, existsb_forallb, existsb_exists in Htoo_far_1. - destruct Htoo_far_1 as (g_f, (Hin_f, Hfalse)). - rewrite negb_true_iff in *. - destruct (negb (get_ident _ =? get_ident _)) eqn : Hsame_g_f in Hfalse; - try (simpl in *; discriminate). - rewrite Nat.ltb_nlt in Hfalse. - destruct Hfalse. - specialize (Hdist_DP_D g_f (In_Gnames g_f)). - rewrite negb_true_iff, 2 andb_false_iff in *. - rewrite filter_In in Hin_f. - rewrite andb_true_iff in Hin_f. - unfold get_based, get_alive in *; simpl in Hin_f; rewrite Nat.eqb_neq in Hsame_g_f. - destruct (ident_EqDec (Good g_f) (Good g_e)). - unfold get_ident in Hsame_g_f; simpl in Hsame_g_f; - rewrite e in Hsame_g_f; destruct Hsame_g_f; reflexivity. - destruct Hin_f as (_,(Hb_f,_)). - specialize (Hsmall g_f c Hb_f). - unfold get_ident in *; rewrite Hconf_e; simpl in *; auto. - simpl. - specialize (Hap0 g_e). - rewrite Hconf_e in Hap0; simpl in Hap0. - specialize (Hap0 Hbased_e). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite R2dec_bool_false_iff in Hpos. - destruct Hpos. - rewrite Hconf_e. - destruct Hap0 as (Ha, Hp0). - rewrite <- Hp0. - rewrite <- (@retraction_section R2 _ (@frame_choice_bijection Loc _ _ (change_frame da conf g_e)) p_e) at 3. - now simpl. - * exists g_e. - repeat split; - auto. - ++ rewrite round_good_g; try apply Hpred. - simpl. - unfold upt_aux, map_config. - destruct (conf (Good g_e)) as (p_e, (((i_e, l_e), a_e), b_e)) eqn : Hconf_e. - simpl. - unfold get_based in Hbased_e; simpl in Hbased_e; rewrite Hbased_e. - destruct (upt_robot_too_far _) eqn : Htoo_far; - unfold get_alive; simpl; try auto. - specialize (Hap0 g_e); - unfold get_based, get_alive in *; rewrite Hconf_e in *; simpl in *; specialize (Hap0 Hbased_e); - intuition. - ++ rewrite round_good_g in *; try apply Hpred. - simpl in *. - unfold upt_aux, map_config in *. - destruct (conf (Good g_e)) as (p_e, (((i_e, l_e), a_e), b_e)) eqn : Hconf_e. - simpl in *. - unfold get_based in Hbased_e; simpl in Hbased_e; rewrite Hbased_e in *. - destruct (upt_robot_too_far _) eqn : Htoo_far; - unfold get_alive; simpl; try auto. - specialize (Hap0 g_e); - unfold get_based, get_alive in *; rewrite Hconf_e in *; simpl in *; - specialize (Hap0 (reflexivity _)); - intuition. - rewrite retraction_section, H0. - simpl. unfold Dp; generalize D_0, Dmax_7D. - replace ((0 + - 0) * (0 + - 0) + (0 + - 0) * (0 + - 0))%R with 0%R by lra. - rewrite sqrt_0. - simpl. lra. - now unfold get_based in *; simpl in Hbased_e_r. - - intros g g' Hbased_r Hbased_r'. - assert (Hbased : get_based (conf (Good g)) = true). - apply still_based_means_based in Hbased_r; try apply Hpred; auto. - assert (Hng : g <> g'). - intro Hf. - rewrite Hf in Hbased_r. - rewrite Hbased_r in *; discriminate. - rewrite 2 round_good_g in *; try apply Hpred. - simpl (let _ := _ in _) in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf; - destruct (conf (Good g')) as (pos', (((ident', light'), alive'), based')) eqn : Hconf'. - unfold upt_aux, map_config in *. - rewrite Hconf', Hconf in *. - unfold get_based in Hbased; simpl in Hbased; try rewrite Hbased in *. - unfold get_based in Hbased_r; simpl in Hbased_r. - destruct (upt_robot_too_far _ _) eqn : Htoo_far_g in Hbased_r; - simpl in Hbased_r; try discriminate. - simpl (fst _) ; simpl (snd _). - rewrite Htoo_far_g. - simpl (fst _) in Hbased_r'. - destruct (based'); simpl in *. - destruct (upt_robot_too_far _ _) eqn : Htoo_far_g' in Hbased_r'; - simpl in Hbased_r; try (unfold get_based in *; discriminate). - assert (Htoo_far_g'_aux := Htoo_far_g'). - unfold upt_robot_too_far in Htoo_far_g'. - destruct (R2dec_bool _ _); try (simpl; discriminate). - destruct (forallb _ _) eqn : Htoo_far1 in Htoo_far_g'; try (simpl; discriminate); - assert (Htoo_far1_aux := Htoo_far1). - unfold too_far_aux_1 in Htoo_far1. - rewrite forallb_forall in Htoo_far1. - unfold get_based, get_alive in Htoo_far1; simpl (snd _) in Htoo_far1. - assert (Hin_g: List.In g - (List.filter - (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g). - rewrite Hconf. - simpl. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in Hap0; unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - specialize (Htoo_far1 g Hin_g). - destruct (negb _) eqn : Hnegb_1 in Htoo_far1; try discriminate. - destruct (upt_robot_too_far _ _); - unfold get_ident in *; simpl in *; rewrite Nat.ltb_lt in *; simpl in *; auto. - destruct (upt_robot_too_far _ _). - simpl in *; rewrite Hconf, Hconf' in *; simpl in *; auto. - simpl in *; rewrite Hconf, Hconf' in *; simpl in *; auto. - destruct (upt_robot_too_far _ _); - simpl in *; rewrite Hconf, Hconf' in *; simpl in *; auto. - unfold get_ident; simpl. - rewrite negb_false_iff, Nat.eqb_eq in Hnegb_1. - rewrite Htoo_far_g'_aux. - simpl. - destruct Hbased_higher as (?,(?,?)). - unfold too_far_aux_1 in *. - assert (Hfalse := ident_unique conf Hng). - unfold get_ident in *; simpl in *; rewrite Hnegb_1 in *; auto. - now destruct Hfalse. - destruct Hbased_higher as (?,(?,?)). - specialize (H1 g g'); rewrite Hconf, Hconf' in *. - unfold get_based in *; simpl in *; specialize (H1 (reflexivity _) (reflexivity _)). - unfold get_ident in *; simpl in *. - destruct (upt_robot_dies_b _); simpl; lia. -Qed. - - - - -Lemma light_off_means_alive_next : forall conf da g, - da_predicat da -> - path_conf conf -> - get_alive (conf (Good g)) = true -> - get_light (conf (Good g)) = false -> - executed_means_light_on conf -> - get_alive (round rbg_ila da conf (Good g)) = true. -Proof. - intros conf da g Hpred Hpath Halive Hlight Hexec. - destruct (get_alive (round rbg_ila da conf (Good g))) eqn : Halive_round; try auto. - specialize (Hexec g da Hpred Halive Halive_round). - rewrite Hexec in *; discriminate. -Qed. - - - -Lemma round_target_alive : - forall conf da, - da_predicat da -> - path_conf conf-> - based_higher conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - target_alive conf -> - target_alive (round rbg_ila da conf). -Proof. - intros conf da Hpred Hpath Hbased_higher Hcoll Hexecutioner_off Hexecuted_on Hexists_at Ht_alive g Hn_0 Halive. - unfold target_alive in *. - rewrite <- (ident_preserved conf g Hpred) in Hn_0. - specialize (Ht_alive g Hn_0 (still_alive_means_alive g Hpred Hbased_higher Halive)). - assert (Halive_bis := Halive). - rewrite round_good_g in Halive; try apply Hpred. - simpl in Halive. - unfold upt_aux, map_config in Halive. - destruct (conf (Good g)) as (pos,(((ident, light), alive), based)) eqn : Hconf. - destruct based eqn : Hbased. - *** - simpl in Halive. - destruct (get_based (round rbg_ila da conf (Good g))) eqn : Hbased_r. - - assert (Hexists_aux : exists g, get_based (conf (Good g)) = true). - { exists g. - rewrite Hconf; simpl; auto. - } - assert (Hbased_higher_round := based_higher_round Hpred Hpath Hbased_higher Hexecutioner_off Hexecuted_on Hexists_aux). - destruct Hbased_higher_round as (Hap0_r, (Hex_r, Hhigher_r)). - destruct Hex_r as (g_ex, (Hbased_ex_r, (Halive_ex_r, Hdist_ex_r))). - exists g_ex. - repeat split; auto. - specialize (Hap0_r g Hbased_r). - destruct Hap0_r as (Ha_r, Hp_r). - rewrite <- Hp_r in Hdist_ex_r. - unfold Dp in *; generalize D_0, Dmax_7D. - transitivity (Dmax - D)%R; try lra. - rewrite dist_sym in Hdist_ex_r. - lra. - - destruct (Hbased_higher) as (Hap0, ((g_ex, (Hbased_ex, (Halive_ex, Hdist_ex))), Hhigher)). - destruct (get_alive (round rbg_ila da conf (Good g_ex))) eqn : Halive_ex_r. - * exists g_ex. - repeat split; auto. - rewrite <- 2 ident_preserved; try auto. - assert (Hbased_aux : get_based (conf (Good g)) = true) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhigher g g_ex Hbased_aux Hbased_ex). - auto. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (0,0)%R (get_location (conf (Good g_ex)))). - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_ex)))). - transitivity (dist (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) + - dist (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_ex))))%R; try auto. - transitivity (dist (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) + D)%R. - apply Rplus_le_compat_l. - assert (Hround := dist_round_max_d g_ex Hpred Hpath Halive_ex). - unfold equiv, bool_Setoid, eq_setoid in Hround. - rewrite Rle_bool_true_iff in Hround. - lra. - replace Dmax with (Dp + D)%R. - apply Rplus_le_compat_r. - transitivity ( dist (get_location (round rbg_ila da conf (Good g))) (0, 0) + - dist (0, 0) (get_location (conf (Good g_ex))))%R; try auto. - replace Dp with (D + (Dp - D))%R by lra. - transitivity (D + dist (0,0) (get_location (conf (Good g_ex))))%R. - apply Rplus_le_compat_r. - assert (Hbased_aux : get_based (conf (Good g)) = true) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hap0 g Hbased_aux). - destruct Hap0 as (_, Hp0). rewrite <- Hp0, dist_sym. - assert (Hround := dist_round_max_d g Hpred Hpath (still_alive_means_alive g Hpred Hbased_higher Halive_bis)). - unfold equiv, bool_Setoid, eq_setoid in Hround. - rewrite Rle_bool_true_iff in Hround. - lra. - apply Rplus_le_compat_l. - now rewrite dist_sym. - unfold Dp; lra. - * - assert (get_ident (conf (Good g_ex)) <> 0). - intro Hfalse_aux. - rewrite (ident_preserved conf g_ex Hpred) in Hfalse_aux. - assert (Hfalse := ident_0_alive (round rbg_ila da conf) g_ex Hfalse_aux). - now rewrite Hfalse in *. - destruct (Hpath g_ex Halive_ex) as [Hfalse|Hpath']. - contradiction. - assert (exists g0 : G, - get_alive (conf (Good g0)) = true /\ - get_based (conf (Good g0)) = false /\ - get_ident (conf (Good g0)) < get_ident (conf (Good g_ex)) /\ - (dist (get_location (conf (Good g0))) (get_location (conf (Good g_ex))) <= Dmax)%R). - destruct Hpath' as (gp, Hp); exists gp; intuition. - rewrite dist_sym; auto. - assert (Htest := executed_means_alive_near g_ex Hpred Halive_ex H Halive_ex_r Hpath H0). - clear H0. - destruct (Htest) as (g_test, (Hident_test, (Hdist_test, (Hbased_test, Halive_test)))). - exists g_test. - repeat split. - apply light_off_means_alive_next; try auto. - apply (Hexecutioner_off g_test da Hpred); try auto. - exists g_ex. - repeat split; try auto. - now rewrite dist_sym. - rewrite <- 2 ident_preserved; try auto. - assert (Hbased_aux : get_based (conf (Good g)) = true) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhigher g g_test Hbased_aux Hbased_test). - auto. - assert ((get_location (round rbg_ila da conf (Good g))) = (0,0)%R). - rewrite round_good_g. - simpl. - unfold upt_aux, map_config. - rewrite Hconf; simpl. - assert (Hbased_aux : get_based (conf (Good g)) = true) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hap0 g Hbased_aux). - destruct Hap0 as (_, Hp0). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - destruct (upt_robot_too_far _); simpl; rewrite retraction_section; - now rewrite <- Hp0, Hconf; simpl in *. - apply Hpred. - rewrite H0; clear H0. - assert (Htri := RealMetricSpace.triang_ineq (0,0)%R (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_test)))). - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (conf (Good g_ex))) (get_location (conf (Good g_test))) (get_location (round rbg_ila da conf (Good g_test)))). - transitivity (dist (0, 0) (get_location (conf (Good g_ex))) + - dist (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_test))))%R; - try auto. - replace Dmax with (((Dp - D) + D) + D)%R. - transitivity (((Dp - D) + - dist (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_test))))%R). - apply Rplus_le_compat_r. - now rewrite dist_sym. - replace (Dp - D + D + D)%R with ((Dp - D) + (D + D))%R by lra. - apply Rplus_le_compat_l. - transitivity (dist (get_location (conf (Good g_ex))) (get_location (conf (Good g_test))) + - dist (get_location (conf (Good g_test))) (get_location (round rbg_ila da conf (Good g_test))))%R; try auto. - transitivity (D+ - dist (get_location (conf (Good g_test))) (get_location (round rbg_ila da conf (Good g_test))))%R. - apply Rplus_le_compat_r. - auto. - apply Rplus_le_compat_l. - assert (Hround := dist_round_max_d g_test Hpred Hpath Halive_test). - unfold equiv, bool_Setoid, eq_setoid in Hround; - rewrite Rle_bool_true_iff in *; auto. - unfold Dp; generalize D_0; lra. -*** - destruct (upt_robot_dies_b _) eqn : Hbool. - - unfold get_alive in Halive; now simpl in Halive. - - unfold rbg_fnc in Halive. - set (new_pos := choose_new_pos _ _) in *. - destruct (move_to _) eqn : Hmove. - + assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_choose_dp, Hdist_chose_d). - assert (Hmove_Some := move_to_Some_zone Hmove). - set (target := choose_target _) in *. - assert (Htarget_in_obs := @choose_target_in_obs _ target (reflexivity _)). - specialize (Hmove_Some _ Htarget_in_obs). - unfold obs_from_config, Obs_ILA in Htarget_in_obs. - rewrite make_set_spec, filter_InA, config_list_InA in Htarget_in_obs. - * rewrite 3 andb_true_iff, Rle_bool_true_iff in Htarget_in_obs. - destruct Htarget_in_obs as (([g_target|b_target],Htarget_spec), - (((Htarget_Dmax, Htarget_alive), Htarget_aux), Htarget_ident)); - try (assert (Hfalse := In_Bnames b_target); - now simpl in * ). - (* si target est allumé, tous els robots visibles sont allumés, et il exists donc un robot à moins de Dp (cf Hexists_at). donc target est à moins de Dp, donc elle sera vivante au tours prochain car: - elle est éteinte donc elle ne bouge pas. - elle ne peux pas mourrir car son bourreau potentiel serait à moins de D de elle, or un bourreau est éteint, et donc est à plus de Dmax de nous. donc le bourreau doit etre à plus de D de target, contradiction. - - - si target est étiente, elle ne meurt aps au prochain tour (executed_mean_light_on) donc comme on à Hdist_Dp, on a dist new_pos (round ... g_target) <= Dmax. - - *) - destruct (get_light target) eqn : Htarget_light. - -- assert (Hall_lighted := @light_off_first _ target (reflexivity _) Htarget_light). - assert (Hless_that_Dp := Hexists_at). - specialize (Hexists_at g (still_alive_means_alive g Hpred Hbased_higher Halive_bis)). - revert Hexists_at; intro Hexists_at. - assert (Hforall_lighted : (forall g_near, - get_alive (conf (Good g_near)) = true -> - (dist (get_location (conf (Good g_near))) (get_location (conf (Good g))) <= Dmax)%R -> - get_ident (conf (Good g_near)) < get_ident (conf (Good g)) -> - get_light (conf (Good g_near)) = true)). - { - unfold For_all in Hall_lighted. - intros g_near Halive_near Hdist_near Hident_near. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - specialize (Hall_lighted ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near)))). - apply Hall_lighted. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g_near). - destruct b; reflexivity. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_near. - unfold frame_choice_bijection, Frame in Hdist_near. unfold Datatypes.id in *. - simpl (_ ∘ _) in Hdist_near. - rewrite Hconf in Hdist_near. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - destruct b; simpl in *; lra. - unfold get_alive in *; now simpl in *. - repeat rewrite andb_true_iff. - unfold get_alive in *; now simpl in *. - - rewrite Nat.leb_le, <-Hconf . - unfold get_ident in *; simpl in *. - lia. - intros x y Hxy. - rewrite (fst_compat Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - } - assert (H_not_0 : get_ident (pos, ((ident, light, alive), based)) > 0). - { - - rewrite (Nat.neq_0_lt_0 _) in Hn_0. - unfold get_ident in *; simpl in *; lia. - - } - rewrite Hconf in *. - specialize (Hexists_at H_not_0 Hforall_lighted). - assert (Hexists_at_less_Dp := Hexists_at). - - destruct Hexists_at. - - assert (Hlight_close := @light_close_first _ target (reflexivity _) Htarget_light). - (* par rapprot à notre position ACTUELE elle est à moins de Dp. *) - assert (dist (0,0) (get_location target) <= Dp)%R. - destruct (Rle_dec (dist (0,0) (@get_location Loc _ _ target))%R Dp); try auto. - specialize (Hlight_close (Rnot_le_gt _ _ n0)). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - specialize (Hlight_close ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good x))), snd (conf (Good x)))). - assert (Hfalse : ~ (dist (get_location (conf (Good g))) (get_location (conf (Good x))) <= Dp)%R). - { - apply Rgt_not_le. - rewrite (frame_dist _ _ (r,t,b)). - assert (Hpred_bis := Hpred). - revert Hpred; intro Hpred; destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - rewrite Horigin. - assert (Hin : @In (prod R2 ILA) - (@state_Setoid (@make_Location R2 R2_Setoid R2_EqDec) (prod R2 ILA) State_ILA) - (@state_EqDec (@make_Location R2 R2_Setoid R2_EqDec) (prod R2 ILA) State_ILA) - (@FSetList.SetList (prod R2 ILA) - (@state_Setoid (@make_Location R2 R2_Setoid R2_EqDec) - (prod R2 ILA) State_ILA) - (@state_EqDec (@make_Location R2 R2_Setoid R2_EqDec) (prod R2 ILA) State_ILA)) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers x)))) - (@snd R2 ILA (conf (@Good Identifiers x)))) - (@obs_from_config (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) (@fst R2 ILA (conf id))) (@snd R2 ILA (conf id))) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) - (@fst R2 ILA - (@pair R2 ILA pos - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident light) - alive) false)))) - (@snd R2 ILA - (@pair R2 ILA pos - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident light) - alive) false)))))). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA. - rewrite (@config_list_InA Loc _ State_ILA), 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good x). - destruct b; reflexivity. - rewrite (frame_dist _ _ (r,t,b)) in H. - unfold frame_choice_bijection, Frame in H. unfold Datatypes.id in *. - simpl (_ ∘ _) in H. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - destruct b; unfold Dp in H; generalize Dmax_7D, D_0; rewrite dist_sym; simpl in *; lra. - unfold get_alive in *; now simpl in *. - unfold get_alive in *; now simpl in *. - rewrite Nat.leb_le. - rewrite Hconf in *. - unfold get_ident in *; simpl in *. - lia. - intros x' y Hxy. - rewrite (fst_compat Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - specialize (Hlight_close Hin). - destruct Hlight_close as [Hlight_close|Hlight_close]. - simpl in *; destruct b; apply Hlight_close. - assert (x <> g). - { - destruct (Geq_dec x g). - rewrite e in H. - rewrite Hconf in H. - unfold get_ident in H; simpl in H; auto. - lia. - auto. - } - specialize (Hcoll x g H0). - rewrite Hconf in Hcoll. - unfold get_based in Hcoll. - simpl (snd _) in Hcoll. - assert (get_based (conf (Good x)) = false). - { - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (_,(_,Hbh)). - destruct (get_based (conf (Good x))) eqn : Hfalse_based. - specialize (Hbh x g Hfalse_based). - rewrite Hconf in Hbh. - unfold get_based in Hbh; simpl (snd _) in Hbh; specialize (Hbh (reflexivity _)). - lia. - reflexivity. - } - unfold get_based in H1; specialize (Hcoll H1 (reflexivity _)). - destruct H as (?,?). rewrite <- Hconf in Hcoll. - specialize (Hcoll H (still_alive_means_alive g Hpred_bis Hbased_higher Halive_bis)). - rewrite <- Horigin in Hlight_close. - rewrite (frame_dist _ _ (r,t,b)) in Hcoll. - revert Hcoll; intro Hcoll. - rewrite <- Hlight_close in Hcoll. - destruct Hcoll. - simpl (frame_choice_bijection _); simpl (get_location _). - unfold Datatypes.id. - assert (Htrue := dist_same (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good x))))). - rewrite <- Htrue. - repeat f_equiv. - destruct b; auto. - } - destruct H as (?,(?,?)); rewrite Hconf in *; contradiction. - exists g_target. - repeat split. - ++ destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - rewrite round_good_g; try apply Hpred. - simpl. - unfold upt_aux, map_config. - destruct (upt_robot_dies_b _ g_target) eqn : Hbool_target. - ** assert (Halive_target_round : - get_alive (round rbg_ila da conf (Good g_target)) = false). - { rewrite round_good_g; try apply Hpred. - simpl. - unfold upt_aux, map_config. - assert (get_based (conf (Good g_target)) = false). - destruct (get_based (conf (Good g_target)))eqn : Hfalse; auto. - destruct (Hbased_higher) as (Hap0, (Hex, Hhig)). - assert (Hbased_aux : get_based (conf (Good g)) = false) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhig g_target g Hfalse Hbased_aux). - rewrite Htarget_spec in Htarget_ident. - unfold equiv, bool_Setoid, eq_setoid in Htarget_ident. - unfold get_ident in *; rewrite Nat.leb_le , Hconf, Hconf_target in *; simpl in *. - lia. - rewrite Hconf_target in *; unfold get_based in *; simpl in H1; rewrite H1. - simpl. - now rewrite Hbool_target; unfold get_alive; simpl. - } - assert ( Hsame_pos_round_target : - get_location (conf (Good g_target)) - == get_location (round rbg_ila da conf (Good g_target))). - { apply executed_means_not_moving; try auto. - rewrite Htarget_spec in Htarget_alive. - unfold get_alive in *; rewrite Hconf_target; now simpl. - } - unfold upt_robot_dies_b in Hbool_target. - rewrite orb_true_iff in Hbool_target. - destruct Hbool_target as [Htoo_close_so_lighted|Htoo_far_but_path_conf]. - - - rewrite existsb_exists in *. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - destruct Htoo_close_so_lighted as (too_close, (Hin_too_close, Hdist_too_close)). - +++ unfold executioner_means_light_off in *. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff in Hin_too_close. - destruct Hin_too_close as (([g_too_close | byz], (Hspec_too_close, Hnames_too_close)), ((Hident_too_close, Halive_too_close), Hbased_too_close)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - specialize (Hexecutioner_off g_too_close da Hpred). - rewrite <- Hspec_too_close in Halive_too_close. - unfold get_alive in *. - simpl (snd (snd _)) in *. - specialize (Hexecutioner_off Halive_too_close). - assert (Hex : (exists g' : G, - snd (fst (snd (conf (Good g')))) = true /\ - get_based (conf (Good g')) = false /\ - get_ident (conf (Good g_too_close)) < get_ident (conf (Good g')) /\ - (dist (get_location (conf (Good g_too_close))) - (get_location (conf (Good g'))) <= D)%R)). - { exists g_target; - repeat split; try now simpl. - rewrite <- Htarget_alive. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd; destruct Htarget_spec_snd as (_,(_,(Htarget_spec_alive,_))). - rewrite Htarget_spec_alive. - now rewrite Hconf_target; simpl. - destruct (get_based (conf (Good g_target)))eqn : Hfalse; auto. - destruct (Hbased_higher) as (Hap0, (Hex, Hhig)). - assert (Hbased_aux : get_based (conf (Good g)) = false) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhig g_target g Hfalse Hbased_aux). - rewrite Htarget_spec in Htarget_ident. - unfold equiv, bool_Setoid, eq_setoid in Htarget_ident. - unfold get_ident in *; rewrite Nat.leb_le , Hconf, Hconf_target in *; simpl in *. - lia. - rewrite Nat.ltb_lt in Hident_too_close. - rewrite <- Hspec_too_close in Hident_too_close. - unfold get_ident in *; simpl in *; auto. - rewrite <- Hspec_too_close, Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite (frame_dist _ _ (change_frame da conf g_target)). - - now simpl in *. - } - specialize (Hexecutioner_off Hex). - clear Hex. - - assert (Hlight_off_first := @light_off_first _ target (reflexivity _) Htarget_light). - specialize (Hlight_off_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))). - unfold equiv, bool_Setoid, eq_setoid in Hlight_off_first. - rewrite <- Hlight_off_first, <- Hexecutioner_off. - assert (Hbased_target : get_based (conf (Good g_target))= false). - { destruct (get_based (conf (Good g_target)))eqn : Hfalse; auto. - destruct (Hbased_higher) as (Hap0, (Hex, Hhig)). - assert (Hbased_aux : get_based (conf (Good g)) = false) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhig g_target g Hfalse Hbased_aux). - rewrite Htarget_spec in Htarget_ident. - unfold equiv, bool_Setoid, eq_setoid in Htarget_ident. - unfold get_ident in *; rewrite Nat.leb_le , Hconf, Hconf_target in *; simpl in *. - lia. - } - unfold get_based in Hbased_target; rewrite Hconf_target in *; simpl in Hbased_target; rewrite Hbased_target. - - unfold get_light. now simpl in *. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. repeat split. - exists (Good g_too_close). - destruct b; reflexivity. - rewrite 2 andb_true_iff. - rewrite Rle_bool_true_iff. - replace (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close))))%R - with - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))))%R. - unfold Datatypes.id. - assert (Hframe := frame_dist (fst (conf (Good g_too_close))) pos ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (dist (fst (conf (Good g_too_close))) pos <= Dmax)%R. - rewrite Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite <- Hspec_too_close in Hdist_too_close. - assert ((dist (fst (conf (Good g_too_close))) - (fst (conf (Good g_target))) <= D)%R). - rewrite (frame_dist _ _ (change_frame da conf g_target)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - now simpl; simpl in Hdist_too_close. - (* *) - specialize (Hless_that_Dp g). - unfold get_alive in Hless_that_Dp. - simpl in Halive. - rewrite Hconf, Halive in Hless_that_Dp; - simpl (snd ( _)) in Hless_that_Dp. - specialize (Hless_that_Dp (reflexivity _)). - destruct (Rle_bool (dist pos (fst (conf (Good g_target)))) Dp) eqn : Hhow_far. - rewrite Rle_bool_true_iff, dist_sym in Hhow_far. - assert (Hti := RealMetricSpace.triang_ineq (fst (conf (Good g_too_close))) (fst (conf (Good g_target))) pos ). - rewrite Hconf_target in Hti at 2. - simpl ( (fst _)) in Hti. - rewrite dist_sym in H1. - generalize (D_0), Dmax_7D. - unfold Dp in *. - rewrite Hconf_target in *; simpl (fst _) in *. - rewrite dist_sym in H1. - lra. - rewrite Rle_bool_false_iff in Hhow_far. - revert H0; intro H0. - destruct Hhow_far. - destruct Hpred as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite <- Horigin, Htarget_spec in H0. - rewrite (frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame in *. - fold Frame in *. rewrite Hchange in *. - rewrite Hconf_target, Hconf in *. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - now destruct b; simpl in *. - split ; try (destruct b; now rewrite <- Hframe). - unfold get_alive; simpl in *; auto. - now simpl in *. - rewrite Nat.leb_le. - transitivity (get_ident (conf (Good g_target))). - rewrite Nat.ltb_lt, <- Hspec_too_close in Hident_too_close. - unfold get_ident in *; simpl in *; lia. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intro Horigin. - set (obs := obs_from_config _ _) in *. - assert (Htarget_ident' := @target_lower obs target). - specialize (Htarget_ident' (((frame_choice_bijection (change_frame da conf g)) (get_location (conf (Good g))) ), snd(conf (Good g)))). - rewrite Htarget_spec in Htarget_ident'. - rewrite Hconf_target, Hconf in *. - apply Nat.lt_le_incl. - unfold get_ident in *; simpl in *; apply Htarget_ident'. - unfold obs. - rewrite Hchange. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, (@config_list_InA Loc _ State_ILA), 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g). - destruct b; rewrite Hconf; simpl; repeat split; reflexivity. - unfold Datatypes.id. - - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - - intro Hdist_0. - - simpl in Hdist_0; simpl. - destruct b; simpl in *; rewrite Hdist_0. - lra. - lra. - unfold get_alive; simpl in *; assumption. - unfold get_alive; simpl in *; assumption. - rewrite Nat.leb_le. - reflexivity. - intros x1 y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - rewrite Hcompat. - reflexivity. - rewrite Hchange in *. - rewrite <- Horigin. - destruct b; auto. - fold target. - apply Htarget_spec. - intros x1 y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - rewrite Hcompat. - reflexivity. - +++ - specialize (Hpath g_target). - assert (get_alive (conf (Good g_target)) == true). - rewrite <- Htarget_alive. - rewrite Htarget_spec. - rewrite Hconf_target. - unfold get_alive; now simpl. - simpl in H1. - specialize (Hpath H1); clear H1. - destruct Hpath as [Hident_0|Hpath_backup]. - rewrite (ident_preserved _ _ Hpred) in Hident_0. - assert (get_alive (round rbg_ila da conf (Good g_target)) = true). - apply ident_0_alive; intuition. - rewrite H1 in *; discriminate. - rewrite forallb_existsb, forallb_forall in Htoo_far_but_path_conf. - destruct Hpath_backup as (g_too_close, (Halive_close,( Hdist_close, (Hident_close, Hbased_closed)))). - specialize (Htoo_far_but_path_conf - ((((let (p, b) := change_frame da conf g_target in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))))). - rewrite negb_true_iff, Rle_bool_false_iff in Htoo_far_but_path_conf. - destruct Htoo_far_but_path_conf. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - **** exists (Good g_too_close). - split. - destruct (change_frame da conf g_target) as (?,b0); destruct b0; - now simpl. - apply In_names. - **** rewrite Nat.ltb_lt. - unfold get_ident in *; simpl in Hident_close; simpl. - auto. - **** rewrite <- Halive_close. - now unfold get_alive; simpl. - **** unfold get_based in *; simpl in *; rewrite negb_true_iff; auto. - **** rewrite dist_sym, (frame_dist _ _ (change_frame da conf g_target)) in Hdist_close. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - unfold frame_choice_bijection, Frame in Hdist_close; fold Frame in *. - revert Hdist_close; intros Hdist_close. - destruct (change_frame _ _ g_target) as ((r_f, t_f), b_f) - eqn : Hchange_far. - now destruct b_f; simpl in *. - ** unfold get_alive in *; simpl. - rewrite Hconf_target. - simpl. - rewrite <- Htarget_alive. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd; destruct Htarget_spec_snd as (Htarget_spec_ident,(_,(Htarget_spec_alive,Hbased_spec_target))). - simpl. - rewrite Htarget_spec_alive. - assert (Hbased_target : get_based (conf (Good g_target))= false). - { destruct (get_based (conf (Good g_target)))eqn : Hfalse; auto. - destruct (Hbased_higher) as (Hap0, (Hex, Hhig)). - assert (Hbased_aux : get_based (conf (Good g)) = false) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhig g_target g Hfalse Hbased_aux). - unfold equiv, bool_Setoid, eq_setoid in Htarget_ident. - unfold get_ident in *; rewrite Nat.leb_le , Hconf, Hconf_target in *; simpl in *. - rewrite Htarget_spec_ident in *. - lia. - } - unfold get_based in Hbased_target; rewrite Hconf_target in *; simpl in Hbased_target; rewrite Hbased_target. - simpl. - - - reflexivity. - ++ rewrite <- 2 ident_preserved; try apply Hpred. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - set (obs := obs_from_config _ _) in *. - assert (Htarget_ident' := @target_lower obs target). - specialize (Htarget_ident' (((frame_choice_bijection (change_frame da conf g)) (get_location (conf (Good g))) ), snd(conf (Good g)))). - rewrite Htarget_spec in Htarget_ident'. - rewrite Hconf in *. - unfold get_ident in *; simpl in *; apply Htarget_ident'. - unfold obs. - rewrite Hchange. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, (@config_list_InA Loc _ State_ILA), 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g). - destruct b; rewrite Hconf; simpl; repeat split; reflexivity. - unfold Datatypes.id. - - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - - intro Hdist_0. - - simpl in Hdist_0; simpl. - destruct b; simpl in *; rewrite Hdist_0. - lra. - lra. - - unfold get_alive; simpl in *; assumption. - unfold get_alive; simpl in *; assumption. - rewrite Nat.leb_le. - reflexivity. - intros x1 y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - rewrite Hcompat. - reflexivity. - rewrite Hchange in *. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intro Horigin. - rewrite <- Horigin. - rewrite Hconf. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - unfold frame_choice_bijection. - rewrite Hchange. - simpl; destruct b; reflexivity. - fold target. - apply Htarget_spec. - ++ - revert Hdist_choose_dp; intros Hdist_choose_dp. - rewrite round_good_g at 1; try apply Hpred. - simpl (get_location _) at 1. - unfold upt_aux, map_config. - rewrite Hbool. - unfold rbg_fnc. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - rewrite Hconf. - fold target new_pos. - rewrite Hmove. - simpl (fst _). - assert (Hdist_choose_dp_aux : (dist ((retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos)) - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) (get_location target)) <= Dp)%R). - rewrite (frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame. - destruct b; rewrite 2 section_retraction; - assumption. - unfold Datatypes.id in *. - rewrite Htarget_spec in Hdist_choose_dp_aux. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_choose_dp_aux. - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g_target)))). - unfold frame_choice_bijection, Frame in Hret_sec. - assert (Hdist_choose_dp' : (dist (retraction - (comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) new_pos) (fst (conf (Good g_target))) <= Dp)%R). - rewrite <- Hret_sec. - destruct b; simpl in *; apply Hdist_choose_dp_aux. - assert (Hti := RealMetricSpace.triang_ineq (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (fst (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g_target)))). - assert ((dist - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (get_location (round rbg_ila da conf (Good g_target))) <= - Dp + - dist (fst (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g_target))))%R). - lra. - assert ((dist - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (get_location (round rbg_ila da conf (Good g_target)))) <= Dp + D)%R. - assert (Hdist_round := @dist_round_max_d g_target conf da Hpred Hpath). - assert (get_alive (conf (Good g_target)) == true). - rewrite <- Htarget_alive. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - simpl in Htarget_spec_snd. - destruct Htarget_spec_snd as (_,(_,(Htarget_spec_alive,_))). - rewrite Htarget_spec_alive. - simpl; auto. - specialize (Hdist_round H2). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round. - rewrite Rle_bool_true_iff in Hdist_round. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - generalize D_0. - transitivity (Dp + dist (fst (conf (Good g_target))) (fst (round rbg_ila da conf (Good g_target))))%R. - apply H1. - apply Rplus_le_compat_l. - apply Hdist_round. - unfold Dp in *. - replace (Dmax) with (Dmax - D + D)%R by lra. - cbn in *. - changeR2. - destruct b; apply H2. - - -- - exists g_target. - assert (Htarget_alive' : get_alive (conf (Good g_target)) = true). - { - rewrite <- Htarget_alive. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - destruct Htarget_spec_snd as (_,(_,(Htarget_spec_alive,_))). - rewrite Htarget_spec_alive. - simpl; auto. - - } - repeat split. - ++ apply light_off_means_alive_next; try auto. - rewrite <- Htarget_light. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - destruct Htarget_spec_snd as (_,(Htarget_spec_light,_)). - rewrite Htarget_spec_light. - simpl; auto. - ++ -rewrite <- 2 ident_preserved; try apply Hpred. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - set (obs := obs_from_config _ _) in *. - assert (Htarget_ident' := @target_lower obs target). - specialize (Htarget_ident' (((frame_choice_bijection (change_frame da conf g)) (get_location (conf (Good g))) ), snd(conf (Good g)))). - rewrite Htarget_spec in Htarget_ident'. - rewrite Hconf in *. - unfold get_ident in *; simpl in *; apply Htarget_ident'. - unfold obs. - rewrite Hchange. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, (@config_list_InA Loc _ State_ILA), 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g). - destruct b; rewrite Hconf; simpl; repeat split; reflexivity. - unfold Datatypes.id. - - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - - intro Hdist_0. - - simpl in Hdist_0; simpl. - destruct b; simpl in *; rewrite Hdist_0. - lra. - lra. - - unfold get_alive; simpl in *; assumption. - unfold get_alive; simpl in *; assumption. - rewrite Nat.leb_le. - reflexivity. - intros x1 y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - rewrite Hcompat. - reflexivity. - rewrite Hchange in *. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intro Horigin. - rewrite <- Horigin. - rewrite Hconf. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - unfold frame_choice_bijection. - rewrite Hchange. - simpl; destruct b; reflexivity. - fold target. - apply Htarget_spec. - ++ - revert Hdist_choose_dp; intros Hdist_choose_dp. - rewrite round_good_g at 1; try apply Hpred. - simpl (get_location _) at 1. - unfold upt_aux, map_config. - rewrite Hbool. - unfold rbg_fnc. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - rewrite Hconf. - fold target new_pos. - rewrite Hmove. - simpl (fst _). - assert (Hdist_choose_dp_aux : (dist ((retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos)) - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) (get_location target)) <= Dp)%R). - rewrite (frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame. - destruct b; rewrite 2 section_retraction; - assumption. - unfold Datatypes.id in *. - rewrite Htarget_spec in Hdist_choose_dp_aux. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_choose_dp_aux. - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g_target)))). - unfold frame_choice_bijection, Frame in Hret_sec. - assert (Hdist_choose_dp' : (dist (retraction - (comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) new_pos) (fst (conf (Good g_target))) <= Dp)%R). - rewrite <- Hret_sec. - destruct b; simpl in *; apply Hdist_choose_dp_aux. - assert (Hti := RealMetricSpace.triang_ineq (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (fst (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g_target)))). - assert ((dist - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (get_location (round rbg_ila da conf (Good g_target))) <= - Dp + - dist (fst (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g_target))))%R). - lra. - assert ((dist - (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos) - (get_location (round rbg_ila da conf (Good g_target)))) <= Dp + D)%R. - assert (Hdist_round := @dist_round_max_d g_target conf da Hpred Hpath). - assert (get_alive (conf (Good g_target)) == true). - rewrite <- Htarget_alive. - destruct Htarget_spec as (_,Htarget_spec_snd). - destruct target as (p_t,(((i_t,l_t),a_t), b_t)). - simpl in Htarget_spec_snd. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - destruct Htarget_spec_snd as (_,(_,(Htarget_spec_alive,_))). - rewrite Htarget_spec_alive. - simpl; auto. - specialize (Hdist_round H0). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round. - rewrite Rle_bool_true_iff in Hdist_round. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - generalize D_0. - transitivity (Dp + dist (fst (conf (Good g_target))) (fst (round rbg_ila da conf (Good g_target))))%R. - apply H. - apply Rplus_le_compat_l. - apply Hdist_round. - unfold Dp in *. - replace (Dmax) with (Dmax - D + D)%R by lra. - cbn in *. - changeR2. - destruct b; apply H0. - * intros x y Hxy. - rewrite (fst_compat Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - + - assert (Hdist_same : get_location (conf (Good g)) == get_location (round rbg_ila da conf (Good g))). - { - rewrite round_good_g; try auto. - simpl (get_location _). - unfold upt_aux, map_config; rewrite Hbool. - unfold rbg_fnc. - unfold new_pos in *; - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - rewrite Hconf, Hmove. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intro Horigin. - rewrite <- Horigin. - rewrite Hconf. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hchange. - assert (Hret_sec := retraction_section (frame_choice_bijection (r,t,b)) (fst (conf (Good g)))). - unfold frame_choice_bijection, Frame in Hret_sec. - rewrite Hconf in Hret_sec. - rewrite <- Hret_sec. - destruct b; simpl in *; auto. - } - assert (Hmove_None := move_to_None Hmove). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - - assert (Hsame_pos : get_location (round rbg_ila da conf (Good g)) = get_location (conf (Good g))). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config; rewrite Hchange, Hbool in *. - unfold rbg_fnc; unfold new_pos in *; rewrite Hconf, Hmove. - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in *. - simpl (fst _) . - unfold frame_choice_bijection, Frame in Horigin. - rewrite <- Horigin. - rewrite retraction_section. - now cbn. - } - destruct Hmove_None as (other,(Hother_obs, Hmove_other)). - -- unfold obs_from_config, Obs_ILA in Hother_obs. - rewrite make_set_spec, filter_InA, config_list_InA in Hother_obs. - rewrite 3 andb_true_iff, Rle_bool_true_iff in Hother_obs. - destruct Hother_obs as (([g_other|b_other],Hother_spec), - (((Hother_Dmax, Hother_alive), Hother_alive_g), Hother_ident)). - destruct (get_alive (round rbg_ila da conf (Good g_other))) - eqn : Halive_other_round. - ++ exists g_other. - repeat split; try auto. - rewrite Hother_spec in Hother_ident. - rewrite <- 2 ident_preserved; try auto. - unfold get_ident in *; simpl in Hother_ident. - rewrite Nat.leb_le in Hother_ident. - destruct Hmove_other as (Hmove_other, Hdist_other_round_2D). - destruct Hmove_other as (other1, (Hother_in,(*Hother_pos,*) Hother_ident')). - revert Hcoll; intro Hcoll. - unfold no_collision_conf in Hcoll. - unfold obs_from_config, Obs_ILA, map_config in Hother_in. - rewrite make_set_spec, filter_InA, config_list_InA in Hother_in. - rewrite 3 andb_true_iff in Hother_in. - destruct Hother_in as (([g_other'|byz], Hother1_spec), (((Hother1_dist,Hother1_alive),Hother1_aliveg), Hother1_ident)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - (* nouvelle façon de faire *) - - assert (Hident_g :get_ident (conf (Good g_other')) <= get_ident (conf (Good g))). - { rewrite Hother1_spec in Hother1_ident. - unfold get_ident in *; simpl in *; auto. - rewrite Nat.leb_le, Hconf in *. - simpl in *; auto. - } - assert (get_ident (other) < get_ident (other1)). - unfold get_ident in *; auto. - rewrite Hother_spec, Hother1_spec in H. - unfold get_ident in H; simpl in H. - apply (Nat.lt_le_trans (fst (fst (fst (snd (conf (Good g_other)))))) ( fst (fst (fst (snd (conf (Good g_other')))))) (fst (fst (fst (snd (conf (Good g))))))); auto. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - rewrite (fst_compat Hother_spec) in Hother_Dmax. - destruct Hmove_other. - rewrite Hother_spec in H0. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - assert (dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))) <= 3 * D)%R. - unfold map_config in *. - assert (Htri_new_pos := RealMetricSpace.triang_ineq (fst (conf (Good g_other))) (retraction (frame_choice_bijection (r,t,b)) new_pos) (fst (round rbg_ila da conf (Good g)))). - assert (Hnew_correct := choose_new_pos_correct (reflexivity new_pos)). - destruct Hnew_correct as (_,Hdist_new_D). - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hdist_new_D. - assert (Hdist_new_D_aux : (dist (retraction (frame_choice_bijection (r, t, b)) new_pos) - (fst (round rbg_ila da conf (Good g))) <= D)%R). - { unfold ILA in *. - unfold State_ILA in *. - rewrite <- Hdist_same. - rewrite Hconf. - rewrite (frame_dist _ _ (r,t,b)). - rewrite section_retraction. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in Hdist_new_D. - generalize D_0; simpl in *; lra. - } - assert (Hdist_other_new : (dist (fst (conf (Good g_other))) - (retraction (frame_choice_bijection (r, t, b)) new_pos) <= 2*D)%R). - { - rewrite (frame_dist _ _ (r,t,b)), section_retraction. - unfold frame_choice_bijection, Frame; - destruct b; simpl in *; lra. - } - generalize D_0; simpl in *; lra. - assert (Dp > 4*D)%R. - { - generalize Dmax_7D, D_0. unfold Dp. lra. - } - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_other))) (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g)))). - transitivity (dist (fst (round rbg_ila da conf (Good g_other))) (fst (conf (Good g_other))) + - dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))))%R. - auto. - rewrite dist_sym at 1; apply Htri. - assert (Hdist_round := dist_round_max_d g_other Hpred Hpath). - unfold equiv, bool_Setoid, eq_setoid in *; - rewrite Rle_bool_true_iff in *. - rewrite Hother_spec in Hother_alive; unfold get_alive in *; simpl in Hother_alive. - specialize (Hdist_round Hother_alive). - rewrite dist_sym in Hdist_round. - transitivity (D + dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))))%R. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - apply Rplus_le_compat_r. - apply Hdist_round. - transitivity (D + 3*D)%R. - apply Rplus_le_compat_l. - apply H1. - generalize Dmax_7D, D_0. - lra. - ++ assert (Halive_near := @executed_means_alive_near conf g_other da Hpred). - assert (Hother_alive_aux : get_alive (conf (Good g_other)) = true). - { - rewrite <- Hother_alive. - rewrite Hother_spec; - unfold get_alive; simpl; auto. - } - revert Hpath; intro Hpath_backup. - destruct (nat_eq_eqdec (get_ident (conf (Good g_other))) 0). - rewrite (ident_preserved conf g_other Hpred) in e. - assert (Hfalse := ident_0_alive (round rbg_ila da conf) g_other e). - rewrite Hfalse in *; discriminate. - rename c into Hother_else. - specialize (Halive_near Hother_alive_aux Hother_else Halive_other_round Hpath_backup). - destruct (Hpath_backup g_other Hother_alive_aux). - destruct Hother_else. - now rewrite H. - destruct Halive_near as (g_near, (Hnear_ident, (Hnear_dist, (Hnear_based, Hnear_alive)))). - destruct H as (g_0, (Halive_0, (Hdist_0, (Hident_0, Hbased_0)))). - exists g_0; repeat split; auto. - now rewrite dist_sym. - rename H into Hother_path. - assert (get_alive (round rbg_ila da conf (Good g_near)) = true). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config. - destruct (upt_robot_dies_b _ g_near) eqn : Hbool_near. - - assert (Hfalse : get_alive (round rbg_ila da conf (Good g_near)) = false). - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config. - rewrite Hbool_near. - unfold get_alive ; simpl. - destruct (conf (Good g_near)) as (?, (((?,?),?),[|])); unfold get_based in *; simpl; simpl in Hnear_based. - discriminate. - auto. - simpl. - assert (Hlight_false : get_light (conf (Good g_near)) = true). - apply (Hexecuted_on g_near da Hpred Hnear_alive Hfalse). - assert (Hlight_true : get_light (conf (Good g_near)) = false). - apply (Hexecutioner_off g_near da Hpred Hnear_alive). - exists g_other. - repeat split; try auto. - destruct (get_based (conf (Good g_other)))eqn : Hfalse_based; auto. - destruct (Hbased_higher) as (Hap0, (Hex, Hhig)). - assert (Hbased_aux : get_based (conf (Good g)) = false) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhig g_other g Hfalse_based Hbased_aux). - unfold equiv, bool_Setoid, eq_setoid in Hother_ident. - unfold get_ident in *; rewrite Nat.leb_le in Hother_ident. - destruct (Hother_spec) as (?,?). - simpl in H0. - destruct other as (?,(((id_other,?),?),?)). - simpl in *. - destruct (conf (Good g_other)) as (?,(((id_other',?),?),?)). - simpl in *. - destruct H0 as (Hident_other,_). - rewrite Hident_other, Hconf in *. - simpl in *; lia. - - - rewrite dist_sym. auto. - rewrite Hlight_true in *. - discriminate. - - unfold get_alive in *; - destruct (conf (Good g_near)) as (?, (((?,?),?),?)) eqn : Hconf_near; - simpl; auto. - unfold get_based in *; simpl in *; rewrite Hnear_based. - simpl. - auto. - } - exists g_near. - repeat split; auto. - rewrite <- ident_preserved; auto. - transitivity (get_ident (conf (Good g_other))); auto. - destruct Hmove_other as ((copy, (Hcopy_spec, (*(Hcopy_pos,*) Hcopy_ident(*)*))), _). - unfold obs_from_config, Obs_ILA in Hcopy_spec. - rewrite make_set_spec, filter_InA, config_list_InA in Hcopy_spec. - rewrite 3 andb_true_iff, Rle_bool_true_iff in Hcopy_spec. - destruct Hcopy_spec as (([g_inf|byz],Hinf_spec), ((Hinf_dist, Hinf_alive), Hinf)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - rewrite Nat.leb_le in Hinf. - rewrite <- ident_preserved; try auto. - apply (Nat.lt_le_trans _ (get_ident copy) _). - rewrite Hother_spec in Hcopy_ident. - unfold get_ident in *; now simpl in *. - rewrite Hconf in *. - unfold get_ident in *; now simpl in *. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - assert (Hdist_round_g_near := @dist_round_max_d g_near conf da Hpred Hpath_backup Hnear_alive). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round_g_near; - rewrite Rle_bool_true_iff in Hdist_round_g_near. - destruct Hmove_other as (?,Hdist_other). - rewrite Hother_spec, dist_sym in Hdist_other. - revert Hdist_other; intro Hdist_other. - assert (dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))) <= 3 * D)%R. - unfold map_config in *. - assert (Htri_new_pos := RealMetricSpace.triang_ineq (fst (conf (Good g_other))) (retraction (frame_choice_bijection (r,t,b)) new_pos) (fst (round rbg_ila da conf (Good g)))). - assert (Hnew_correct := choose_new_pos_correct (reflexivity new_pos)). - destruct Hnew_correct as (_,Hdist_new_D). - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hdist_new_D. - assert (Hdist_new_D_aux : (dist (retraction (frame_choice_bijection (r, t, b)) new_pos) - (fst (round rbg_ila da conf (Good g))) <= D)%R). - { unfold ILA in *. - unfold State_ILA in *. - rewrite <- Hdist_same. - rewrite Hconf. - rewrite (frame_dist _ _ (r,t,b)). - rewrite section_retraction. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in Hdist_new_D. - apply Hdist_new_D. - } - assert (Hdist_other_new : (dist (fst (conf (Good g_other))) - (retraction (frame_choice_bijection (r, t, b)) new_pos) <= 2*D)%R). - { - rewrite (frame_dist _ _ (r,t,b)), section_retraction. - rewrite dist_sym. - unfold frame_choice_bijection, Frame; - destruct b; simpl in *; lra. - } - generalize D_0; simpl in *; lra. - assert (Dp > 4*D)%R. - { - generalize Dmax_7D, D_0. unfold Dp. lra. - } - unfold Dp in *. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_other))) (get_location (conf (Good g_near)))). - clear Hdist_other. - assert (Htri': (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) <= 4*D)%R). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - replace (4*D)%R with (3*D + D)%R by lra. - transitivity ((dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g_other))) + D)%R); try (now generalize D_0; lra). - transitivity ( - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g_other))) + - dist (fst (conf (Good g_other))) (fst (conf (Good g_near))))%R ; - try auto. - apply Rplus_le_compat_l. - apply Hnear_dist. - apply Rplus_le_compat_r. - rewrite dist_sym. - apply H1. - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_near))) - (get_location (round rbg_ila da conf (Good g_near)))). - unfold Dp. - transitivity (4*D + D)%R. - transitivity (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) + - dist (get_location (conf (Good g_near))) - (get_location (round rbg_ila da conf (Good g_near))))%R. - apply Htri2. - transitivity ((dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) + D))%R. - apply Rplus_le_compat_l. - apply Hdist_round_g_near. - apply Rplus_le_compat_r. - apply Htri'. - generalize Dmax_7D, D_0; lra. - ++ - try (assert (Hfalse := In_Bnames b_other); - now simpl in *). - ++ intros x y Hxy. - rewrite (fst_compat Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. -Qed. - - - -Lemma path_round : - forall conf da, - da_predicat da -> - path_conf conf -> - based_higher conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - path_conf (round rbg_ila da conf). -Proof. - intros conf da Hpred Hpath Hbased_higher Hcoll Hexecutioner_off Hexecuted_on Hexists_at. - unfold path_conf in *. - assert (target_alive conf). - { unfold target_alive. - intros g0 Hn0 Halive0. - specialize (Hpath g0 Halive0). - destruct Hpath. - intuition. - destruct H as (g', (Hpath_alive,( Hpath_Dmax, (Hpath_ident, Hpath_based)))). - exists g'. - repeat split; simpl in *; auto;try lra. - } - assert (Ht_alive := round_target_alive Hpred Hpath Hbased_higher Hcoll Hexecutioner_off Hexecuted_on Hexists_at H). - unfold target_alive in Ht_alive. - intros g' Halive_r'. - specialize (Ht_alive g'). - destruct (nat_eq_eqdec ( get_ident (round rbg_ila da conf (Good g'))) 0). - left; now simpl in *. - specialize (Ht_alive c Halive_r'). - right. - destruct (Ht_alive) as (?,(?,(?,?))). - destruct (get_based (round rbg_ila da conf (Good g'))) eqn : Hbased_r'. - assert (Hbased_higher_round := based_higher_round Hpred Hpath Hbased_higher Hexecutioner_off Hexecuted_on). - assert (exists g, get_based (conf (Good g)) = true). - exists g'. - apply (still_based_means_based conf g' Hpred Hbased_r'). - specialize (Hbased_higher_round H3). - clear H3. - destruct Hbased_higher_round as (Hap0_r, (Hex_r, Hhig_r)). - destruct Hex_r as (g_ex, (Hb_ex, (Ha_ex, Hd_ex))). - exists g_ex. - repeat split; auto. - assert (Hpos_r : (get_location (round rbg_ila da conf (Good g'))) = (0,0)%R). - specialize (Hap0_r g' Hbased_r'). - intuition. - rewrite Hpos_r in *. - rewrite dist_sym; unfold Dp in *; generalize D_0, Dmax_7D; lra. - destruct (get_based (round rbg_ila da conf (Good x))) eqn : Hbased_rx. - assert (Hbased_higher_round := based_higher_round Hpred Hpath Hbased_higher Hexecutioner_off Hexecuted_on). - assert (exists g, get_based (conf (Good g)) = true). - exists x. - apply (still_based_means_based conf x Hpred Hbased_rx). - specialize (Hbased_higher_round H3). - clear H3. - destruct Hbased_higher_round as (?,(?,?)). - specialize (H5 x g' Hbased_rx Hbased_r'). - lia. - exists x; repeat split; simpl in *; auto. - -Qed. - - -Lemma executed_means_light_on_round : - forall conf da, - da_predicat da -> - path_conf conf -> - based_higher conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - executed_means_light_on (round rbg_ila da conf). -Proof. - intros conf da Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hless_that_Dp. - intros g da' Hpred' Halive_r Halive_rr. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - destruct (round rbg_ila da conf (Good g)) as (pos_r, (((ident_r, light_r), alive_r), based_r)) eqn : Hconf_r. - rewrite round_good_g in *; try apply Hpred; try apply Hpred'. - assert (Hconfr : round rbg_ila da conf (Good g) == (pos_r, (ident_r, light_r, alive_r, based_r))) by now rewrite Hconf_r. - rewrite <- Hconf_r in *. - rewrite round_good_g in *; try apply Hpred; try apply Hpred'. - unfold get_light; unfold get_alive in Halive_rr, Halive_r; simpl; simpl in *. - unfold upt_aux in *. - unfold map_config in *. - destruct based eqn : Hbased. - *** rewrite Hconf in *; simpl in *. - destruct (upt_robot_too_far _ _ ) eqn: Htoo_far; simpl in *. - 2: { - rewrite Hconf_r in Halive_rr. - simpl in Halive_rr. - destruct Hconfr as (Hpos_r, (Hident_r, (Hlight_r, (Halive_r', Hbased_r)))). - rewrite <- Hbased_r in *. - destruct (upt_robot_too_far _ _) in Halive_rr; - simpl in Halive_rr; - rewrite Halive_r, Halive_rr in Halive_r'; discriminate. - } - rewrite Hconf_r in Halive_rr. - simpl in Halive_rr. - destruct Hconfr as (Hpos_r, (Hident_r, (Hlight_r, (Halive_r', Hbased_r)))). - rewrite <- Hbased_r in *. - destruct (upt_robot_dies_b _) eqn : Hbool_r in Halive_rr; try now simpl in *; rewrite Halive_rr, Halive_r in Halive_r'; discriminate. - unfold upt_robot_dies_b, upt_robot_too_far in *. - rewrite orb_true_iff in Hbool_r. - destruct (R2dec_bool _ _ ) eqn : Hpos_too_far in Htoo_far; - try (destruct (forallb _ _) eqn : Htoo_far_1 in Htoo_far); - try (destruct (forallb _ _) eqn : Htoo_far_2 in Htoo_far); - try (simpl; discriminate). - rewrite retraction_section in Hpos_r. - rewrite <- Hpos_r, <- Hident_r, <- Hlight_r, <- Halive_r' in Hconf_r. - rewrite Hconf_r in *. - assert (Hpos_too_far' : R2dec_bool pos (0,0)%R = true). - { assert (Hret_sec := retraction_section (frame_choice_bijection (change_frame da conf g)) (fst (conf (Good g)))); - unfold frame_choice_bijection, Frame in Hret_sec; - rewrite Hconf in Hret_sec; - simpl (_ ∘ _) in Hret_sec; - simpl (fst _) in Hret_sec; - rewrite <- Hret_sec at 1. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite Hconf in *; simpl in *; apply Hpos_too_far. - } - rewrite R2dec_bool_true_iff in Hpos_too_far'. - simpl in Hpos_too_far'. - unfold too_far_aux_1, too_far_aux_2 in *. - destruct Hbool_r as [Hnear|Hfar]. - - rewrite existsb_exists in Hnear. - destruct Hnear as (near, (Hin_near, Hdist_near)). - rewrite filter_In in Hin_near. - destruct Hin_near as (Hin_near, Hprop_near). - rewrite config_list_spec, in_map_iff in Hin_near. - destruct Hin_near as (near_id, (near_comp, Hnear_names)). - rewrite 2 andb_true_iff in Hprop_near. - rewrite <- near_comp in *. - destruct near_id as [g_near|byz]; - try ( - assert (Hfalse := In_Bnames byz); - now simpl in *). - rewrite forallb_forall in Htoo_far_2. - specialize (Htoo_far_2 g_near (In_Gnames g_near)). - rewrite negb_true_iff in Htoo_far_2. - rewrite 2 andb_false_iff in Htoo_far_2. - rewrite Rle_bool_true_iff in *. - rewrite Rle_bool_false_iff in *. - destruct Hprop_near as ((Hident_near, Halive_near), Hbased_near). - destruct Htoo_far_2 as [Hbased_false|[Halive_false |Hdist_false]]. - rewrite negb_false_iff in *. - rewrite negb_true_iff in *. - unfold get_based in *; simpl in *. - rewrite Nat.ltb_lt in *. - assert (ident_near_aux : get_ident (round rbg_ila da conf (Good g_near)) < get_ident (conf (Good g))). - rewrite Hconf; unfold get_ident in *; simpl in *; auto. - - rewrite (ident_preserved conf g Hpred ) in ident_near_aux. - rewrite forallb_forall in Htoo_far_1. - assert (Hin_g_near: List.In g_near - (List.filter - (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g_near). - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g_near). - unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - specialize (Htoo_far_1 g_near Hin_g_near). - destruct (negb _) eqn : Hnegb_1 in Htoo_far_1; try (simpl in Htoo_far_1; discriminate). - rewrite <- 2 ident_preserved in ident_near_aux; try apply Hpred. - rewrite Nat.ltb_lt in Htoo_far_1. - unfold get_ident in *; simpl in *; lia. - rewrite negb_false_iff, Nat.eqb_eq in Hnegb_1. - rewrite <- 2 ident_preserved in ident_near_aux; try auto. - unfold get_ident in *; simpl in *; rewrite Hnegb_1 in *; lia. - assert (get_alive (round rbg_ila da conf (Good g_near)) = true) by - now (unfold get_alive in *; simpl in *; auto). - apply (still_alive_means_alive) in H; try apply Hpred; try apply Hbased_higher. - unfold get_alive in *; simpl in *; rewrite H in *; discriminate. - destruct Hdist_false. - assert (dist (fst (round rbg_ila da conf (Good g_near))) (pos) <= D)%R. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite (frame_dist _ _ (change_frame da' (round rbg_ila da conf) g)). - now simpl in *. - assert ((dist (fst (conf (Good g_near))) (fst (conf (Good g))) <= Dp -3* D)%R -> - (dist - (get_location - ((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near)))) - (get_location - ((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))) <= Dp - 3* D)%R). - intros Hx. - rewrite (frame_dist _ _ (change_frame da conf g)) in Hx. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - simpl (frame_choice_bijection _) in *. - destruct (change_frame _) as ((r,t),b) eqn : Hchange. - destruct b; simpl in *; auto. - destruct (change_frame _) as ((r,t),b) eqn : Hchange. - rewrite Hconf in H0. - simpl (fst _) in H0. - assert (dist (fst (conf (Good g_near))) pos <= 2*D)%R. - assert ( Htri := RealMetricSpace.triang_ineq (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g_near))) pos). - replace (2*D)%R with (D+D)%R by (generalize D_0;lra). - transitivity (dist (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g_near))) + - dist (fst (round rbg_ila da conf (Good g_near))) pos)%R; try auto. - transitivity ((dist (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g_near))) + D)%R). - try lra. - apply Rplus_le_compat_r. - rewrite <- Rle_bool_true_iff. - apply dist_round_max_d. - apply Hpred. - apply Hpath. - assert (get_alive (round rbg_ila da conf (Good g_near)) = true). - unfold get_alive in *; simpl in *; auto. - now apply still_alive_means_alive in H1. - rewrite Hconf. simpl (fst _). - destruct b; apply H0; unfold Dp; generalize Dmax_7D, D_0; lra. - - - assert (Hpath_conf_round := path_round Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hless_that_Dp). - specialize (Hpath_conf_round g). - unfold get_alive in Hpath_conf_round; rewrite Hconf_r in Hpath_conf_round; simpl (snd _) in Hpath_conf_round. - specialize (Hpath_conf_round (reflexivity _)). - destruct Hpath_conf_round as [Hfalse_0 | (gp, (Halivep_r', (Hdistp_r', (Hidentp_r', Hbasedp_r'))))]. - assert (Hfalse_based := ident_0_not_based conf g). - unfold get_based, get_ident in *; rewrite Hconf in *; simpl in *. - now specialize (Hfalse_based Hfalse_0). - rewrite forallb_existsb, forallb_forall in Hfar. - destruct (change_frame da' (round rbg_ila da conf) g) as ((r_r,t_r),b_r). - specialize (Hfar (comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id)) - (fst (round rbg_ila da conf (Good gp))), snd (round rbg_ila da conf (Good gp)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hfar. - destruct Hfar. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff, Nat.ltb_lt, negb_true_iff. - repeat split; try auto. - exists (Good gp). - split; try auto. - destruct b_r; reflexivity. - apply (In_names (Good gp)). - rewrite dist_sym, (frame_dist _ _ ((r_r,t_r,b_r))) in Hdistp_r'. - destruct b_r; simpl in *; auto; apply Hdistp_r'. - - simpl in Halive_rr. - rewrite Halive_rr in *; try discriminate. - *** destruct (upt_robot_dies_b _) eqn : Hbool; rewrite Hconf in *; try now simpl in *. - simpl. - simpl in Halive_r, Halive_rr. - destruct (upt_robot_dies_b - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) - (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila - da' - (@round (prod R2 ILA) Loc State_ILA Identifiers - Obs_ILA (prod R2 NoCollisionAndPath.light) - (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice - inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf) g - return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid - (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec - VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid - R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid - R2_EqDec VS - (@Normed2Metric R2 R2_Setoid - R2_EqDec VS - (@Euclidean2Normed R2 - R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers - Obs_ILA (prod R2 NoCollisionAndPath.light) - (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice - inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf id))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool - bool robot_choice_RL Frame Choice - inactive_choice_ila UpdFun inactive_ila rbg_ila - da conf id))) g) eqn : Hbool_r; try now simpl in *; - try (simpl in *; - unfold rbg_fnc in *; - rewrite Hconf_r in *; - simpl in Halive_rr, Hconf_r; - destruct Hconfr as (_,(_,(_,Halive_false))); - rewrite Halive_r, Halive_rr in Halive_false; discriminate). - rewrite Hconf_r in *. - simpl in *. - - - unfold rbg_fnc. - destruct (move_to _) eqn : Hmove; try now simpl in *. - simpl in Halive_r. - assert (Hbool_backup := Hbool). - unfold upt_robot_dies_b in Hbool. - - rewrite orb_false_iff in *. - - assert (Hbool_safe := Hbool). - - destruct Hbool as (Hexists, Hfar). - rewrite negb_false_iff in Hfar. - rewrite <- negb_true_iff in Hexists. - rewrite forallb_existsb, forallb_forall in Hexists. - rewrite existsb_exists in Hfar. - destruct Hfar as (far, (Hfar_in, Hfar_dist)). - assert (Hmsz := move_to_Some_zone Hmove). - simpl. - exfalso. - assert (Hpath_backup := Hpath). - specialize (Hpath g). - unfold get_alive at 1 in Hpath. - rewrite Hconf in Hpath at 1. - specialize (Hpath Halive_r). - destruct Hpath. - - rewrite (ident_preserved conf g Hpred) in H. - rewrite (ident_preserved (round rbg_ila da conf) g Hpred') in H. - assert (H0 := ident_0_alive (round rbg_ila da' (round rbg_ila da conf)) g H). - rewrite round_good_g in H0. - simpl in H0. - unfold upt_aux in H0; - unfold map_config in *; - rewrite Hbool_r in H0. - unfold get_alive in H0; - simpl in H0. - rewrite Hconf_r in H0; simpl in H0. - assert (Hbased_r : based_r = false). - { - assert (Hbased_r_aux : get_based (round rbg_ila da conf (Good g)) = false). - rewrite round_good_g; try auto. - simpl; unfold upt_aux, map_config, get_based; simpl. - rewrite Hbool_backup. - now rewrite Hconf; simpl. - rewrite Hconf_r in *; unfold get_based in *; simpl in *; auto. - } - rewrite Hbased_r in *. - discriminate. - apply Hpred'. - - destruct H as (g', (Halive_g', (Hdist_d_g', (Hident_g', Hbased_g')))). - destruct (R2_EqDec pos pos_r). - + (* monter que si il bouge aps, comme il s'élimine a (round (round)), c'est pas possible car selon Hforall, il ne peux pas en y avoir a moins de 2D, donc au round suivant il n'y en a aps a moins de D, et selon dist_some, il y en a un a moins de DP, donc il y en a un amoins de Dmax a round d'apres.*) - clear Hexists Hfar_in. - unfold upt_robot_dies_b in Hbool_r. - rewrite orb_true_iff in *. - destruct Hbool_r as [Hexists|Hfar]. - * rewrite existsb_exists in Hexists. - destruct Hexists as (exec, (Hexists_in, Hexists_dist)). - rewrite filter_In in Hexists_in. - destruct Hexists_in as (Hexists_in, Hexists_prop). - rewrite config_list_spec, in_map_iff in Hexists_in. - destruct Hexists_in as (exec_id, (exec_comp, Hexec_names)). - rewrite 2 andb_true_iff in Hexists_prop. - rewrite <- exec_comp in *. - destruct exec_id as [g_exec|byz]. - - - destruct Hexists_prop as ((Hident_exec, Halive_exec), Hbased_exec). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in Hexists_dist. - destruct (change_frame da' _) as ((r_r, t_r), b_r) eqn : Hchange_r. - assert (Rle_bool (dist (fst (round rbg_ila da conf (Good g_exec))) - (fst (round rbg_ila da conf (Good g)))) D = true). - rewrite <- Hexists_dist. - assert (Hframe := frame_dist (fst (round rbg_ila da conf (Good g_exec))) (fst (round rbg_ila da conf (Good g))) ((r_r, t_r), b_r)). - unfold frame_choice_bijection, Frame in Hframe. - rewrite Hframe; simpl in *. reflexivity. - assert (get_alive ((comp (bij_rotation r_r) - (comp (bij_translation t_r) - (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_exec))), - snd (round rbg_ila da conf (Good g_exec))) == true). - rewrite <-Halive_exec. - reflexivity. - assert ((fst (round rbg_ila da conf (Good g_exec))) == - fst ((let frame_choice := change_frame da conf g_exec in - let new_frame := frame_choice_bijection frame_choice in - let local_config := - map_config - (lift - (existT precondition new_frame (precondition_satisfied da conf g))) - conf in - let obs := obs_from_config local_config (local_config (Good g_exec)) in - let local_robot_decision := rbg_ila obs in - let choice := choose_update da local_config g_exec local_robot_decision in - let new_local_state := - update local_config g_exec frame_choice local_robot_decision choice in - lift - (existT precondition (new_frame â»Â¹) - (precondition_satisfied_inv da conf g_exec)) new_local_state))). - rewrite (RelationPairs.fst_compat _ _ _ _ (round_good_g g_exec _ Hpred)). - reflexivity. - assert (Hexec_comp : ((comp (bij_rotation r_r) - (comp (bij_translation t_r) - (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_exec))), - snd (round rbg_ila da conf (Good g_exec))) == exec). - rewrite <- exec_comp. - split; simpl; destruct b_r; auto; - destruct (round rbg_ila da conf (Good g_exec)) as (?,(((?,?),?),?)) eqn : Hconf_r_exec; simpl; auto. - destruct ((round rbg_ila da conf (Good g_exec))) as (pos', (((ide', lid'), ali'), bas')) eqn : Hconf'. - simpl in H1. - rewrite H1 in Hexec_comp. - clear H1. - simpl in Hexec_comp. - assert ((snd (round rbg_ila da conf (Good g_exec))) == - snd ((let frame_choice := change_frame da conf g_exec in - let new_frame := frame_choice_bijection frame_choice in - let local_config := - map_config - (lift - (existT precondition new_frame (precondition_satisfied da conf g))) - conf in - let obs := obs_from_config local_config (local_config (Good g_exec)) in - let local_robot_decision := rbg_ila obs in - let choice := choose_update da local_config g_exec local_robot_decision in - let new_local_state := - update local_config g_exec frame_choice local_robot_decision choice in - lift - (existT precondition (new_frame â»Â¹) - (precondition_satisfied_inv da conf g_exec)) new_local_state))). - rewrite (RelationPairs.snd_compat _ _ _ _ (round_good_g g_exec _ Hpred)). - reflexivity. - rewrite <- exec_comp in Hexec_comp. - rewrite Hconf' in H1. - simpl in H1. - simpl in Hexec_comp. - destruct (change_frame da conf g_exec) as ((r',t'),b') eqn : Hchange'. - unfold upt_aux in Hexec_comp, H1. - unfold map_config in *; - destruct (upt_robot_dies_b _) eqn : Hbool'. - -- destruct (conf (Good g_exec)) as (?,((?,?),?)); - simpl in H1. - unfold get_alive in Halive_exec; simpl in Halive_exec. - now rewrite Halive_exec in *. - -- destruct (conf (Good g_exec)) as (pos_exec,(((ident_exec,light_exec),alive_exec), based_exec)) eqn : Hconf_exec. - unfold rbg_fnc in Hexec_comp. - destruct (move_to _) eqn : Hmove'. - auto. - specialize (Hmsz ((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_exec))), snd (conf (Good g_exec)))). - set (new_pos := choose_new_pos _ _) in *. - - assert ( let (pos_x, _) := - ((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_exec))), snd (conf (Good g_exec))) in - (dist new_pos pos_x > 2 * D)%R). - { apply Hmsz. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff. - destruct (change_frame da conf g) as ((?,?),?). - repeat split. - * - exists (Good g_exec); repeat split; simpl. - destruct b; now simpl. - now destruct (snd (conf (Good g_exec))) as (((?,?),?),?). - * unfold Datatypes.id. - assert ((@dist (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS ES)) - (@fst (@location Loc) ILA - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS r0) - match b return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) (@fst R2 ILA (conf (@Good Identifiers g_exec)))) - (@snd R2 ILA (conf (@Good Identifiers g_exec))))) - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS r0) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) pos)) == dist (fst (conf (Good g_exec))) pos). - rewrite (frame_dist _ pos ((r , r0), b)). - unfold frame_choice_bijection, Frame; destruct b; now simpl. - rewrite H2. - clear H2. - assert (Rle_bool - (dist - (fst (round rbg_ila da conf (Good g_exec))) - (fst (round rbg_ila da conf (Good g)))) D = true). - rewrite <- Hexists_dist. - rewrite (frame_dist (fst (round rbg_ila da conf (Good g_exec))) _ ((r_r,t_r), b_r)). - unfold frame_choice_bijection, Frame; destruct b_r; now simpl in *. - rewrite Rle_bool_true_iff in *. - assert (Rle_bool (dist (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g)))) D = true)%R. - apply dist_round_max_d. - apply Hpred. - apply Hpath_backup. - rewrite Hconf. - now unfold get_alive; simpl. - assert (Rle_bool (dist (get_location (conf (Good g_exec))) (get_location (round rbg_ila da conf (Good g_exec)))) D = true)%R. - apply dist_round_max_d. - apply Hpred. - apply Hpath_backup. - unfold get_alive in *; simpl in Halive_exec. - assert (get_alive (round rbg_ila da conf (Good g_exec)) = true). - rewrite Hconf'. - now unfold get_alive; simpl. - apply still_alive_means_alive in H4. - now unfold get_alive in *. - apply Hpred. - apply Hbased_higher. - generalize Dmax_7D. - rewrite Rle_bool_true_iff in *. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in H3, H4. - rewrite Hconf in H3. - simpl (fst (pos, (ident, light, alive))) in *. - set (u:= (fst (conf (Good g_exec)))) in *; - set (v := (fst (round rbg_ila da conf (Good g_exec)))) in *; - set (w := (fst (round rbg_ila da conf (Good g)))) in *. - assert (dist pos w <= D)%R by now unfold w; apply H3. - assert (dist u v <= D)%R by now unfold v. - clear H3 H4. - assert (dist u w <= 2*D)%R. - assert (Htri := RealMetricSpace.triang_ineq u v w). - lra. - assert (Htri := RealMetricSpace.triang_ineq u w pos). - generalize D_0. - rewrite dist_sym in H5; lra. - * - apply (still_alive_means_alive g_exec Hpred Hbased_higher). - unfold get_alive in *; rewrite Hconf'; simpl (snd _) in *. - auto. - - * unfold get_alive; simpl in *; auto. - * assert (Hident_exec' : (get_ident - (comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id)) - (fst (pos', (ide', lid', ali', bas'))), snd (pos', (ide', lid', ali',bas')))) <? - get_ident (round rbg_ila da conf (Good g)) = true). - now unfold get_ident in *; simpl in *. - rewrite <- ident_preserved in Hident_exec'; try auto. - auto. - unfold ILA in *. - assert (Haux : get_ident (round rbg_ila da conf (Good g_exec)) = ide'). - unfold ILA in *; - rewrite Hconf'; unfold get_ident; simpl; auto. - unfold get_ident in Hident_exec'. simpl in Hident_exec'. - rewrite <- Haux in Hident_exec'. - rewrite <- ident_preserved, Hconf in Hident_exec'; try auto. - unfold get_ident in *; simpl in *. - rewrite Nat.leb_le, Nat.ltb_lt in *. - unfold ILA in *; lia. - - * intros x y Hxy. - rewrite (get_alive_compat Hxy). - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_ident_compat Hxy). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - reflexivity. - * rewrite <- Hconf' in *. - assert (Hident_p := ident_preserved conf ). - unfold get_ident in Hident_exec, Hident_p. - simpl in Hident_p, Hident_exec. - rewrite Nat.ltb_lt in Hident_exec. - rewrite <- 2 Hident_p in Hident_exec; try apply Hpred. - destruct (Geq_dec g g_exec). rewrite e0 in Hident_exec. - lia. - specialize (Hcol _ _ n0). - assert (forall (x y : R2), dist x y <> 0%R -> dist x y > 0%R)%R. - intros. - generalize (dist_nonneg x y). lra. - apply H2. - assert (dist (get_location (conf (Good g))) (get_location (conf (Good g_exec))) <> 0)%R. - apply Hcol. - unfold get_based; rewrite Hconf; now simpl. - rewrite negb_true_iff in Hbased_exec. - destruct (get_based (conf (Good g_exec))) eqn : Hbased_exec_before; try auto. - destruct Hbased_higher as (_,(_,Hbased_ident)). - specialize (Hbased_ident g_exec g Hbased_exec_before). - rewrite Hconf in *. - specialize (Hbased_ident (reflexivity _)). - unfold get_ident in *; simpl in *. lia. - rewrite Hconf; unfold get_alive; simpl; auto. - apply (still_alive_means_alive g_exec Hpred); try auto. - destruct Hpred as (Horigin,_). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - specialize (Horigin conf g (r,t,b) Hchange). - simpl in Horigin. - unfold Datatypes.id in Horigin. - rewrite <- Horigin. - rewrite dist_sym in H3. - rewrite (frame_dist _ _ (r,t,b)) in H3. - destruct b; simpl in *; apply H3. - } - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - clear H1. - assert (new_pos == ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g))))). - unfold round. - destruct (Hpred) as (?,?). - specialize (H3 (Good g)). - rewrite H3. - simpl ((fst - (lift - _ _))). - rewrite Hchange. - unfold upt_aux. - unfold map_config in *; - clear Hbool'. - destruct (upt_robot_dies_b _ g) eqn : Hfalse. - rewrite <- orb_false_iff in Hbool_safe. - simpl in Hfalse. - unfold upt_robot_dies_b in Hfalse. - revert Hbool_safe; intro Hbool_safe. - simpl in *. - assert (false = true). - - rewrite <- Hfalse, <- Hbool_safe. - destruct b; simpl; reflexivity. - discriminate. - clear Hfalse. - unfold rbg_fnc. rewrite Hconf. simpl (move_to _ _); simpl in Hmove'. unfold new_pos in Hmove'. simpl in *; rewrite Hmove'. - rewrite <- (section_retraction (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos). - now destruct b; simpl in *. - rewrite H1 in H2. - assert (Hdist_round_g_to_exec : (dist - (fst (round rbg_ila da conf (Good g))) - (fst (conf (Good g_exec))) > 2 * D)%R). - rewrite (frame_dist _ _ ((r,t),b)). - unfold frame_choice_bijection, Frame. - destruct b; now simpl in *. - clear H1 H2. - assert ( Rle_bool - (dist (fst (round rbg_ila da conf (Good g_exec))) - (fst (round rbg_ila da conf (Good g)))) D = true)%R. - rewrite (frame_dist _ _ ((r_r,t_r),b_r)). - unfold frame_choice_bijection, Frame. - destruct (b_r); now simpl in *. - rewrite Rle_bool_true_iff in H1. - assert (Rle_bool (dist (get_location (conf (Good g_exec))) (get_location (round rbg_ila da conf (Good g_exec)))) D = true). - apply dist_round_max_d. - apply Hpred. - apply Hpath_backup. - rewrite <- Hconf' in Halive_exec. - unfold get_alive in Halive_exec. - assert (get_alive (round rbg_ila da conf (Good g_exec)) = true). - rewrite Hconf'. - now unfold get_alive; simpl. - apply still_alive_means_alive in H2. - unfold get_alive in *; now simpl in *. - apply Hpred. - apply Hbased_higher. - rewrite Rle_bool_true_iff in H2. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g))) (get_location (round rbg_ila da conf (Good g_exec))) (fst (conf (Good g_exec)))). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite dist_sym in H1, H2. - lra. - - discriminate. - -- - assert (Hfalse := In_Bnames byz). - now simpl in *. - * rewrite forallb_existsb, forallb_forall in Hfar. - set (target := choose_target _) in *. - assert (Hin_obs := @choose_target_in_obs (obs_from_config _ _) - (target) (reflexivity _)). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange; - destruct (change_frame da' (round rbg_ila da conf) g) as ((r_r, t_r),b_r) eqn : Hchange_r. - - unfold obs_from_config, Obs_ILA in Hin_obs. - rewrite make_set_spec, filter_InA in Hin_obs. - destruct Hin_obs. - rewrite config_list_InA in H. - destruct H as (id, H). - destruct id as [g_far|byz]. - rewrite andb_true_iff in H0; destruct H0. - assert (Hfar':=Hfar). - specialize (Hfar ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_far))), snd (round rbg_ila da conf (Good g_far)))). - - assert (@List.In (prod R2 ILA) - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r_r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t_r) - match b_r return (@bijection R2 R2_Setoid) with - | true => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) reflexion - | false => - @sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - (@Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - end)) - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf (@Good Identifiers g_far)))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila rbg_ila - da conf (@Good Identifiers g_far)))) - (@List.filter (prod R2 ILA) - (fun elt : prod R2 ILA => - andb - (andb - (Nat.ltb (get_ident elt) - (get_ident - (@pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r_r) - (@comp R2 R2_Setoid - (@bij_translation R2 R2_Setoid R2_EqDec VS t_r) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b_r - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) - (prod (prod R R2) bool) bool bool robot_choice_RL Frame - Choice inactive_choice_ila UpdFun inactive_ila rbg_ila da - conf (@Good Identifiers g)))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) - (prod (prod R R2) bool) bool bool robot_choice_RL Frame Choice - inactive_choice_ila UpdFun inactive_ila rbg_ila da conf - (@Good Identifiers g)))))) (get_alive elt)) - (negb (get_based elt))) - (@config_list Loc (prod R2 ILA) State_ILA Identifiers - (fun id : @Identifiers.ident Identifiers => - @pair R2 ILA - (@section R2 R2_Setoid - (@comp R2 R2_Setoid (bij_rotation r_r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t_r) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b_r - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end))) - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf id))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf id)))))). - { rewrite filter_In, config_list_spec, in_map_iff. - repeat split. - - exists (Good g_far). - split; try (now destruct b_r; simpl). - apply In_names. - - rewrite 2 andb_true_iff. - repeat split. - assert (Hident := ident_preserved). - unfold get_ident in *. - simpl (snd _). - rewrite <- 2 Hident; try apply Hpred. - rewrite Nat.ltb_lt. - assert (Hlower_target := target_lower). - changeR2. - fold obs_from_config in *. - set (obs := obs_from_config _ _) in *. - specialize (@Hlower_target obs - target - ((let (p, b) := change_frame da conf g in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))) - - ). - unfold get_ident in Hlower_target. - simpl (snd _) in Hlower_target. - unfold obs_from_config, Obs_ILA in Hlower_target. - destruct (change_frame _) as ((rr,tt),bb) eqn : Hchangee. - assert (((r,t),b) == ((rr,tt),bb)). - rewrite <- Hchange, <- Hchangee. - reflexivity. - do 2 destruct H2. - simpl in H2, H3, H4. - rewrite <- H2, <- H4, <- H3 in Hlower_target. - assert ((fst (fst (fst (snd ( - ((comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_far))), snd (conf (Good g_far)))))))) <? - (fst (fst (fst (snd (conf (Good g)))))) == true). - rewrite <- Nat.ltb_lt in Hlower_target. - rewrite <- Hlower_target. - f_equiv. - destruct H as (?,Hsnd). - unfold location, Loc, make_Location in Hsnd. - destruct (target) as (p_target, (((i_target,l_target), a_target), b_target)) eqn : Hconf_target. - destruct ((conf (Good g_far))) as (?,(((?,?),?),?)) eqn : Hconf_bij. - simpl in Hsnd. - destruct Hsnd as (?,(?,?)). - - rewrite H. now simpl. - unfold obs, obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g). - split; destruct b; simpl; - auto. - now rewrite Hconf; simpl. - now rewrite Hconf; simpl. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - unfold Datatypes.id. - rewrite Hconf in *. - split. - assert (Hframe :=(frame_dist (fst (pos, (ident, light, alive))) pos ((r,t),b))). - unfold frame_choice_bijection, Frame in Hframe. - simpl in *. - destruct b; simpl in *; - rewrite <- Hframe; - repeat rewrite Rplus_opp_r; - replace (0*0+0*0)%R with 0%R by lra; - generalize Dmax_7D D_0; - rewrite sqrt_0; try ( - intros; split; try (lra); - unfold get_alive; simpl; auto; - intros; split; try (lra); - unfold get_alive; simpl; - auto). - unfold get_ident; simpl; rewrite Nat.leb_le; lia. - intros x y Hxy. - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy), (get_location_compat _ _ Hxy). - f_equiv. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - rewrite Hchangee. - rewrite H2, H4, H3. - destruct bb; simpl; auto. - reflexivity. - rewrite <- Nat.ltb_lt in *. - now simpl in *. - unfold get_alive; simpl. - unfold round. - assert(Hpred_backup := Hpred). - destruct Hpred as (?,?). - specialize (H3 (Good g_far)). - rewrite H3. - simpl ((lift _ _)). - unfold upt_aux. - unfold map_config in *. - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - destruct (upt_robot_dies_b _ g_far) eqn : Hbool_far; try now simpl. - simpl. - specialize (Hexecuted (g_far)). - assert (Halive_far : get_alive (round rbg_ila da conf (Good g_far)) == false). - rewrite round_good_g. - simpl ; unfold upt_aux; unfold map_config. - rewrite Hbool_far. - unfold get_alive; simpl. - assert (get_based (conf (Good g_far)) = false). - { - rewrite Hconf_far. - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. lia. - } - unfold get_based in H4. - - rewrite Hconf_far in *; simpl in *; rewrite H4; - now simpl. - apply Hpred_backup. - assert (get_light (conf (Good g_far)) == true). - { - apply (Hexecuted da). - apply Hpred_backup. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - destruct b; rewrite <- H4, <- H; rewrite 2 andb_true_iff in H0; destruct H0; - try (now rewrite H6); - now simpl in *. - auto. - } - assert (Htarget_light : get_light (target) = true). - unfold target, obs_from_config, Obs_ILA. - unfold target in *. - rewrite H, <- Hconf_far. - now unfold get_light in *; simpl in *. - assert (Hlight_off_first := light_off_first (reflexivity _) Htarget_light). - unfold For_all in *. - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - unfold obs_from_config, Obs_ILA in Hdist_dp. - assert ((dist pos p_far) <= Dp)%R. - { destruct Hconfr as (Hpos_rl,_). - set (obs := obs_from_config _ _) in *. - assert (Hpos_rl' : @equiv R2 R2_Setoid (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst - (rbg_fnc - obs))) pos_r) by (rewrite <- Hpos_rl; destruct b; now simpl). - rewrite <- Inversion in Hpos_rl'. - destruct b; - rewrite (RelationPairs.fst_compat _ _ _ _ H) in Hdist_dp; - unfold rbg_fnc in Hpos_rl'; unfold new_pos, obs, target, obs in *; rewrite Hmove in Hpos_rl'; - simpl (fst _) in Hpos_rl'; - unfold equiv, R2_Setoid in Hpos_rl'; - simpl in *; rewrite <- Hpos_rl' in Hdist_dp; - rewrite (@frame_dist pos p_far (@change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g)); - unfold frame_choice_bijection, Frame; fold Frame; try rewrite Hchange; - simpl in *; rewrite e in *. lra. lra. - } - unfold upt_robot_dies_b in Hbool_far. - rewrite orb_true_iff in Hbool_far. - destruct Hbool_far as [Htoo_close_so_lighted|Htoo_far_but_path_conf]. - rewrite existsb_exists in *. - destruct Htoo_close_so_lighted as (too_close, (Hin_too_close, Hdist_too_close)). - ++ unfold executioner_means_light_off in *. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff in Hin_too_close. - destruct Hin_too_close as (([g_too_close | byz], (Hspec_too_close, Hnames_too_close)), ((Hident_too_close, Halive_too_close), Hbased_too_far)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - specialize (Hexecutioner g_too_close da Hpred_backup). - rewrite <- Hspec_too_close in Halive_too_close. - unfold get_alive in *. - simpl (snd ) in *. - - assert (get_based (conf (Good g_far)) = false). - { - rewrite Hconf_far. - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. lia. - } - unfold get_based in *; rewrite Hconf_far in H6; simpl in H6; - rewrite H6 in *. - specialize (Hexecutioner Halive_too_close). - assert (Hex : (exists g' : G, - snd (fst (snd (conf (Good g')))) = true /\ - snd (snd (conf (Good g'))) = false /\ - get_ident (conf (Good g_too_close)) < get_ident (conf (Good g')) /\ - (dist (get_location (conf (Good g_too_close))) - (get_location (conf (Good g'))) <= D)%R)). - { exists g_far. - repeat split; try now simpl. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - - revert H H0; intros H H0. - rewrite 2 andb_true_iff in H0; destruct H0 as ((H0, H0_ident), H0'). - unfold get_alive in *. rewrite <- H7. - rewrite <- (get_alive_compat H). - now unfold get_alive. - rewrite Hconf_far; simpl; auto. - rewrite Nat.ltb_lt in Hident_too_close. - rewrite <- Hspec_too_close in Hident_too_close. - unfold get_ident in *; simpl in *; auto. - rewrite <- Hspec_too_close, Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite (frame_dist _ _ (change_frame da conf g_far)). - unfold frame_choice_bijection, Frame; now simpl in *. - } - specialize (Hexecutioner Hex). - clear Hex. - specialize (Hlight_off_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))). - unfold equiv, bool_Setoid, eq_setoid in Hlight_off_first. - rewrite <- Hlight_off_first, <- Hexecutioner. - unfold get_light; now simpl in *. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. repeat split. - exists (Good g_too_close). - destruct b; reflexivity. - rewrite 2 andb_true_iff; repeat split. - rewrite Rle_bool_true_iff. - replace (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close))))%R - with - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))))%R. - unfold Datatypes.id. - assert (Hframe := frame_dist (fst (conf (Good g_too_close))) pos ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (dist (fst (conf (Good g_too_close))) pos <= Dmax)%R. - rewrite Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite <- Hspec_too_close in Hdist_too_close. - assert ((dist (fst (conf (Good g_too_close))) - (fst (conf (Good g_far))) <= D)%R). - rewrite (frame_dist _ _ (change_frame da conf g_far)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - now simpl; simpl in Hdist_too_close. - assert (Hti := RealMetricSpace.triang_ineq (fst (conf (Good g_too_close))) (fst (conf (Good g_far))) pos ). - rewrite Hconf_far in Hti at 2. - simpl ( (fst _)) in Hti. - rewrite dist_sym in H5. - generalize (D_0), Dmax_7D. - unfold Dp in *. - lra. - now destruct b; simpl in *; rewrite <- Hframe. - now destruct b; simpl in *. - unfold get_alive; now simpl. - unfold get_alive; now simpl. - assert (Htarget_lower := target_lower). - set (obs := obs_from_config _ _) in *. - specialize (Htarget_lower obs ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far))) ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - rewrite Hconf in Htarget_lower. - assert (get_ident (conf (Good g_far)) < get_ident (conf (Good g))). - unfold get_ident in *; simpl (fst _) in *; rewrite Hconf, Hconf_far; apply Htarget_lower. - unfold obs, obs_from_config, Obs_ILA. - - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff. - repeat split. - exists (Good g). - now destruct b; rewrite Hconf. - unfold Datatypes.id. - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - rewrite Rle_bool_true_iff. - intro Hdist_0. - simpl (snd _). - destruct b; rewrite dist_same; lra. - simpl in *; assumption. - simpl in *; assumption. - unfold get_ident in *; rewrite Nat.leb_le; simpl; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - assert (Hcompat' := get_ident_compat Hxy); unfold get_ident in Hcompat'; rewrite Hcompat'. - reflexivity. - destruct Hpred_backup as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - rewrite Hchange. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hconf; simpl. destruct b; simpl; reflexivity. - unfold target, obs_from_config, Obs_ILA in *. - unfold get_alive in *. - unfold location, Loc, make_Location. - rewrite H. - rewrite H6; destruct b; reflexivity. - rewrite Nat.leb_le. - transitivity (get_ident (conf (Good g_far))). - rewrite <- Hspec_too_close in Hident_too_close; - unfold get_ident in *; simpl in Hident_too_close; - rewrite Nat.ltb_lt in Hident_too_close; simpl. - lia. - unfold get_ident in *; simpl in *; rewrite Hconf in *; simpl in *; - lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - ++ - specialize (Hpath_backup g_far). - assert (get_alive (conf (Good g_far)) == true). - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - revert H H0; intros H H0. - rewrite 2 andb_true_iff in H0; destruct H0 as ((H0, H0_ident), H0'). - rewrite <- H6. - rewrite <- (get_alive_compat H). - now unfold get_alive. - simpl in H6. - specialize (Hpath_backup H6); clear H6. - destruct Hpath_backup as [Hident_0|Hpath_backup]. - rewrite (ident_preserved _ _ Hpred_backup) in Hident_0. - assert (get_alive (round rbg_ila da conf (Good g_far)) = true). - apply ident_0_alive; intuition. - rewrite H6 in *; discriminate. - rewrite forallb_existsb, forallb_forall in Htoo_far_but_path_conf. - destruct Hpath_backup as (g_too_close, (Halive_close,( Hdist_close, (Hident_close, Hbased_close)))). - specialize (Htoo_far_but_path_conf - ((((let (p, b) := change_frame da conf g_far in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))))). - rewrite negb_true_iff, Rle_bool_false_iff in Htoo_far_but_path_conf. - destruct Htoo_far_but_path_conf. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - ** exists (Good g_too_close). - split. - destruct (change_frame da conf g_far) as (?,b0); destruct b0; - now simpl. - apply In_names. - ** rewrite Nat.ltb_lt. - unfold get_ident in *; simpl in Hident_close; simpl. - auto. - ** rewrite <- Halive_close. - now unfold get_alive; simpl. - ** rewrite negb_true_iff; unfold get_based in *; simpl in *; auto. - ** rewrite dist_sym, (frame_dist _ _ (change_frame da conf g_far)) in Hdist_close. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - unfold frame_choice_bijection, Frame in Hdist_close; fold Frame in *. - revert Hdist_close; intros Hdist_close. - destruct (change_frame _ _ g_far) as ((r_f, t_f), b_f) - eqn : Hchange_far. - now destruct b_f; simpl in *. - ++ simpl. - - rewrite 2 andb_true_iff in H0. - destruct H0 as ((H0,H0_aux), H0'). - - assert (get_based (conf (Good g_far)) = false). - { - rewrite Hconf_far. - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. lia. - } - unfold get_based in *; rewrite Hconf_far in H4; simpl in H4; - rewrite H4 in *. - simpl. - rewrite (get_alive_compat H) in H0_aux. - unfold get_alive in *; simpl in *;auto. - ++ - assert (get_based (conf (Good g_far)) = false). - { - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. lia. - } - rewrite negb_true_iff. - unfold get_based; simpl. - unfold round. - destruct (Hpred); try auto. - rewrite (H4 (Good g_far)). - simpl. - unfold upt_aux, map_config. - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - unfold get_based in *; simpl in H2. - rewrite H2. - simpl. - destruct (upt_robot_dies_b _ g_far); simpl; try auto. - } - specialize (Hfar H2). - rewrite negb_true_iff, Rle_bool_false_iff in Hfar. - clear H2. - destruct Hfar. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - rewrite 2 andb_true_iff in H0. - destruct H0 as ((H0, H0_aux), H0'). - rewrite Rle_bool_true_iff in H0. - rewrite (RelationPairs.fst_compat _ _ _ _ H) in H0. - revert H0; intros H0. - assert (Hdist_round : (dist (fst (conf (Good g_far))) (fst (round rbg_ila da conf (Good g_far))) <= D)%R). - rewrite <- Rle_bool_true_iff. - assert (Hd := dist_round_max_d g_far Hpred Hpath_backup). - destruct (conf (Good g_far)) as (p_far,(((i_far, l_far), a_far), b_far)) eqn : Hconf_far. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - rewrite <- (get_alive_compat H) in H2. - rewrite Hconf_far in H2. - rewrite H0_aux in H2. - now specialize (Hd (symmetry H2)). - assert (Hconf_r_equiv : round rbg_ila da conf (Good g) == (pos_r, (ident_r, light_r, alive_r, based_r))) by now rewrite Hconf_r. - rewrite (RelationPairs.fst_compat _ _ _ _ Hconf_r_equiv). - rewrite <- e. - simpl ((fst (pos, (ident_r, light_r, alive_r, based_r)))). - assert (dist ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_far)))) - ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - pos) <= Dmax)%R. - assert (dist (fst (round rbg_ila da conf (Good g_far))) pos <= Dmax)%R. - destruct (conf (Good g_far)) as (p_far, (((i_far, l_far), a_far), b_far)) eqn : Hconf_far. - assert ((dist pos p_far) <= Dp)%R. - { destruct Hconfr as (Hpos_rl,_). - set (obs := obs_from_config _ _) in *. - assert (Hpos_rl' : @equiv R2 R2_Setoid (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst - (rbg_fnc - obs))) - pos_r) by (rewrite <- Hpos_rl; destruct b; now simpl). - rewrite <- Inversion in Hpos_rl'. - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - - destruct b; - rewrite (RelationPairs.fst_compat _ _ _ _ H) in Hdist_dp; - unfold rbg_fnc in Hpos_rl'; unfold new_pos, obs, target, obs in *; rewrite Hmove in Hpos_rl'; - simpl (fst _) in Hpos_rl'; - unfold equiv, R2_Setoid in Hpos_rl'; - simpl in *; rewrite <- Hpos_rl' in Hdist_dp; - rewrite (@frame_dist pos p_far (@change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g)); - unfold frame_choice_bijection, Frame; fold Frame; try rewrite Hchange; - simpl in *; rewrite e in *; lra. - } - simpl (fst _) at 1 in Hdist_round. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_far))) p_far pos). - rewrite dist_sym in Hdist_round, H2. - generalize D_0, Dmax_7D. - unfold Dp in *. lra. - rewrite (frame_dist _ _ ((r_r, t_r), b_r)) in H2. - unfold frame_choice_bijection, Frame in H2; fold Frame in H2. - destruct b_r; now simpl in *. - destruct b_r; now simpl in *. - assert (Hfalse := In_Bnames byz). - now simpl in *. - - - - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), - (get_alive_compat Hxy), (get_ident_compat Hxy); - reflexivity. - + unfold rbg_fnc in Hconfr. - rewrite Hmove in Hconfr. - unfold upt_robot_dies_b in Hbool_r. - rewrite (orb_true_iff) in Hbool_r. - destruct Hbool_r as [Hnear|Hfar]. - * rewrite existsb_exists in Hnear. - destruct Hnear as (near, (Hin_near, Hdist_near)). - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff in Hin_near. - destruct Hin_near as (([g_near|byz],(Hspec_near, Hnames_near)), ((Hident_near, Halive_near), Hbased_near)); try (assert (Hfalse := In_Bnames byz); - simpl in *; auto). - rewrite <- Hspec_near in *. - destruct Hconfr as (Hpos_rl, _). - destruct (change_frame _ _ g) as ((r,t),b) eqn : Hchange. - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - assert (Hpos_rl_equiv : retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (new_pos, false)) == pos_r). - rewrite <- Hpos_rl. - destruct b; reflexivity. - rewrite <- Inversion in Hpos_rl_equiv. - simpl (fst _) in Hpos_rl_equiv. - rewrite <- Hpos_rl_equiv in *. - specialize (Hmsz ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g_near))), snd (conf (Good g_near)))). - apply Rgt_not_le in Hmsz. - assert (dist pos_r (fst (conf (Good g_near))) <= 2*D)%R. - revert Hdist_near; intros Hdist_near. - rewrite Hconf_r in Hdist_near. - assert (dist (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g_near))) <= D)%R. - rewrite <- Rle_bool_true_iff. - apply dist_round_max_d. - apply Hpred. - apply Hpath_backup. - assert (get_alive ((round rbg_ila da conf (Good g_near))) == true). - rewrite <- Halive_near; unfold get_alive; now simpl. - apply still_alive_means_alive in H; try apply Hpred. - auto. - auto. - rewrite Rle_bool_true_iff in Hdist_near. - - assert ((@dist (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS ES)) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da' - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf) g return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf (@Good Identifiers g_near)))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf (@Good Identifiers g_near))))) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da' - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf) g return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@fst R2 ILA - (@pair R2 ILA pos_r - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident_r light_r) - alive_r) based_r)))) - (@snd R2 ILA - (@pair R2 ILA pos_r - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident_r light_r) - alive_r) based_r)))))) == - dist (fst (round rbg_ila da conf (Good g_near))) - (fst (pos_r, (ident_r, light_r, alive_r, based_r)))). - rewrite (frame_dist (fst (round rbg_ila da conf (Good g_near))) _ (change_frame da' (round rbg_ila da conf) g)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - - unfold frame_choice_bijection, Frame. fold Frame. - destruct (change_frame _ (round _ _ _) g) as ((?,?), xb) eqn : Hchange'; destruct xb; - simpl in *; auto. - assert (dist (fst (round rbg_ila da conf (Good g_near))) - (fst (pos_r, (ident_r, light_r, alive_r))) <= D)%R. - rewrite <- H0. apply Hdist_near. - simpl (fst _) in H1. - rewrite dist_sym in H1. - assert (Htri := RealMetricSpace.triang_ineq pos_r (fst (round rbg_ila da conf (Good g_near))) (fst (conf (Good g_near)))). - generalize D_0. rewrite dist_sym in H. lra. - - destruct Hmsz. - rewrite (frame_dist _ _ ((r,t),b)) in H. - unfold frame_choice_bijection, Frame in H. - destruct b; simpl in *; lra. - assert (dist pos_r (fst (conf (Good g_near))) <= 2*D)%R. - revert Hdist_near; intros Hdist_near. - rewrite Hconf_r in Hdist_near. - assert (dist (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g_near))) <= D)%R. - rewrite <- Rle_bool_true_iff. - apply dist_round_max_d. - apply Hpred. - apply Hpath_backup. - assert (get_alive ((round rbg_ila da conf (Good g_near))) == true). - rewrite <- Halive_near; unfold get_alive; now simpl. - apply still_alive_means_alive in H; try apply Hpred. - auto. - auto. - rewrite Rle_bool_true_iff in Hdist_near. - assert ( (@dist (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Normed2Metric (@location Loc) (@location_Setoid Loc) (@location_EqDec Loc) VS - (@Euclidean2Normed (@location Loc) (@location_Setoid Loc) - (@location_EqDec Loc) VS ES)) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da' - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf) g return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@fst R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf (@Good Identifiers g_near)))) - (@snd R2 ILA - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf (@Good Identifiers g_near))))) - (@get_location Loc (prod R2 ILA) State_ILA - (@pair R2 ILA - (@section R2 R2_Setoid - match - @change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da' - (@round (prod R2 ILA) Loc State_ILA Identifiers Obs_ILA - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila UpdFun inactive_ila - rbg_ila da conf) g return (@bijection R2 R2_Setoid) - with - | pair p b => - match p return (@bijection R2 R2_Setoid) with - | pair r t => - @comp R2 R2_Setoid (bij_rotation r) - (@comp R2 R2_Setoid (@bij_translation R2 R2_Setoid R2_EqDec VS t) - (@sim_f R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - match - b - return - (@similarity R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES))) - with - | true => reflexion - | false => - @Similarity.id R2 R2_Setoid R2_EqDec VS - (@Normed2Metric R2 R2_Setoid R2_EqDec VS - (@Euclidean2Normed R2 R2_Setoid R2_EqDec VS ES)) - end)) - end - end - (@fst R2 ILA - (@pair R2 ILA pos_r - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident_r light_r) - alive_r) based_r)))) - (@snd R2 ILA - (@pair R2 ILA pos_r - (@pair - (prod (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive) NoCollisionAndPath.based - (@pair (prod identifiants NoCollisionAndPath.light) - NoCollisionAndPath.alive - (@pair identifiants NoCollisionAndPath.light ident_r light_r) - alive_r) based_r)))))) == - dist (fst (round rbg_ila da conf (Good g_near))) - (fst (pos_r, (ident_r, light_r, alive_r)))). - rewrite (frame_dist (fst (round rbg_ila da conf (Good g_near))) _ (change_frame da' (round rbg_ila da conf) g)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - - unfold frame_choice_bijection, Frame. fold Frame. - destruct (change_frame _ (round _ _ _) g) as ((?,?), xb) eqn : Hchange'; destruct xb; - simpl in *; auto. - assert (dist (fst (round rbg_ila da conf (Good g_near))) - (fst (pos_r, (ident_r, light_r, alive_r))) <= D)%R. - rewrite <- H0. apply Hdist_near. - simpl (fst _) in H1. - rewrite dist_sym in H1. - assert (Htri := RealMetricSpace.triang_ineq pos_r (fst (round rbg_ila da conf (Good g_near))) (fst (conf (Good g_near)))). - generalize D_0. rewrite dist_sym in H. lra. - - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. - repeat split. - exists (Good g_near). - destruct b; reflexivity. - rewrite 2 andb_true_iff. - repeat split. - rewrite Rle_bool_true_iff. - assert (Hdist_round := dist_round_max_d g Hpred Hpath_backup). - rewrite Hconf in Hdist_round; unfold get_alive in Hdist_round. - simpl (snd _ ) in Hdist_round. - rewrite Halive_r in *. - specialize (Hdist_round (reflexivity _)). - simpl (equiv) in Hdist_round. - rewrite Rle_bool_true_iff in Hdist_round. - assert (dist (fst (conf (Good g_near))) (pos)<= Dmax)%R. - assert (Htri := RealMetricSpace.triang_ineq (fst (conf (Good g_near))) (fst (round rbg_ila da conf (Good g))) pos). - rewrite Hconf_r in *. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - simpl (fst _) in *. - rewrite dist_sym in H, Hdist_round. - generalize D_0, Dmax_7D. - lra. - rewrite (frame_dist _ _ ((r,t),b)) in H0. - now destruct b; simpl in *. - assert (get_alive (round rbg_ila da conf (Good g_near)) = true). - rewrite <- Halive_near. - now unfold get_alive; simpl in *. - apply still_alive_means_alive in H0; try apply Hpred. - unfold get_alive in *; now simpl in *. - auto. - unfold get_alive in *; simpl in *; auto. - rewrite Nat.leb_le. unfold get_ident; simpl. - rewrite Nat.ltb_lt in Hident_near. - revert Hident_near; intro Hident_near. - assert (Hident_near' : get_ident ( (round rbg_ila da conf (Good g_near))) < - get_ident ( (round rbg_ila da conf (Good g))))by - (unfold get_ident in *; simpl in *; lia). - rewrite <- 2 ident_preserved in Hident_near'; try apply Hpred. - rewrite Hconf in Hident_near'; unfold get_ident in *; simpl in *; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - assert (Hident_p := ident_preserved conf ). - unfold get_ident in Hident_near, Hident_p. - simpl in Hident_p, Hident_near. - rewrite Nat.ltb_lt in Hident_near. - rewrite <- 2 Hident_p in Hident_near; try apply Hpred. - destruct (Geq_dec g g_near). rewrite e in Hident_near. - lia. - specialize (Hcol _ _ n0). - assert (forall (x y : R2), dist x y <> 0%R -> dist x y > 0%R)%R. - intros. - generalize (dist_nonneg x y). lra. - apply H. - assert (dist (get_location (conf (Good g))) (get_location (conf (Good g_near))) <> 0)%R. - apply Hcol. - unfold get_based; rewrite Hconf; now simpl. - rewrite negb_true_iff in Hbased_near. - destruct (get_based (conf (Good g_near))) eqn : Hbased_near_before; try auto. - destruct Hbased_higher as (_,(_,Hbased_ident)). - specialize (Hbased_ident g_near g Hbased_near_before). - rewrite Hconf in *. - specialize (Hbased_ident (reflexivity _)). - unfold get_ident in *; simpl in *. lia. - rewrite Hconf; unfold get_alive; simpl; auto. - apply (still_alive_means_alive g_near Hpred); try auto. - destruct Hpred as (Horigin,_). - specialize (Horigin conf g (r,t,b) Hchange). - simpl in Horigin. - unfold Datatypes.id in Horigin. - rewrite <- Horigin. - rewrite dist_sym in H0. - rewrite (frame_dist _ _ (r,t,b)) in H0. - destruct b; simpl in *; apply H0. - * set (obs := obs_from_config _ _) in *. - assert (Hin_obs := @choose_target_in_obs obs - (choose_target obs) (reflexivity _)). - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange; - destruct (change_frame da' (round rbg_ila da conf) g) as ((r_r, t_r),b_r) eqn : Hchange_r. - - unfold obs, obs_from_config, Obs_ILA in Hin_obs. - rewrite make_set_spec, filter_InA in Hin_obs. - destruct Hin_obs. - rewrite config_list_InA in H. - destruct H as (id, H). - destruct id as [g_far|byz]; - try now ( assert (Hfalse := In_Bnames byz); - now simpl in *). - - rewrite andb_true_iff in H0; destruct H0. - assert (Hfar':=Hfar). - rewrite forallb_existsb, forallb_forall in Hfar. - - specialize (Hfar ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_far))), snd (round rbg_ila da conf (Good g_far)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hfar. - destruct Hfar. - { rewrite filter_In, config_list_spec, in_map_iff. - repeat split. - - exists (Good g_far). - split; try (now destruct b_r; simpl). - apply In_names. - - rewrite 2 andb_true_iff. - repeat split. - assert (Hident := ident_preserved). - unfold get_ident in *. - simpl (snd _). - rewrite <- 2 Hident; try apply Hpred. - rewrite Nat.ltb_lt. - assert (Hlower_target := target_lower). - changeR2. - fold obs_from_config in *. - specialize (@Hlower_target obs (choose_target obs) ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - unfold obs, obs_from_config, Obs_ILA in *. - rewrite H in Hlower_target. - unfold get_ident in *. - unfold obs_from_config, Obs_ILA in Hlower_target. - destruct (change_frame _) as ((rr,tt),bb) eqn : Hchangee. - assert (((r,t),b) == ((rr,tt),bb)). - rewrite <- Hchange, <- Hchangee. - reflexivity. - do 2 destruct H2. - simpl in H2, H3, H4. - assert ((fst (fst (fst (snd ( - ((comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_far))), snd (conf (Good g_far)))))))) <? - (fst (fst (fst (snd (conf (Good g)))))) == true). - rewrite <- Nat.ltb_lt in Hlower_target. - rewrite <- Hlower_target. - f_equiv. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. - repeat split. - exists (Good g). - destruct b; auto. - unfold Datatypes.id. - rewrite 2 andb_true_iff; repeat split. - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - rewrite Rle_bool_true_iff. - intro Hdist_0. - rewrite Hconf. - simpl (fst _). - destruct b; rewrite dist_same; - lra. - unfold get_alive; rewrite Hconf; simpl in *; assumption. - unfold get_alive; simpl in *; assumption. - rewrite Hconf; simpl in *; rewrite Nat.leb_le; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - rewrite Hconf; simpl. - fold Frame. fold Obs_ILA. - fold Frame in Hchangee. fold Obs_ILA in Hchangee. - destruct (change_frame _) eqn : Heq. - destruct p. - rewrite H3. - assert (tt = r1). assert (snd (fst (rr,tt,bb)) = (snd (fst (r0,r1,b0)))). - auto. rewrite Hchangee. auto. - now simpl in H6. - auto. - simpl. rewrite H4, H6; clear H6. - assert (rr = r0). assert (fst (fst (rr,tt,bb)) = (fst (fst (r0,r1,b0)))). - auto. rewrite Hchangee. auto. - auto. - rewrite H2, H6; clear H6. - assert (bb = b0). assert (snd (rr,tt,bb) = (snd (r0,r1,b0))). - auto. rewrite Hchangee. auto. - now simpl in H6. - rewrite H6. - destruct b0; simpl; reflexivity. - reflexivity. - simpl in H5. - rewrite Nat.ltb_lt in H5. - apply H5. - unfold get_alive; simpl. - unfold round. - assert(Hpred_backup := Hpred). - destruct Hpred as (?,?). - specialize (H3 (Good g_far)). - rewrite H3. - simpl ((lift _ _)). - unfold upt_aux, map_config. - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - destruct (upt_robot_dies_b _ g_far) eqn : Hbool_far; try now simpl. - assert (get_based (conf (Good g_far)) = false). - { - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - rewrite Hconf_far. - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. destruct H1; lia. - - } - rewrite Hconf_far in H4; unfold get_based in H4; simpl in H4; rewrite H4. - - - simpl. - specialize (Hexecuted (g_far)). - assert (Halive_far : get_alive (round rbg_ila da conf (Good g_far)) == false). - rewrite round_good_g. - simpl ; unfold upt_aux, map_config. - rewrite Hbool_far. - unfold get_alive; simpl. - rewrite Hconf_far, H4; - now simpl. - apply Hpred_backup. - assert (get_light (conf (Good g_far)) == true). - { - apply (Hexecuted da). - apply Hpred_backup. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far,b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - set (target := choose_target _) in *. - - rewrite <- H5. rewrite 2 andb_true_iff in H0. destruct H0 as ((?,?),?). - rewrite <- (get_alive_compat H). auto. - now simpl in *. - } - set (target := choose_target _) in *. - - assert (Htarget_light : get_light target = true). - rewrite (get_light_compat H). - rewrite Hconf_far in H5. - now unfold get_light; simpl in *. - assert (Hlight_off_first := light_off_first (reflexivity _) Htarget_light). - unfold For_all in *. - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - unfold obs_from_config, Obs_ILA in Hdist_dp. - assert ((dist pos_r p_far) <= Dp)%R. - { destruct Hconfr as (Hpos_rl,_). - revert Hpos_rl; intros Hpos_rl. - assert (Hpos_rl' : @equiv R2 R2_Setoid (retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (new_pos, false))) pos_r) - by (rewrite <- Hpos_rl; destruct b; now simpl). - rewrite <- Inversion in Hpos_rl'. - destruct b; - rewrite (RelationPairs.fst_compat _ _ _ _ H) in Hdist_dp; - - simpl (fst _) in Hpos_rl'; - unfold equiv, R2_Setoid in Hpos_rl'; - rewrite <- Hpos_rl' in Hdist_dp; - rewrite (frame_dist pos_r p_far (@change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g)); - unfold frame_choice_bijection, Frame; fold Frame; try rewrite Hchange; - simpl in *; lra. - } - unfold upt_robot_dies_b in Hbool_far. - rewrite orb_true_iff in Hbool_far. - destruct Hbool_far as [Htoo_close_so_lighted|Htoo_far_but_path_conf]. - rewrite existsb_exists in *. - destruct Htoo_close_so_lighted as (too_close, (Hin_too_close, Hdist_too_close)). - ++ unfold executioner_means_light_off in *. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff in Hin_too_close. - destruct Hin_too_close as (([g_too_close | byz], (Hspec_too_close, Hnames_too_close)), ((Hident_too_close, Halive_too_close), Hbased_too_closed)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - specialize (Hexecutioner g_too_close da Hpred_backup). - rewrite <- Hspec_too_close in Halive_too_close. - unfold get_alive in *. - simpl (snd (snd _)) in *. - specialize (Hexecutioner Halive_too_close). - assert (Hex : (exists g' : G, - snd (fst (snd (conf (Good g')))) = true /\ - get_based (conf (Good g')) = false /\ - get_ident (conf (Good g_too_close)) < get_ident (conf (Good g')) /\ - (dist (get_location (conf (Good g_too_close))) - (get_location (conf (Good g'))) <= D)%R)). - { exists g_far. - repeat split; try now simpl. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - rewrite <- (get_alive_compat H) in H7. - rewrite 2 andb_true_iff in H0. - destruct H0 as ((H0,H0_aux), H0'). - unfold get_alive in H7; simpl in H7; rewrite H0_aux in H7. - now symmetry. - rewrite Hconf_far, H4; unfold get_based; now simpl. - rewrite Nat.ltb_lt in Hident_too_close. - rewrite <- Hspec_too_close in Hident_too_close. - unfold get_ident in *; simpl in *; auto. - rewrite <- Hspec_too_close, Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite (frame_dist _ _ (change_frame da conf g_far)). - unfold frame_choice_bijection, Frame; now simpl in *. - } - specialize (Hexecutioner Hex). - clear Hex. - - - specialize (Hlight_off_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))). - unfold equiv, bool_Setoid, eq_setoid in Hlight_off_first. - rewrite <- Hlight_off_first, <- Hexecutioner. - unfold get_light; now simpl in *. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. repeat split. - exists (Good g_too_close). - destruct b; reflexivity. - rewrite 2 andb_true_iff. - rewrite Rle_bool_true_iff. - replace (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close))))%R - with - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))))%R. - unfold Datatypes.id. - assert (Hframe := frame_dist (fst (conf (Good g_too_close))) pos ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (dist (fst (conf (Good g_too_close))) pos <= Dmax)%R. - rewrite Rle_bool_true_iff in Hdist_too_close. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_too_close. - rewrite <- Hspec_too_close in Hdist_too_close. - assert ((dist (fst (conf (Good g_too_close))) - (fst (conf (Good g_far))) <= D)%R). - rewrite (frame_dist _ _ (change_frame da conf g_far)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - now simpl; simpl in Hdist_too_close. - rewrite (RelationPairs.fst_compat _ _ _ _ H) in H0. - simpl (fst _) in H0. - revert H0; intros H0. - specialize (Hless_that_Dp g). - unfold get_alive in Hless_that_Dp; - rewrite Hconf, Halive_r in Hless_that_Dp; - simpl (snd (snd _)) in Hless_that_Dp. - specialize (Hless_that_Dp (reflexivity _)). - destruct (Rle_bool (dist pos (fst (conf (Good g_far)))) Dp) eqn : Hhow_far. - rewrite Rle_bool_true_iff, dist_sym in Hhow_far. - assert (Hti := RealMetricSpace.triang_ineq (fst (conf (Good g_too_close))) (fst (conf (Good g_far))) pos ). - rewrite Hconf_far in Hti at 2. - simpl ( (fst _)) in Hti. - rewrite dist_sym in H6. - generalize (D_0), Dmax_7D. - unfold Dp in *. - rewrite Hconf_far in *; simpl (fst _) in *. - lra. - rewrite Rle_bool_false_iff in Hhow_far. - assert (Hlight_close_first := - @light_close_first - _ _ (reflexivity _) Htarget_light). - apply Rnot_le_gt in Hhow_far. - assert ((dist (0, 0) - (get_location target) > Dp)%R). - destruct Hpred_backup as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite <- Horigin. revert H; intro H; rewrite H. - rewrite (frame_dist _ _ (r,t,b)) in Hhow_far. - unfold frame_choice_bijection, Frame in *. - fold Frame in *. rewrite Hchange in *. - revert Hhow_far. rewrite Hconf_far, Hconf. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - now destruct b; simpl in *. - specialize (Hlight_close_first H8). - clear H8. - unfold For_all in Hlight_close_first. - revert Hless_that_Dp; intros Hless_that_Dp. - assert (get_alive (conf (Good g_far)) = true). - { rewrite 2 andb_true_iff in H0; destruct H0 as ((H0, Haux),H0'). - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - rewrite <- H8. - rewrite <- (get_alive_compat H). - unfold get_alive in *; auto. - } - specialize (Hexecuted da Hpred_backup H8 Halive_far). - destruct Hless_that_Dp as (g_less, (Halive_less, (Hident_less, Hpos_less))). - rewrite Hconf in *. rewrite Halive_r in *. lia. - rewrite <- Nat.le_succ_l in Hident_g'. - intros g_near Halive_near Hdist_near Hindent_near. - simpl. - assert (Hlight_off_first_2 := @light_off_first obs _ (reflexivity _)). - unfold obs, target in *. - rewrite (get_light_compat H) in Hlight_off_first_2. - unfold get_light in *; rewrite Hconf_far in Hexecuted; - simpl ( snd (fst _)) in *. - specialize (Hlight_off_first_2 Hexecuted). - unfold For_all in Hlight_off_first_2. - specialize (Hlight_off_first_2 (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near))))). - destruct (conf (Good g_near)) as (p_near, (((i_near, l_near), a_near), b_near)) eqn : Hconf_near. - simpl (snd (fst _)) in *. - apply Hlight_off_first_2. - { - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, - 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g_near). - rewrite Hconf_near; destruct b; now simpl. - assert (dist p_near pos <= Dmax)%R. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - now simpl in *. - rewrite (frame_dist _ _ (r,t,b)) in H9. - unfold frame_choice_bijection, Frame in H9; fold Frame in H9. - destruct b; simpl in *. lra. - lra. - now simpl in *. - now simpl in *. - - unfold get_ident in *; simpl in *; auto. rewrite Nat.leb_le; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. - } - - assert (Hpos_less' := Hpos_less). - apply Rle_not_gt in Hpos_less. - destruct Hpos_less. - - destruct Hpred_backup as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *; - simpl (fst _) in *. - - specialize (Hlight_close_first (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_less))), snd (conf (Good g_less))))). - clear H8. - assert (dist (0, 0)%R - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_less))), - snd (conf (Good g_less)))) == - dist pos (fst (conf (Good g_less)))). - { - rewrite <- Horigin. - rewrite (frame_dist pos _ (change_frame da conf g)). - unfold frame_choice_bijection, Frame in *. - destruct (change_frame _) as ((?,?),xb) eqn : Hxchange; - destruct xb, b; try discriminate; - try (assert (fst (fst (r0, r1, true)) = fst (fst (r, t, true)) /\ snd (fst (r0, r1, true)) = snd (fst (r, t, true))) - by ( - rewrite Hchange; - split; reflexivity)); - try (assert (fst (fst (r0, r1, false)) = fst (fst (r, t, false)) /\ snd (fst (r0, r1, false)) = snd (fst (r, t, false))) - by ( - rewrite Hchange; - split; reflexivity)); - simpl in H8; - destruct H8 as (H8_a, H8_b); - try rewrite H8_a ,H8_b; - try now simpl in *. - } - rewrite <- H8. - assert (Htemp : (dist (0, 0) - (fst - (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) - (fst (conf (Good g_less))), snd (conf (Good g_less)))) > Dp)%R \/ - fst - (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) - (fst (conf (Good g_less))), snd (conf (Good g_less))) = (0%R, 0%R)). - { - apply Hlight_close_first. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, - 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists ((Good g_less)). - destruct b; reflexivity. - simpl (fst _). - rewrite (frame_dist _ _ (r,t,b)) in Hpos_less'. - unfold frame_choice_bijection, Frame in *; fold Frame in *. - unfold Dp in *; generalize D_0; destruct b; - rewrite dist_sym in Hpos_less'; - simpl in *; - lra. - now simpl in *. - now simpl in *. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat (Hxy)). - reflexivity. - } - destruct Htemp; try auto. - assert (g_less <> g). - { - destruct (Geq_dec g_less g). - rewrite e in Hident_less. - rewrite Hconf in Hident_less. - unfold get_ident in *; simpl in *; lia. - - auto. - } - specialize (Hcol g_less g H10). - rewrite Hconf in Hcol. - unfold get_based in Hcol. - simpl (snd _) in Hcol. - assert (get_based (conf (Good g_less)) = false). - { - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (_,(_,Hbh)). - destruct (get_based (conf (Good g_less))) eqn : Hfalse_based. - specialize (Hbh g_less g Hfalse_based). - rewrite Hconf in Hbh. - unfold get_based in Hbh; simpl (snd _) in Hbh; specialize (Hbh (reflexivity _)). - unfold get_ident in *; simpl in *; lia. - reflexivity. - } - unfold get_based in H11; specialize (Hcol H11 (reflexivity _)). - unfold get_alive in *; simpl (snd _) in Hcol; - specialize (Hcol Halive_less Halive_r). - rewrite <- Horigin in H9. - rewrite (frame_dist _ _ (r,t,b)) in Hcol. - revert Hcol; intro Hcol. - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - simpl (fst _) in Hcol. - rewrite <- Hchange in Hcol. - exfalso. - apply Hcol. - rewrite <- H9. - rewrite Hchange. - simpl (frame_choice_bijection _); simpl (get_location _). - unfold Datatypes.id. - assert (Htrue := dist_same (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good g_less))))). - rewrite <- Htrue. - repeat f_equiv. - destruct b; auto. - - rewrite (frame_dist _ _ ((r,t),b)) in H7. - now unfold frame_choice_bijection, Frame; destruct b; simpl in *. - now simpl. - rewrite Nat.leb_le. - rewrite Nat.ltb_lt in *. - transitivity (get_ident (((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far))))). - unfold get_ident in *; rewrite <- Hspec_too_close, Hconf_far in *; - simpl in *. lia. - assert (get_ident - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far))) < - get_ident - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (pos, (ident, light, alive, false))), snd (pos, (ident, light, alive, false)))). - assert (Htarget_lower:= target_lower). - specialize (Htarget_lower obs - - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far))) - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (pos, (ident, light, alive, false))), snd (pos, (ident, light, alive, false))) - - - ). - unfold get_ident in *; apply Htarget_lower. - unfold obs, obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. - repeat split. - exists (Good g). - destruct b; rewrite Hconf; auto. - unfold Datatypes.id. - - rewrite 2 andb_true_iff; repeat split. - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - rewrite Rle_bool_true_iff. - intro Hdist_0. - destruct b; rewrite dist_same; - lra. - unfold get_alive; simpl in *; assumption. - unfold get_alive; simpl in *; assumption. - rewrite Nat.leb_le. - reflexivity. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - unfold get_ident in Hcompat. - rewrite Hcompat. - reflexivity. - destruct Hpred_backup as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hchange, Hconf; destruct b; simpl; reflexivity. - destruct b; rewrite <- H; unfold target, obs, obs_from_config, Obs_ILA; - unfold get_alive, get_ident; reflexivity. - unfold get_ident in *; simpl in *; - destruct b; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat (Hxy)). - reflexivity. - ++ specialize (Hpath_backup g_far). - assert (get_alive (conf (Good g_far)) == true). - rewrite 2 andb_true_iff in H0. - destruct H0 as ((H0, H0_aux), H0'). - rewrite <- H0_aux, (get_alive_compat H). - unfold get_alive; rewrite Hconf_far; now simpl. - simpl in H7. - specialize (Hpath_backup H7); clear H7. - destruct Hpath_backup as [Hident_0|Hpath_backup]. - rewrite (ident_preserved _ _ Hpred_backup) in Hident_0. - assert (get_alive (round rbg_ila da conf (Good g_far)) = true). - apply ident_0_alive; intuition. - rewrite H7 in *; discriminate. - rewrite forallb_existsb, forallb_forall in Htoo_far_but_path_conf. - destruct Hpath_backup as (g_too_close, (Halive_close,( Hdist_close, (Hident_close, Hbased_close)))). - specialize (Htoo_far_but_path_conf - ((((let (p, b) := change_frame da conf g_far in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_too_close))), snd (conf (Good g_too_close)))))). - rewrite negb_true_iff, Rle_bool_false_iff in Htoo_far_but_path_conf. - destruct Htoo_far_but_path_conf. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - ** exists (Good g_too_close). - split. - destruct (change_frame da conf g_far) as (?,b0); destruct b0; - now simpl. - apply In_names. - ** rewrite Nat.ltb_lt. - unfold get_ident in *; simpl in Hident_close; simpl. - auto. - ** rewrite <- Halive_close. - now unfold get_alive; simpl. - ** rewrite negb_true_iff; unfold get_based in *; simpl in *; auto. - ** rewrite dist_sym, (frame_dist _ _ (change_frame da conf g_far)) in Hdist_close. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - unfold frame_choice_bijection, Frame in Hdist_close; fold Frame in *. - revert Hdist_close; intros Hdist_close. - destruct (change_frame _ _ g_far) as ((r_f, t_f), b_f) - eqn : Hchange_far. - now destruct b_f; simpl in *. - ++ simpl. - rewrite 2 andb_true_iff in H0. - destruct H0 as ((H0, H0_aux), H0'). - assert (get_based (conf (Good g_far)) = false). - { - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - rewrite Hconf_far. - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. destruct H1; lia. - - } - rewrite Hconf_far in H4; unfold get_based in H4; simpl in H4; rewrite H4. - simpl. - - rewrite <- H0_aux. - rewrite (get_alive_compat H). - now unfold get_alive; simpl. - ++ - assert (get_based (conf (Good g_far)) = false). - { - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - destruct (Hbased_higher) as (Hap0,(Hex, Hhigher)). - unfold get_based; - destruct b_far; try now simpl. - specialize (Hhigher g_far g). - rewrite Hconf_far, Hconf in Hhigher. - rewrite (get_ident_compat H) in H1. - unfold get_based, get_ident in *. - simpl in *; specialize (Hhigher (reflexivity _) (reflexivity _)). - rewrite Nat.leb_le in *. lia. - } - rewrite negb_true_iff. - unfold get_based; simpl. - unfold round. - destruct (Hpred); try auto. - rewrite (H4 (Good g_far)). - simpl. - unfold upt_aux, map_config. - destruct (conf (Good g_far)) as (p_far, (((i_far,l_far),a_far), b_far)) eqn : Hconf_far. - unfold get_based in *; simpl in H2. - rewrite H2. - simpl. - destruct (upt_robot_dies_b _ g_far); simpl; try auto. - - } - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - rewrite 2 andb_true_iff in H0; destruct H0 as ((H0, H0'), Haux). - rewrite Rle_bool_true_iff in H0. - rewrite (RelationPairs.fst_compat _ _ _ _ H) in H0. - revert H0; intros H0. - assert (Hdist_round : (dist (fst (conf (Good g_far))) (fst (round rbg_ila da conf (Good g_far))) <= D)%R). - rewrite <- Rle_bool_true_iff. - assert (Hd := dist_round_max_d g_far Hpred Hpath_backup). - destruct (conf (Good g_far)) as (p_far,(((i_far, l_far), a_far), b_far)) eqn : Hconf_far. - assert (get_alive ( ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_far, (i_far, l_far, a_far, b_far))), snd (p_far, (i_far, l_far, a_far, b_far)))) - == - get_alive (conf (Good g_far))). - unfold get_alive; rewrite Hconf_far; now simpl. - rewrite <- (get_alive_compat H) in H2. - rewrite Hconf_far in H2. - rewrite H0' in H2. - now specialize (Hd (symmetry H2)). - assert (Hconf_r_equiv : round rbg_ila da conf (Good g) == (pos_r, (ident_r, light_r, alive_r, based_r))) by now rewrite Hconf_r. - rewrite (fst_compat Hconf_r_equiv). - - simpl ((fst (pos_r, (ident_r, light_r, alive_r)))). - assert (dist ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_far)))) - ((comp (bij_rotation r_r) - (comp (bij_translation t_r) (if b_r then reflexion else Similarity.id))) - pos_r) <= Dmax)%R. - assert (dist (fst (round rbg_ila da conf (Good g_far))) pos_r <= Dmax)%R. - destruct (conf (Good g_far)) as (p_far, (((i_far, l_far), a_far), b_far)) eqn : Hconf_far. - assert ((dist pos_r p_far) <= Dp)%R. - { destruct Hconfr as (Hpos_rl,_). - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - - assert (Hpos_rl' : retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (new_pos, false)) == pos_r) by (rewrite <- Hpos_rl; destruct b; now simpl). - rewrite <- Inversion in Hpos_rl'. - simpl (fst _) in Hpos_rl'. - unfold equiv, R2_Setoid in Hpos_rl'; - simpl (obs_from_config _ _) in *. - rewrite <- Hpos_rl' in Hdist_dp. - destruct b; unfold new_pos, obs in *; - rewrite (fst_compat H) in Hdist_dp; - rewrite (@frame_dist pos_r p_far (@change_frame (prod R2 ILA) Loc State_ILA Identifiers - (prod R2 NoCollisionAndPath.light) (prod (prod R R2) bool) bool bool - robot_choice_RL Frame Choice inactive_choice_ila da conf g)); - unfold frame_choice_bijection, Frame; fold Frame; try rewrite Hchange; - simpl in *; lra. - } - simpl (fst _) at 1 in Hdist_round. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_far))) p_far pos_r). - rewrite dist_sym in Hdist_round, H2. - generalize D_0, Dmax_7D. - unfold Dp in *. lra. - rewrite (frame_dist _ _ ((r_r, t_r), b_r)) in H2. - unfold frame_choice_bijection, Frame in H2; fold Frame in H2. - destruct b_r; now simpl in *. - destruct b_r; now simpl in *. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), - (get_alive_compat Hxy), (get_ident_compat Hxy); - reflexivity. - - assert (get_based (round rbg_ila da conf (Good g)) = false). - rewrite (round_good_g); try auto. - simpl. - unfold upt_aux, map_config. - simpl. - rewrite Hconf. - unfold get_based; simpl. - now destruct (upt_robot_dies_b _); simpl. - rewrite Hconf_r in H; unfold get_based in H; simpl in H; rewrite H in *. - simpl in *. - destruct (upt_robot_too_far _) eqn : Htoo_far; simpl in *; - rewrite Hconf_r in *; - simpl in *; - intuition; rewrite Halive_r, Halive_rr in *; discriminate. -Qed. - - - -Lemma executioner_means_light_off_round : - forall conf da, - da_predicat da -> - path_conf conf -> - based_higher conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - executioner_means_light_off (round rbg_ila da conf). - -Proof. - intros conf da Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hexists. - assert (Hexecuted_round := executed_means_light_on_round Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hexists). - intros g da' Hpred' Halive_r (g_dead, (Halive_dead, (Hbased_dead, (Hident_dead, Hdist_dead)))). - destruct (get_based (conf (Good g))) eqn : Hbased. - - destruct (get_based (round rbg_ila da conf (Good g))) eqn : Hfalse_r. - assert (Haux : exists g', get_based (conf (Good g')) = true) by now exists g. - assert (Hbased_higher_round := based_higher_round Hpred Hpath Hbased_higher Hexecutioner Hexecuted Haux). - clear Haux. - destruct Hbased_higher_round as (Hap0_r, (Hex_r, Hhi_r)). - specialize (Hhi_r g g_dead Hfalse_r Hbased_dead). - lia. - destruct (get_based (conf (Good g_dead))) eqn : Hfalse. - + rewrite round_good_g in Hbased_dead, Hfalse_r; try auto. - simpl in *. - unfold upt_aux, map_config in *. - destruct (conf (Good g_dead)) as (p_d,(((i_d, l_d), a_d), b_d)) eqn : Hconf_dead. - destruct (conf (Good g)) as (pos, (((ide,lig),ali),bas)) eqn : Hconf. - unfold get_based in *; simpl in *; rewrite Hfalse, Hbased in *; simpl in *. - - destruct (upt_robot_too_far _) eqn: Htoo_far in Hbased_dead; try now simpl; auto. - destruct (upt_robot_too_far _) eqn: Htoo_far' in Hfalse_r; try now simpl. - unfold upt_robot_too_far in *. - destruct (R2dec_bool _ _); try now simpl in *. - destruct (R2dec_bool _ _) in Htoo_far'; try now simpl in *. - - destruct (forallb _ _) eqn : Htoo_far_1 in Htoo_far; try now simpl in *. - destruct (forallb _ _) eqn : Htoo_far_1' in Htoo_far'; try now simpl in *. - - rewrite forallb_forall in *. - - unfold get_based, get_alive in Htoo_far_1. - simpl (snd (_)) in Htoo_far_1. - assert (Hin_g' : List.In g_dead (List.filter (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) - Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g_dead). - rewrite Hconf_dead. - simpl. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g_dead). - rewrite Hconf_dead in Hap0; unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - assert (Hin_g : List.In g (List.filter (fun g' : G => snd (snd (conf (Good g'))) && snd (fst (snd (conf (Good g'))))) - Gnames)). - { - rewrite filter_In. - split; try apply (In_Gnames g). - rewrite Hconf. - simpl. - destruct Hbased_higher as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in Hap0; unfold get_alive, get_based in Hap0; simpl in Hap0. - intuition. - } - specialize (Htoo_far_1 g Hin_g). - unfold get_based, get_alive in Htoo_far_1'; simpl (snd _) in Htoo_far_1'. - specialize (Htoo_far_1' g_dead Hin_g'). - unfold too_far_aux_1 in *. - destruct (negb _) eqn : Hident_far_1 in Htoo_far_1; simpl in Htoo_far_1; try discriminate. - unfold get_ident in Htoo_far_1; simpl in Htoo_far_1. - destruct (negb _) eqn : Hident_far_1' in Htoo_far_1'; simpl in Htoo_far_1'; try discriminate. - unfold get_ident in Htoo_far_1'; simpl in Htoo_far_1'. - rewrite Nat.ltb_lt in *. - lia. - rewrite negb_true_iff , negb_false_iff in *. - rewrite Nat.eqb_eq, Nat.eqb_neq in *. - destruct Hident_far_1; unfold get_ident in *; simpl in *; auto. - assert (Hg : g <> g_dead). - intro H; rewrite H in *. - lia. - destruct (ident_unique conf Hg). - rewrite negb_false_iff, Nat.eqb_eq in *; unfold get_ident in *; simpl in *; auto. - + destruct (Hbased_higher) as (?,(?, Hhi)). - specialize (Hhi g g_dead Hbased Hfalse). - rewrite <- 2 ident_preserved in Hident_dead; try auto. - lia. - - - destruct (get_based (conf (Good g_dead))) eqn : Hbased_false. - * destruct (Hbased_higher) as (Hap0,(Hex, Hhi)). - specialize (Hap0 g_dead Hbased_false). - destruct (conf (Good g_dead)) as (p_d,(((i_d, l_d), a_d), b_d)) eqn : Hconf_dead. - destruct (conf (Good g)) as (pos, (((ide,lig),ali),bas)) eqn : Hconf. - rewrite round_good_g in Hbased_dead; try auto. - simpl in Hbased_dead; unfold upt_aux, map_config in Hbased_dead; unfold get_based in *; - simpl in Hbased_false; simpl in Hbased_dead; rewrite Hconf_dead, Hbased_false in Hbased_dead. - simpl in Hbased_dead. - destruct (upt_robot_too_far _) eqn : Htoo_far; try (now simpl in *). - unfold upt_robot_too_far in *. - destruct (R2dec_bool _) eqn : Hpos_dead; try now simpl in *. - destruct (forallb _) eqn : Htoo_far_1 in Htoo_far; try (now simpl); - destruct (forallb _) eqn : Htoo_far_2 in Htoo_far; try now simpl. - unfold too_far_aux_2 in Htoo_far_2. - rewrite forallb_forall in Htoo_far_2. - specialize (Htoo_far_2 g (In_Gnames g)). - rewrite negb_true_iff, 2 andb_false_iff in *. - destruct Htoo_far_2 as [Hbased_g_false| [Halive_g_false|Hdist_g_false]]. - rewrite negb_false_iff in *. - unfold get_based in *; rewrite Hconf in *; simpl in *; rewrite Hbased_g_false in *; discriminate. - apply still_alive_means_alive in Halive_r; try auto. - unfold get_alive in *; rewrite Hconf in *; simpl in *; rewrite Halive_g_false in *; discriminate. - assert (dist pos p_d > (Dp -3*D))%R. - rewrite (frame_dist _ _ (change_frame da conf g_dead)). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite Hconf, Hconf_dead, Rle_bool_false_iff in Hdist_g_false; - simpl in *. - lra. - assert (dist pos p_d > 3*D)%R. - unfold Dp in *. - generalize Dmax_7D, D_0; lra. - assert (dist pos p_d <= 3 * D)%R. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g))) p_d). - transitivity (dist (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g))) + - dist (get_location (round rbg_ila da conf (Good g))) p_d)%R. - rewrite Hconf in *; simpl in *; auto. - transitivity (D + dist (get_location (round rbg_ila da conf (Good g))) p_d)%R. - apply Rplus_le_compat_r. - assert (Htrue := dist_round_max_d g Hpred Hpath (still_alive_means_alive g Hpred Hbased_higher Halive_r)). - unfold equiv, bool_Setoid, eq_setoid in *; rewrite Rle_bool_true_iff in Htrue; auto. - replace (3*D)%R with (D+ (D+D))%R by lra. - apply Rplus_le_compat_l. - replace p_d with (get_location (conf (Good g_dead))). - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (round rbg_ila da conf (Good g_dead))) (get_location (conf (Good g_dead)))). - transitivity (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (round rbg_ila da conf (Good g_dead))) + - dist (get_location (round rbg_ila da conf (Good g_dead))) (get_location (conf (Good g_dead))))%R; try auto. - transitivity (D+ dist (get_location (round rbg_ila da conf (Good g_dead))) (get_location (conf (Good g_dead))))%R. - now apply Rplus_le_compat_r. - apply Rplus_le_compat_l. - assert (Htrue := dist_round_max_d g_dead Hpred Hpath (still_alive_means_alive g_dead Hpred Hbased_higher Halive_dead)). - unfold equiv, bool_Setoid, eq_setoid in *; rewrite Rle_bool_true_iff in Htrue; auto. - now rewrite dist_sym. - now rewrite Hconf_dead; simpl. - lra. - * apply (moving_means_light_off conf g Hpred Hbased Halive_r). - specialize (Hexecuted_round g_dead da' Hpred' Halive_dead). - assert (Halive_dead_r: get_alive - (round rbg_ila da' (round rbg_ila da conf) (Good g_dead)) = - false). - { rewrite (round_good_g g_dead (round rbg_ila da conf)). - unfold get_alive; - simpl. - unfold upt_aux, map_config. - destruct (upt_robot_dies_b _) eqn : Hbool_dead. - + destruct ((round rbg_ila da conf (Good g_dead))) as (?,(((?,?),?),?)); - unfold get_based in *; simpl in *; rewrite Hbased_dead; intuition. - + destruct ((round rbg_ila da conf (Good g_dead))) as (pos_dead_r,(((ident_dead_r,light_dead_r),alive_dead_r), based_dead_r)) eqn : Hconf_dead_r; - simpl; intuition. - unfold get_based in *; simpl in Hbased_dead; rewrite Hbased_dead. - unfold upt_robot_dies_b in Hbool_dead. - rewrite orb_false_iff in Hbool_dead. - destruct Hbool_dead as (Hnear, _). - rewrite <- negb_true_iff, forallb_existsb, forallb_forall in Hnear. - unfold get_alive in *; simpl in Halive_dead; rewrite Halive_dead. - destruct (change_frame _ (round _ _ _) g_dead) as ((r_dead, t_dead), b_dead) eqn : Hchange_r'. - specialize (Hnear - (((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if b_dead then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g))), - snd (round rbg_ila da conf (Good g))))). - rewrite <- Hnear. - simpl. - rewrite negb_false_iff, Rle_bool_true_iff. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - clear Hnear. - rewrite <- Hconf_dead_r in Hdist_dead. - assert (dist - ((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if b_dead then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g)))) - ((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if b_dead then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_dead)))) <= D)%R. - rewrite (frame_dist _ _ ((r_dead, t_dead), b_dead)) in Hdist_dead. - unfold frame_choice_bijection, Frame in Hdist_dead. - now destruct b_dead; simpl in *. - now destruct b_dead; simpl in *. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - exists (Good g). - split; try apply In_names; try (destruct b_dead; reflexivity). - rewrite Nat.ltb_lt. - unfold get_ident in *. rewrite <- Hconf_dead_r in *; now simpl in *. - rewrite <- Hconf_dead_r in *; now simpl in *. - rewrite negb_true_iff. - destruct (conf (Good g)) as (pos,(((ide,lig), ali),bas)) eqn : Hconf. - simpl in Hbased. - unfold get_based; simpl. - assert (get_based (round rbg_ila da conf (Good g))=false-> snd (snd (round rbg_ila da conf (Good g))) = false ) by now unfold get_based. - apply H. - rewrite round_good_g; try auto. - unfold get_based; simpl. - unfold upt_aux, map_config; simpl. - rewrite Hconf, Hbased. - simpl. - now destruct (upt_robot_dies_b _). - + apply Hpred'. - } - specialize (Hexecuted_round Halive_dead_r). - assert (Hexecuted_not_moving := light_on_means_not_moving_before g_dead Hpred Hpath Halive_dead Hexecuted_round). - intros Hnot_moving. - rewrite round_good_g in Halive_dead; try apply Hpred. - simpl in Halive_dead. - unfold upt_aux, map_config in *. - destruct (conf (Good g_dead)) as (p_dead, (((i_dead, l_dead), a_dead), b_dead)) eqn : Hconf_dead; simpl in Halive_dead. - unfold get_based in Hbased_false. - simpl in Hbased_false. - rewrite Hbased_false in *. - destruct (upt_robot_dies_b _) eqn : Hbool_r; - unfold get_alive in Halive_dead; simpl in Halive_dead. - + discriminate. - + unfold upt_robot_dies_b in Hbool_r. - rewrite orb_false_iff in Hbool_r. - destruct Hbool_r as (Hnear, _). - rewrite <- negb_true_iff, forallb_existsb, forallb_forall in Hnear. - destruct (change_frame _ _ g_dead) as ((r_dead, t_dead), bb_dead) eqn : Hchange_dead. - specialize (Hnear ((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if bb_dead then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hnear. - destruct Hnear. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - exists (Good g). - split; try apply In_names; try (destruct bb_dead; reflexivity). - rewrite Nat.ltb_lt. - rewrite <- 2 ident_preserved in Hident_dead. - unfold get_ident in *. rewrite <- Hconf_dead in *. now simpl in *. - apply Hpred. - apply Hpred. - apply still_alive_means_alive in Halive_r. - unfold get_alive in *; now simpl in *. - apply Hpred. - auto. - rewrite negb_true_iff. - unfold get_based in *. - now simpl. - assert (dist - ((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if bb_dead then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g)))) - ((comp (bij_rotation r_dead) - (comp (bij_translation t_dead) - (if bb_dead then reflexion else Similarity.id))) - (fst (round rbg_ila da conf (Good g_dead)))) <= D)%R. - rewrite (frame_dist _ _ ((r_dead, t_dead), bb_dead)) in Hdist_dead. - unfold frame_choice_bijection, Frame in Hdist_dead. - now destruct bb_dead; simpl in *. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hnot_moving, Hconf_dead, Hexecuted_not_moving. - now destruct bb_dead; simpl in *. - - - -Qed. - -Parameter conf_init : config. - - -Lemma executioner_next_means_moving : - forall conf g da (* da' *), - executed_means_light_on (round rbg_ila da conf) -> - da_predicat da -> based_higher conf -> path_conf conf -> - get_alive (conf (Good g)) = true -> - (exists g', get_alive (round rbg_ila da conf (Good g')) = true /\ - get_based (round rbg_ila da conf (Good g')) = false /\ - get_ident (round rbg_ila da conf (Good g)) < get_ident (round rbg_ila da conf (Good g')) /\ - (dist (get_location (round rbg_ila da conf (Good g))) (get_location (round rbg_ila da conf (Good g'))) - <= D)%R) -> - get_location (conf (Good g)) <> get_location (round rbg_ila da conf (Good g)). -Proof. - intros conf g da (*da'*) Hex_light Hpred Hbased_higher Hpath Halive (g', (Halive_exec, (Hbased_exec ,(Hident_exec, Hdist_exec)))). - assert (Halive_exec_backup := Halive_exec). - rewrite round_good_g in Halive_exec, Hbased_exec; try apply Hpred. - unfold get_alive in Halive_exec; simpl in Halive_exec, Hbased_exec. - unfold upt_aux, map_config in Halive_exec, Hbased_exec. - destruct (conf (Good g')) as (pos', (((ident', light'), alive'), based')) eqn : Hconf'. - destruct based' eqn : Hbased'. - *** unfold get_based in Hbased_exec; simpl in Hbased_exec. - destruct (conf (Good g)) as (pos,(((ide, lig), ali), bas)) eqn : Hconf. - destruct (upt_robot_too_far _) eqn : Htoo_far; try now simpl in *. - unfold upt_robot_too_far in *. - destruct (R2dec_bool _) eqn : Hpos_dead; try now simpl in *. - destruct (forallb _) eqn : Htoo_far_1 in Htoo_far; try (now simpl); - destruct (forallb _) eqn : Htoo_far_2 in Htoo_far; try now simpl. - unfold too_far_aux_2 in Htoo_far_2. - rewrite forallb_forall in Htoo_far_2. - specialize (Htoo_far_2 g (In_Gnames g)). - rewrite negb_true_iff, 2 andb_false_iff in *. - destruct Htoo_far_2 as [Hbased_g_false| [Halive_g_false|Hdist_g_false]]. - rewrite negb_false_iff in *. - assert (get_based (conf (Good g)) = false). - { destruct (get_based (conf (Good g)))eqn : Hfalse; try auto. - rewrite forallb_forall in Htoo_far_1. - assert ( List.In g - (List.filter - (fun g'0 : G => - get_based - ((let (p, b) := change_frame da conf g' in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g'0))), snd (conf (Good g'0))) && - get_alive - ((let (p, b) := change_frame da conf g' in - let (r, t) := p in - comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g'0))), snd (conf (Good g'0)))) Gnames)). - rewrite filter_In, andb_true_iff; repeat split; try auto. - apply In_Gnames. - rewrite Hconf; apply Halive. - specialize (Htoo_far_1 g H). - unfold too_far_aux_1 in *. - destruct (negb _) eqn : Hnegb_g. - rewrite Nat.ltb_lt in *. - rewrite <- 2 ident_preserved in Hident_exec; try auto. - unfold get_ident in *; rewrite Hconf, Hconf' in *; simpl in *. lia. - rewrite negb_false_iff, Nat.eqb_eq in *. - destruct (Geq_dec g g'). - rewrite e in *; lia. - assert (Hident_d := ident_unique conf n0). - rewrite Hconf, Hconf' in *; unfold get_ident in *; simpl in *; lia. - } - - unfold get_based in *; simpl in *; rewrite H in *; discriminate. - rewrite Hconf in *; unfold get_alive in *; simpl in *; rewrite Halive in *; discriminate. - assert (dist pos pos' > (Dp -3*D))%R. - rewrite (frame_dist _ _ (change_frame da conf g')). - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in *. - rewrite Hconf, Hconf', Rle_bool_false_iff in Hdist_g_false; - simpl in *. - lra. - assert (dist pos pos' > 3*D)%R. - unfold Dp in *. - generalize Dmax_7D, D_0; lra. - assert (dist pos pos' <= 3 * D)%R. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g))) pos'). - transitivity (dist (get_location (conf (Good g))) (get_location (round rbg_ila da conf (Good g))) + - dist (get_location (round rbg_ila da conf (Good g))) pos')%R. - rewrite Hconf in *; simpl in *; auto. - transitivity (D + dist (get_location (round rbg_ila da conf (Good g))) pos')%R. - apply Rplus_le_compat_r. - rewrite <- Hconf in Halive. - assert (Htrue := dist_round_max_d g Hpred Hpath Halive). - unfold equiv, bool_Setoid, eq_setoid in *; rewrite Rle_bool_true_iff in Htrue; auto. - replace (3*D)%R with (D+ (D+D))%R by lra. - apply Rplus_le_compat_l. - replace pos' with (get_location (conf (Good g'))). - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (round rbg_ila da conf (Good g'))) (get_location (conf (Good g')))). - transitivity (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (round rbg_ila da conf (Good g'))) + - dist (get_location (round rbg_ila da conf (Good g'))) (get_location (conf (Good g'))))%R; try auto. - transitivity (D+ dist (get_location (round rbg_ila da conf (Good g'))) (get_location (conf (Good g'))))%R. - now apply Rplus_le_compat_r. - apply Rplus_le_compat_l. - assert (Htrue := dist_round_max_d g' Hpred Hpath (still_alive_means_alive g' Hpred Hbased_higher Halive_exec_backup)). - unfold equiv, bool_Setoid, eq_setoid in *; rewrite Rle_bool_true_iff in Htrue; auto. - now rewrite dist_sym. - now rewrite Hconf'; simpl. - lra. - *** destruct (upt_robot_dies_b _) eqn : Hbool. - - now simpl in *. - - intro Hdist. - assert (dist (get_location (conf (Good g))) (get_location (conf (Good g'))) <= 2*D)%R. - { - rewrite Hdist. - assert (Hdist' := dist_round_max_d g' Hpred Hpath (still_alive_means_alive g' Hpred Hbased_higher Halive_exec_backup)). - assert (Htri := RealMetricSpace.triang_ineq (get_location (conf (Good g'))) (get_location (round rbg_ila da conf (Good g'))) (get_location (round rbg_ila da conf (Good g)))). - apply Rle_bool_true_iff in Hdist'. - assert (dist (get_location (conf (Good g'))) (get_location (round rbg_ila da conf (Good g))) <= - D + - dist (get_location (round rbg_ila da conf (Good g'))) - (get_location (round rbg_ila da conf (Good g))))%R. - lra. - rewrite dist_sym in Hdist_exec. - assert (dist (get_location (conf (Good g'))) (get_location (round rbg_ila da conf (Good g))) <= - D + - D)%R. - lra. - rewrite dist_sym. lra. - } - assert (get_light (round rbg_ila da conf (Good g')) = true). - rewrite round_good_g; try apply Hpred. - simpl; unfold upt_aux, map_config; rewrite Hbool. - unfold rbg_fnc. - destruct (move_to _) eqn : Hmove; try now (unfold get_light; simpl; rewrite Hconf'; simpl). - exfalso. - assert (Hfalse := move_to_Some_zone Hmove). - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - destruct (change_frame da conf g') as ((r, t), b) eqn : Hchange'. - specialize (Hfalse ((fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) (Good g))). - set (obs := obs_from_config _ _) in *. - assert (((fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) (Good g) - ∈ obs)%set). - unfold obs, obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, andb_true_iff. - repeat split. - exists (Good g). - destruct b; reflexivity. - rewrite 2 andb_true_iff. - rewrite Rle_bool_true_iff. - repeat split. - replace (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))%R - with - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))))%R. - unfold Datatypes.id. - assert (Hframe := frame_dist (fst (conf (Good g))) (fst (conf (Good g'))) ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (dist (fst (conf (Good g))) (fst (conf (Good g'))) <= Dmax)%R. - revert H; intro H. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in H. - generalize Dmax_7D D_0. - transitivity (2*D)%R. - apply H. - lra. - rewrite <- Hconf'. - destruct b; simpl in *; lra. - now simpl. - now unfold get_alive in *; simpl in *. - now unfold get_alive in *; simpl in *. - rewrite <- 2 ident_preserved in Hident_exec; auto. - rewrite <- Hconf', Nat.leb_le; unfold get_ident in *; simpl in *; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy). - assert (Hcompat := get_ident_compat Hxy). - rewrite Hcompat. - reflexivity. - assert(Hfalse' : (dist new_pos (fst (( - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))) )) > 2 * D)%R). - unfold obs in *. - rewrite Hconf' in *. - specialize (Hfalse H0); apply Hfalse. - assert (Hpred_backup := Hpred). - destruct Hpred_backup as (Horigin, ?). - specialize (Horigin conf g' (change_frame da conf g') (reflexivity _)). - rewrite Hchange' in *. - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - apply Rnot_le_gt. - intro Hfalse_dist. - unfold upt_robot_dies_b in Hbool. - rewrite orb_false_iff in Hbool. - destruct Hbool as (Hex,_). - rewrite <- negb_true_iff, forallb_existsb, forallb_forall in Hex. - specialize (Hex (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good g))), - snd (conf (Good g)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hex. - apply Hex. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split; try auto. - exists (Good g). split; auto. destruct b; reflexivity. apply In_names. - rewrite Nat.ltb_lt. - rewrite <- 2 ident_preserved in Hident_exec; try apply Hpred. - unfold get_ident in *; simpl in *. - lia. - destruct (get_based (conf (Good g))) eqn : Hbased. - destruct Hbased_higher as (_,(_,Hbased_ident)). - specialize (Hbased_ident g g' Hbased). - rewrite Hconf' in *. - specialize (Hbased_ident (reflexivity _)). - rewrite <- 2 ident_preserved in Hident_exec; try apply Hpred. - unfold get_ident in *; simpl in *. rewrite Hconf' in *; simpl in *. lia. - rewrite negb_true_iff; unfold get_based; simpl in *; auto. - assert (forall x, x <= 0 -> x <= D)%R. - intros; generalize D_0; lra. - apply H2, Hfalse_dist. - clear Hfalse H0; rename Hfalse' into Hfalse. - destruct (round rbg_ila da conf (Good g')) as (pos_round', (((ident_round', light_round'), alive_round'), based_round')) eqn : Hconf_round'. - assert (round rbg_ila da conf (Good g') == (pos_round', (((ident_round', light_round'), alive_round'), based_round'))) by now rewrite Hconf_round'. - rewrite round_good_g in H0; try apply Hpred. - destruct H0 as (Hpos_round',_). - simpl in Hpos_round'. - rewrite Hchange' in Hpos_round'. - unfold upt_aux, map_config in Hpos_round'; rewrite Hbool in Hpos_round'. - unfold rbg_fnc in *. - unfold new_pos in *; rewrite Hmove in *. - rewrite Hconf' in *; simpl (fst _) in Hpos_round'. - assert (Hpos_round_aux : retraction - (comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) new_pos == - pos_round'). - { destruct b; unfold new_pos; rewrite <- Hpos_round'; rewrite Hconf'. f_equiv. unfold comp. simpl. f_equiv. } - rewrite <- Inversion in Hpos_round_aux. - unfold new_pos in *; rewrite Hconf' in *. - rewrite <- Hpos_round_aux in *. - revert Hdist_exec; intro Hdist_exec. - replace (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))%R - with - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))))%R in Hfalse. - unfold Datatypes.id. - assert (Hframe := frame_dist pos_round' (fst (conf (Good g))) ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (dist pos_round' (fst (conf (Good g))) > 2*D)%R. - rewrite Hframe. - destruct b; simpl in *; lra. - clear Hframe. - unfold get_location, State_ILA, AddInfo, OnlyLocation, get_location, Datatypes.id in Hdist_exec. - rewrite <- Hconf_round' in Halive_exec_backup. - assert (Hdist_round := dist_round_max_d g Hpred Hpath (Halive)). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round. - rewrite Rle_bool_true_iff in Hdist_round. - assert (Hd_0 := D_0). - assert (Htri := RealMetricSpace.triang_ineq pos_round' (fst (round rbg_ila da conf (Good g))) - (fst (conf (Good g)))). - assert ((dist pos_round' (fst (conf (Good g))) <= - D + - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g))))%R). - rewrite dist_sym in Hdist_exec. - simpl ( (fst (pos_round', (ident_round', light_round', alive_round')))) in *. - transitivity (dist pos_round' (fst (round rbg_ila da conf (Good g))) + - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g))))%R. - lra. - apply (Rplus_le_compat_r). - apply Hdist_exec. - assert (dist pos_round' (fst (conf (Good g))) <= - D + D)%R. - transitivity (D + dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g))))%R. - lra. - apply (Rplus_le_compat_l). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite dist_sym; apply Hdist_round. - lra. - destruct b; simpl; auto. - assert (Hfalse := light_on_means_not_moving_before g' Hpred Hpath Halive_exec_backup H0). - revert Hbool; intro Hbool. - unfold upt_robot_dies_b in *. - rewrite orb_false_iff in Hbool. - destruct Hbool as (Hnot_near,_). - rewrite <- negb_true_iff, forallb_existsb, forallb_forall in Hnot_near. - destruct (change_frame da conf g') as ((r,t),b) eqn : Hchange'. - - specialize (Hnot_near ( - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))). - rewrite negb_true_iff, Rle_bool_false_iff in Hnot_near. - destruct Hnot_near. - + rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - * exists (Good g). - split; destruct b; auto; try apply In_names. - * rewrite <- 2 ident_preserved in Hident_exec; try apply Hpred. - unfold get_ident; simpl. - now rewrite Nat.ltb_lt. - * unfold get_alive in *; simpl in *; auto. - * rewrite negb_true_iff. - destruct (get_based (conf (Good g))) eqn : Hbased. - destruct (Hbased_higher) as (_, (_, Hhi)). - specialize (Hhi g g' Hbased). - rewrite <- 2 ident_preserved in Hident_exec; try auto. - rewrite Hconf' in *; unfold get_based in Hhi; simpl in Hhi; specialize (Hhi (reflexivity _)). - lia. - unfold get_based in *; simpl in *; auto. - + - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - assert (Hframe := frame_dist (fst (conf (Good g))) (fst (conf (Good g'))) ((r,t),b)). - unfold frame_choice_bijection, Frame in Hframe. - rewrite <- Hfalse, <- Hdist in Hdist_exec. - rewrite Hframe in Hdist_exec. - simpl in *; destruct b; simpl in*; lra. -Qed. - - - - - -Lemma moving_means_alive_next : -forall conf g da, - da_predicat da -> - path_conf conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - get_alive (conf (Good g)) = true -> - get_location (conf (Good g)) <> (get_location (round rbg_ila da conf (Good g))) -> - get_alive (round rbg_ila da conf (Good g)) = true. -Proof. - intros conf g da Hpred Hpath Hcol Hexecutioner_light Hexecuted_light Hexists Halive Hmoving. - rewrite round_good_g in *; try apply Hpred. - destruct (change_frame da conf g) as ((r, t), b) eqn : Hchange_frame. - unfold upt_aux, map_config in *. - unfold rbg_ila in *. - unfold get_alive in *. - unfold update, UpdFun, upt_aux in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - destruct (based) eqn : Hbased. - * simpl. - destruct (upt_robot_too_far _) eqn : Htoo_far. - now simpl. - now simpl in *. - * destruct (upt_robot_dies_b _) eqn : Hbool. - - destruct Hmoving. - unfold map_config. - unfold lift, State_ILA, OnlyLocation, AddInfo, lift. - simpl (fst _). - unfold projT1. - assert (Hbij := (retraction_section (frame_choice_bijection (r, t, b))) pos) . - simpl in *. - now rewrite Hbij. - - simpl in *. - auto. -Qed. - - - - -Lemma exists_at_less_round : forall conf da, - da_predicat da -> - path_conf conf -> - based_higher conf -> - no_collision_conf conf -> - executioner_means_light_off conf -> - executed_means_light_on conf -> - exists_at_less_that_Dp conf -> - exists_at_less_that_Dp (round rbg_ila da conf). -Proof. - intros conf da Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hexists. - assert (Hexecuted_round := executed_means_light_on_round Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hexists). - assert (Hexecutioner_round := executioner_means_light_off_round Hpred Hpath Hbased_higher Hcol Hexecutioner Hexecuted Hexists). - - intros g Halive H_not_0 Hlighted. - assert (Hexists_backup := Hexists). - specialize (Hexists g (still_alive_means_alive g Hpred Hbased_higher Halive)). - assert (Hlight_on_means := light_on_means_not_moving_before). - assert (Hpath_backup := Hpath). - specialize (Hpath g (still_alive_means_alive g Hpred Hbased_higher Halive)). - destruct Hpath as [H_0| H_exists_g']. - + rewrite <- ident_preserved, H_0 in H_not_0; try auto. - lia. - + destruct H_exists_g' as (g', (Halive_g', (Hdist_g', (Hident_g', Hbased_g')))). - assert (Halive' := Halive). - rewrite round_good_g in Halive; try apply Hpred. - unfold get_alive in Halive; simpl in Halive. - unfold upt_aux, map_config in Halive. - destruct (conf (Good g)) as (pos, (((ide, lig), ali), bas)) eqn : Hconf. - destruct bas eqn : Hbased. - - simpl in Halive. - unfold based_higher in *. - destruct (Hbased_higher) as (Hap0,_). - specialize (Hap0 g). - rewrite Hconf in Hap0; unfold get_based in Hap0; simpl in Hap0. - specialize (Hap0 (reflexivity _)). - assert (get_location (round rbg_ila da conf (Good g)) = (0,0)%R). - { - rewrite round_good_g; try auto. - simpl; unfold upt_aux, map_config. - rewrite Hconf; simpl. - destruct (upt_robot_too_far _) eqn : Htoo_far; - simpl; - rewrite retraction_section; - intuition. - } - assert (Haux : exists g, get_based (conf (Good g)) = true). - exists g. - now unfold get_based in *; rewrite Hconf in *; simpl. - assert (Hbased_higher_round := based_higher_round Hpred Hpath_backup Hbased_higher Hexecutioner Hexecuted Haux). - clear Haux. - - destruct (get_based (round rbg_ila da conf (Good g))) eqn : Hbased_r. - destruct Hbased_higher_round as (Hap0_r, ((g_ex, (Hbased_ex_r, (Halive_ex_r, Hdist_ex_r))), Hhi_r)). - exists g_ex. - repeat split; auto. - specialize (Hap0_r g Hbased_r). - destruct Hap0_r as (Ha_r, Hp_r). - rewrite <- Hp_r in Hdist_ex_r. - unfold Dp in *; generalize D_0, Dmax_7D. - transitivity (Dmax - D)%R; try lra. - rewrite dist_sym in Hdist_ex_r. - lra. - destruct (Hbased_higher) as (_, ((g_ex, (Hbased_ex, (Halive_ex, Hdist_ex))), Hhigher)). - (* tous les robots sont allumés, donc aucun ne bouge, donc g_ex ne s'élimine pas, donc GG *) - destruct (get_alive (round rbg_ila da conf (Good g_ex))) eqn : Halive_ex_r. - * exists g_ex. - repeat split; auto. - rewrite <- 2 ident_preserved; try auto. - assert (Hbased_aux : get_based (conf (Good g)) = true) - by now unfold get_based; rewrite Hconf; simpl. - specialize (Hhigher g g_ex Hbased_aux Hbased_ex). - auto. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (0,0)%R (get_location (conf (Good g_ex)))). - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_ex)))). - transitivity (dist (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) + - dist (get_location (conf (Good g_ex))) (get_location (round rbg_ila da conf (Good g_ex))))%R; try auto. - transitivity (dist (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_ex))) + D)%R. - apply Rplus_le_compat_l. - assert (Hround := dist_round_max_d g_ex Hpred Hpath_backup Halive_ex). - unfold equiv, bool_Setoid, eq_setoid in Hround. - rewrite Rle_bool_true_iff in Hround. - lra. - rewrite H, dist_sym. - lra. - * - assert (get_ident (conf (Good g_ex)) <> 0). - intro Hfalse_aux. - rewrite (ident_preserved conf g_ex Hpred) in Hfalse_aux. - assert (Hfalse := ident_0_alive (round rbg_ila da conf) g_ex Hfalse_aux). - now rewrite Hfalse in *. - destruct (Hpath_backup g_ex Halive_ex) as [Hfalse|Hpath']. - contradiction. - assert (exists g0 : G, - get_alive (conf (Good g0)) = true /\ - get_based (conf (Good g0)) = false /\ - get_ident (conf (Good g0)) < get_ident (conf (Good g_ex)) /\ - (dist (get_location (conf (Good g0))) (get_location (conf (Good g_ex))) <= Dmax)%R). - destruct Hpath' as (gp, Hp); exists gp; intuition. - rewrite dist_sym; auto. - assert (Htest := executed_means_alive_near g_ex Hpred Halive_ex H0 Halive_ex_r Hpath_backup H1). - clear H1. - destruct (Htest) as (g_test, (Hident_test, (Hdist_test, (Hbased_test, Halive_test)))). - assert (get_light (conf (Good g_test)) = false). - apply (Hexecutioner g_test da Hpred); try auto. - exists g_ex. - repeat split; auto. - now rewrite dist_sym. - (* g_test est éteint. et g_test est a moins de DP de g, donc a moins de Dmax a round. comme il est a moins de Dmax a round, il est éteint a round et donc n'a aps bouger.*) - assert (dist (get_location (conf (Good g_test))) (pos) <= Dp)%R. - replace Dp with (D + (Dp-D))%R by lra. - assert (Htri := RealMetricSpace.triang_ineq (get_location (conf (Good g_test))) (get_location (conf (Good g_ex))) pos). - transitivity (dist (get_location (conf (Good g_test))) (get_location (conf (Good g_ex))) + - dist (get_location (conf (Good g_ex))) pos)%R; try auto. - transitivity (dist (get_location (conf (Good g_test))) (get_location (conf (Good g_ex))) + (Dp - D))%R. - apply Rplus_le_compat_l. - replace pos with (0,0)%R; try auto. - now destruct Hap0. - apply Rplus_le_compat_r. - now rewrite dist_sym. - assert (get_light (round rbg_ila da conf (Good g_test)) = true). - apply Hlighted. - apply light_off_means_alive_next; try auto. - rewrite H. - replace Dmax with (D + Dp)%R by now unfold Dp; lra. - transitivity (dist (get_location (round rbg_ila da conf (Good g_test))) (get_location (conf (Good g_test))) + (dist (get_location (conf (Good g_test))) (0,0)))%R. - apply RealMetricSpace.triang_ineq. - transitivity (D + (dist (get_location (conf (Good g_test))) (0,0)))%R. - apply Rplus_le_compat_r. - assert (Htrue := dist_round_max_d g_test Hpred Hpath_backup Halive_test). - unfold equiv, bool_Setoid, eq_setoid in Htrue; rewrite Rle_bool_true_iff in *; auto. - now rewrite dist_sym. - apply Rplus_le_compat_l. - replace (0,0)%R with pos by intuition. - auto. - rewrite <- 2 ident_preserved; auto. - transitivity (get_ident (conf (Good g_ex))); try auto. - apply Hhigher. - rewrite Hconf; unfold get_based; auto. - auto. - exists g_test. - repeat split. - apply light_off_means_alive_next; try auto. - rewrite <- 2 ident_preserved; auto. - transitivity (get_ident (conf (Good g_ex))); try auto. - apply Hhigher. - rewrite Hconf; unfold get_based; auto. - auto. - rewrite H, dist_sym. - rewrite <- (Hlight_on_means conf g_test da Hpred Hpath_backup (light_off_means_alive_next g_test Hpred Hpath_backup Halive_test H1 Hexecuted) H3). - replace (0,0)%R with pos by intuition. - auto. - - destruct (upt_robot_dies_b _) eqn : Hbool; simpl in Halive; - try discriminate. - - - - assert (Htarget_in := @choose_target_in_obs (obs_from_config - (fun id : ident => - ((let (p0, b) := change_frame da conf g in - let (r, t) := p0 in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((let (p0, b) := change_frame da conf g in - let (r, t) := p0 in - comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))) _ (reflexivity _)). - unfold obs_from_config at 2, Obs_ILA at 2 in Htarget_in; - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff in Htarget_in; - try (now intros x y Hxy; - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy)). - destruct Htarget_in as (([g_target | byz], Htarget_spec), (((Hdist_target, Halive_target), Halive_conf), Hident_target)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - assert (Hbased_target : get_based (conf (Good g_target)) = false). - { - set (target := choose_target _) in Htarget_spec. - destruct (get_based (target)) eqn : Hbased_target. - rewrite Htarget_spec in Hbased_target. - unfold get_based in *; simpl in Hbased_target. - destruct (Hbased_higher) as (Hap0, (Hex, Hhi)). - specialize (Hhi g_target g). - unfold get_based in *; simpl in Hhi; rewrite Hconf in *; - specialize (Hhi Hbased_target (reflexivity _)). - rewrite <- Hconf in Hident_target. - fold target in Hident_target. - rewrite (get_ident_compat Htarget_spec), Nat.leb_le in Hident_target. - unfold get_ident in *; rewrite Hconf in *; simpl in *. - lia. - rewrite Htarget_spec in Hbased_target; unfold get_based in *; simpl in *; auto. - } - - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - destruct (get_light ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_target))), snd (conf (Good g_target)))) eqn : Hlight_target. - * unfold executed_means_light_on in *. - destruct (conf (Good g_target)) as (p_target, (((i_target,l_target), a_target), b_target)) eqn : Hconf_target. - assert (get_alive (round rbg_ila da conf (Good g_target)) = true). - { rewrite round_good_g; try apply Hpred. - unfold get_alive; simpl. - unfold upt_aux, map_config. - unfold get_based in Hbased_target; simpl in Hbased_target; rewrite Hbased_target in *. - destruct (upt_robot_dies_b _ g_target) eqn : Hbool_target. - simpl. - assert (Halive_target_r : get_alive (round rbg_ila da conf (Good g_target)) = false). - rewrite round_good_g; try apply Hpred. - simpl; unfold upt_aux, map_config, get_alive; rewrite Hbool_target, Hconf_target. - now simpl. - apply robot_dead_means in Halive_target_r. - destruct Halive_target_r as [Hfalse|[(g_near,(Hident_near, (Halive_near, Hbool_near))) | Hfar]]. - -- rewrite Htarget_spec in Halive_target. - unfold get_alive in *; rewrite Hconf_target in *. - simpl in *. - now rewrite Hfalse in *; discriminate. - -- specialize (Hexecutioner g_near da Hpred Halive_near). - revert Hexecutioner; intros Hexecutioner. - assert ((exists g' : G, - get_alive (conf (Good g')) = true /\ - get_based (conf (Good g')) = false /\ - get_ident (conf (Good g_near)) < get_ident (conf (Good g')) /\ - (dist (get_location (conf (Good g_near))) - (get_location (conf (Good g'))) <= D)%R)). - exists g_target. - repeat split. - rewrite Htarget_spec in Halive_target; unfold get_alive in *; simpl in *. - now rewrite Hconf_target in *; simpl in *. - unfold get_based in *; rewrite Hconf_target; simpl in *; auto. - assumption. - rewrite Rle_bool_true_iff in *. - now rewrite dist_sym. - specialize (Hexecutioner H); clear H. - assert ((forall g_near : G, - get_alive (conf (Good g_near)) = true -> - (dist (get_location (conf (Good g_near))) - (get_location (pos, (ide, lig, ali, bas))) <= Dmax)%R -> - get_ident (conf (Good g_near)) < get_ident (pos, (ide, lig, ali, bas)) -> - get_light (conf (Good g_near)) = true)). - intros g0 Halive_0 Hdist_0 Hident_0. - assert (Hoff_first := (@light_off_first (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g))))) _ (reflexivity _ )). - rewrite Htarget_spec in Hoff_first. - specialize (Hoff_first Hlight_target). - unfold For_all in *. - specialize (Hoff_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g0))), snd (conf (Good g0)))). - unfold get_light in *. - simpl (snd (fst _)) in *. - apply Hoff_first. - unfold obs_from_config, Obs_ILA; - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g0); destruct b; reflexivity. - rewrite Hconf. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_0. - unfold frame_choice_bijection, Frame in *; fold Frame in *; - destruct b; now simpl in *. - unfold get_alive in *; now simpl in *. - unfold get_alive in *; now simpl in *. - rewrite Hconf, Nat.leb_le. - destruct b; unfold get_ident in *; simpl in *; lia. - - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy); reflexivity. - assert (H_0 : get_ident (pos, (ide, lig, ali, bas)) > 0). - { - rewrite <- ident_preserved, Hconf in *; auto. - } - specialize (Hexists H_0 H). - clear H. - rewrite Hconf_target; simpl. - destruct Hexists as (g_exi, (Halive_exi, (Hicent_exi, Hdist_exi))). - destruct (Rle_lt_dec (dist p_target pos) Dp). - assert (Hexec_near : (dist (get_location (conf (Good g_near))) pos <= Dmax)%R). - { - rewrite Rle_bool_true_iff, Hconf_target, dist_sym in Hbool_near. - generalize Dmax_7D, D_0. - unfold Dp in *. - assert (Htri := RealMetricSpace.triang_ineq (get_location (conf (Good g_near))) p_target pos). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - simpl (fst _) in Hbool_near. - simpl in *; lra. - } - assert (Hoff_first := @light_off_first (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g)))) _ (reflexivity _)). - rewrite Htarget_spec in Hoff_first. - unfold For_all, get_light in *; simpl (snd (fst _)) in *. - specialize (Hoff_first Hlight_target (((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near))))). - rewrite Hoff_first in Hexecutioner; try discriminate. - { unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g_near); destruct b; reflexivity. - rewrite Hconf. - - rewrite (frame_dist _ _ (r,t,b)) in Hexec_near. - unfold frame_choice_bijection, Frame in *; fold Frame in *; - destruct b; now simpl in *. - unfold get_alive in *; now simpl in *. - unfold get_alive in *; rewrite Hconf in *; simpl in *; assumption. - rewrite Hconf, Nat.leb_le. - assert (Htarget_lower := target_lower). - specialize (Htarget_lower (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g)))) - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target))), - snd (p_target, (i_target, l_target, a_target, b_target))) ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - rewrite Hconf in Htarget_lower. - assert (get_ident (conf (Good g_target)) < get_ident (conf (Good g))). - unfold get_ident in *; simpl (fst _) in *; rewrite Hconf, Hconf_target; apply Htarget_lower. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff. - repeat split. - exists (Good g). - now rewrite Hconf. - unfold Datatypes.id. - generalize (dist_same ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) pos)), Dmax_7D, D_0; - rewrite Rle_bool_true_iff. - intro Hdist_0. - simpl in Hdist_0; simpl. - rewrite Hdist_0. - lra. - simpl in *; assumption. - simpl in *; assumption. - rewrite Nat.leb_le. unfold get_ident; simpl; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Horigin, ?). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - unfold frame_choice_bijection, Frame, Datatypes.id in *. - rewrite <- Horigin. - rewrite Hchange. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hconf; simpl. destruct b; simpl; reflexivity. - rewrite Hbased_target, <- Htarget_spec, Hconf. - reflexivity. - transitivity (get_ident (conf (Good g_target))). - unfold get_ident in *; simpl in *; lia. - unfold get_ident in *; rewrite Hconf in *; simpl in *; lia. - intros x y Hxy; rewrite (RelationPairs.fst_compat _ _ _ _ Hxy), (get_alive_compat Hxy), (get_ident_compat Hxy); reflexivity. - } - assert (Hclose_first := @light_close_first (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g)))) _ (reflexivity _)). - rewrite Htarget_spec in Hclose_first. - unfold For_all, get_light in *; simpl (snd (fst _)) in *. - specialize (Hclose_first Hlight_target). - assert ((dist (0, 0) - (get_location - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target))), - snd (p_target, (i_target, l_target, a_target, false)))) > Dp)%R). - { clear Hclose_first. - destruct Hpred as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite <- Horigin. - clear Horigin. - rewrite (frame_dist _ _ (r,t,b)) in r0. - rewrite dist_sym. - rewrite Hchange. - unfold frame_choice_bijection, Frame in *; - destruct b; simpl in *; lra. - } - specialize (Hclose_first H ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_exi))), snd (conf (Good g_exi)))); - clear H. - destruct Hclose_first as [Hclose_first| Hclose_first]. - - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g_exi); destruct b; reflexivity. - rewrite Hconf. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_exi. - unfold frame_choice_bijection, Frame in *; fold Frame in *; - destruct b; unfold Dp in *; generalize Dmax_7D, D_0; - revert Hdist_exi; intros Hdist_exi; - rewrite dist_sym; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *; - simpl in *; lra. - unfold get_alive in *; now simpl in *. - unfold get_alive in *; rewrite Hconf in *; simpl in *; assumption. - rewrite Hconf, Nat.leb_le. - unfold get_ident in *; simpl in *; - lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - - apply Rgt_not_le in Hclose_first. - destruct Hclose_first. - - destruct Hpred as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite <- Horigin. - clear Horigin. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_exi. - revert Hdist_exi; intros Hdist_exi. - rewrite Hchange. - unfold frame_choice_bijection, Frame in *; - destruct b; simpl in *; lra. - - assert (g_exi <> g). - { - destruct (Geq_dec g_exi g). - rewrite e in Hicent_exi. - rewrite Hconf in Hicent_exi. - unfold get_ident in Hicent_exi; simpl in Hicent_exi; auto. - lia. - auto. - } - specialize (Hcol g_exi g H). - rewrite Hconf in Hcol. - unfold get_based in Hcol. - simpl (snd _) in Hcol. - assert (get_based (conf (Good g_exi)) = false). - { - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (_,(_,Hbh)). - destruct (get_based (conf (Good g_exi))) eqn : Hfalse_based. - specialize (Hbh g_exi g Hfalse_based). - rewrite Hconf in Hbh. - unfold get_based in Hbh; simpl (snd _) in Hbh; specialize (Hbh (reflexivity _)). - lia. - reflexivity. - } - unfold get_based in H0; specialize (Hcol H0 (reflexivity _)). - revert Hcol; intro Hcol. - specialize (Hcol Halive_exi); unfold get_alive in *; simpl (snd _) in Hcol; specialize (Hcol Halive). - destruct Hpred as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite <- Horigin in Hclose_first. - rewrite (frame_dist _ _ (r,t,b)) in Hcol. - - rewrite Hchange in Hclose_first; rewrite <- Hclose_first in Hcol. - destruct Hcol. - simpl (frame_choice_bijection _); simpl (get_location _). - unfold Datatypes.id. - assert (Htrue := dist_same (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good g_exi))))). - rewrite <- Htrue. - destruct b; auto. - - -- specialize (Hpath_backup g_target). - revert Hpath_backup; intro Hpath_backup. - rewrite <- Halive_target, Htarget_spec in Hpath_backup. - unfold get_alive at 1 2 in Hpath_backup. - rewrite Hconf_target in Hpath_backup; - simpl (snd (snd _)) in Hpath_backup. - specialize (Hpath_backup (reflexivity _)). - destruct Hpath_backup as [Hfalse| (g_near,(Halive_near, (Hdist_near, (Hident_near, Hbased_near))))]. - assert (Halive_target_round : get_alive (round rbg_ila da conf (Good g_target)) = false). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config; rewrite Hbool_target. - unfold get_alive; simpl. - now rewrite Hconf_target; simpl. - } - rewrite (ident_0_alive (round rbg_ila da conf) g_target) in Halive_target_round. - discriminate. - rewrite <- ident_preserved; try auto. - now rewrite Hconf_target. - rewrite Halive_target in Halive_near. - rewrite Hconf_target in *. - specialize (Hfar g_near Hident_near Halive_near). - rewrite negb_true_iff, Rle_bool_false_iff in Hfar. - destruct Hfar. - apply Hbased_near. - lra. - -- apply Hpred. - -- rewrite Hconf_target. - simpl. - rewrite Htarget_spec in Halive_target. - unfold get_alive in Halive_target; now simpl in *. - } - - exists g_target. - - repeat split. - -- assumption. - -- rewrite <- 2 ident_preserved. - assert (Hlower_target := target_lower). - specialize (@Hlower_target (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g))))). - assert (In (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))) - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g))))). - { unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA, 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - exists (Good g); destruct b; reflexivity. - rewrite Hconf. - simpl in *. - destruct b; replace (sqrt _) with (sqrt 0) by (f_equal; lra); - rewrite sqrt_0; generalize Dmax_7D D_0; lra. - unfold get_alive in *; rewrite Hconf in *; now simpl in *. - unfold get_alive in *; rewrite Hconf in *; simpl in *; assumption. - rewrite Hconf, Nat.leb_le. unfold get_ident; simpl; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - } - specialize (Hlower_target (choose_target - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g))))) _ H0); clear H0. - rewrite (get_ident_compat Htarget_spec) in Hlower_target. - unfold get_ident in *; rewrite Hconf_target; - simpl (snd (fst _)) in *; apply Hlower_target. - destruct Hpred as (Horigin, _). - specialize (Horigin conf g (change_frame da conf g) (reflexivity _)). - revert Horigin; intros Horigin. - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite <- Horigin. - rewrite Hchange. - unfold frame_choice_bijection, Frame; fold Frame; destruct b; - - reflexivity. - reflexivity. - apply Hpred. - apply Hpred. - -- destruct (round rbg_ila da conf (Good g)) as (pos_round, (((ident_round, light_round), alive_round), based_round)) eqn : Hconf_round. - assert (Hround_equiv : round rbg_ila da conf (Good g) == (pos_round, (((ident_round, light_round), alive_round), based_round))) by now rewrite Hconf_round. - rewrite round_good_g in Hround_equiv; try apply Hpred. - simpl in Hround_equiv. - destruct Hround_equiv as (Hpos_round, Hsnd_round). - unfold upt_aux, map_config in *. rewrite Hchange, Hbool in *. - rewrite Hconf in *. - unfold rbg_fnc in *. - destruct (get_light ((round rbg_ila da conf (Good g_target)))) eqn : Hlight_target_round. - ++ - - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_dp, Hdist_d). - destruct (move_to _) eqn : Hmove. - ** simpl in Hsnd_round. - assert (retraction - (comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - new_pos == - pos_round). - rewrite <- Hpos_round. - destruct b; now simpl in *. - - specialize (Hlighted g_target H). - revert Hlighted; intro Hlighted. - assert (Hsome := move_to_Some_zone Hmove). - rewrite <- Inversion in H0. - - assert ((dist (get_location (round rbg_ila da conf (Good g_target))) - (get_location (pos_round, (ident_round, light_round, alive_round, based_round)))) <= Dmax)%R. - assert (Hdist_round_target := dist_round_max_d g_target Hpred Hpath_backup). - rewrite <- Halive_target in Hdist_round_target at 1. - rewrite Htarget_spec, Hconf_target in Hdist_round_target. - unfold get_alive in Hdist_round_target. - simpl (snd (snd _)) in *. - specialize (Hdist_round_target (reflexivity _)). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round_target. - rewrite Rle_bool_true_iff in Hdist_round_target. - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite <- H0 in Hdist_dp. - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_target))) p_target pos_round). - assert (dist p_target pos_round <= Dp)%R. - { - rewrite (@frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame. - rewrite dist_sym. - destruct b; rewrite (RelationPairs.fst_compat _ _ _ _ Htarget_spec) in Hdist_dp; simpl in *; lra. } - assert(dist (fst (round rbg_ila da conf (Good g_target))) p_target <= D)%R. - { - rewrite dist_sym. - simpl in *; apply Hdist_round_target. - } - transitivity (dist (fst (round rbg_ila da conf (Good g_target))) p_target + - dist p_target pos_round)%R. - simpl in *; apply Htri. - transitivity (dist (fst (round rbg_ila da conf (Good g_target))) p_target + Dp)%R. - lra. - transitivity (D + Dp)%R. - lra. - unfold Dp. - lra. - rewrite <- H0 in Hdist_dp. - assert (Hlight_target_simp : l_target = true) by now simpl in *. - revert Hlighted; intro Hlighted. - specialize (Hlighted H1). - assert (Htarget_lower := target_lower). - assert (get_light (round rbg_ila da conf (Good g_target)) = true). - { - apply Hlighted. - specialize (Htarget_lower (@obs_from_config _ _ _ _ Obs_ILA - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide, lig, ali,bas)))) - ( - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_target))), snd (conf (Good g_target)))) ( - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))). - rewrite <- Hconf_round. - rewrite <- 2 ident_preserved. - rewrite Hconf_target, Hconf in *. - unfold get_ident in *. - simpl (fst (fst( snd _))) in *. - apply Htarget_lower. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g). - rewrite Hconf. - destruct b; simpl; - try now repeat split. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite dist_same. - generalize D_0, Dmax_7D; lra. - now simpl. - unfold get_alive in *; simpl in *; assumption. - rewrite Nat.leb_le. unfold get_ident; simpl; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Hpred1, Hpred3). - unfold change_frame_origin in Hpred1. - rewrite <- (Hpred1 conf g (r,t,b)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id. - destruct b; simpl; - rewrite Hconf; - now simpl. - auto. - revert Htarget_spec; intro Htarget_spec. - rewrite <- Htarget_spec. - f_equiv. - apply Hpred. - apply Hpred. - } - assert (get_location (round rbg_ila da conf (Good g_target)) == get_location (conf (Good g_target))). - { - symmetry. - apply Hlight_on_means. - apply Hpred. - apply Hpath_backup. - apply H. - apply H2. - } - rewrite H3. - assert ((dist - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) pos_round) - (fst - (choose_target - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide,lig,ali,bas)))))) <= Dp)%R <-> - (dist - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) pos_round) - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target))), - snd (p_target, (i_target, l_target, a_target)))) <= Dp)%R). - f_equiv. - revert Htarget_spec; intros Htarget_spec. - rewrite (dist_compat _ _ (reflexivity _) _ _ (fst_compat Htarget_spec)). - f_equiv. - assert ((dist - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - pos_round) - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target, b_target))), - snd (p_target, (i_target, l_target, a_target,b_target)))) <= Dp)%R). - { destruct H4. - unfold Loc, location, make_Location in *; - unfold reflexion, Similarity.id in *; destruct b; simpl in *; - specialize (H4 Hdist_dp); - apply H4. - } - clear H4. - rewrite Hconf_target. - rewrite (frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame. - destruct b; simpl in *; auto; - unfold Datatypes.id; - auto. - - ** rewrite <- Hconf_round in Halive'. - simpl in Hsnd_round. - assert (Hlight_round : get_light (round rbg_ila da conf (Good g)) = true). - { - rewrite Hconf_round; unfold get_light; simpl. - intuition. - } - assert (Hlight_on_round := Hlight_on_means conf g da Hpred Hpath_backup Halive' Hlight_round). - assert (Hlight_on_round_target := Hlight_on_means conf g_target da Hpred Hpath_backup H Hlight_target_round). - assert (Hnear_exists : (forall g_near : G, - get_alive (conf (Good g_near)) = true -> - (dist (get_location (conf (Good g_near))) - (get_location (pos, (ide, lig, ali,bas))) <= Dmax)%R -> - get_ident (conf (Good g_near)) < get_ident (pos, (ide, lig, ali,bas)) -> - get_light (conf (Good g_near)) = true)). - intros g_near Halive_near Hdist_near Hident_near. - assert (Hlight_off_first := light_off_first Htarget_spec Hlight_target). - specialize (Hlight_off_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g_near))), snd (conf (Good g_near)))). - assert ((((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near))) - ∈ obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali,bas)))), snd (pos, (ide,lig,ali,bas))))%set). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g_near). - reflexivity. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_near. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in Hdist_near. destruct b; simpl in *; - apply Hdist_near. - unfold get_alive in *; simpl in *; assumption. - unfold get_alive in *; rewrite Hconf in *; simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - specialize (Hlight_off_first H0); clear H0. - simpl in Hlight_off_first. - now unfold get_light in *; simpl in *. - rewrite <- Hconf_round, <- ident_preserved, Hconf in H_not_0; auto. - specialize (Hexists H_not_0 Hnear_exists). - revert Hexists; intros Hexists. - destruct (Rgt_ge_dec (dist (get_location (pos_round, (ident_round, light_round, alive_round, based_round))) - (get_location (round rbg_ila da conf (Good g_target)))) Dp). - assert (Hlight_close := light_close_first Htarget_spec Hlight_target). - exfalso. - destruct Hexists as (g_false, (Halive_false, (Hident_false, Hdist_false))). - assert (Hdist_dp_target : (dist (0, 0) - (get_location - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target, b_target))), - snd (p_target, (i_target, l_target, a_target,b_target)))) > Dp)%R). - { - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location. - assert (Hinv := Inversion (frame_choice_bijection (r,t,b))). - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin conf g (r,t,b) Hchange). - rewrite <- Horigin. - assert ( - (Datatypes.id - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target))), - snd (p_target, (i_target, l_target, a_target))))) == - (frame_choice_bijection (r,t,b)) (get_location (conf (Good g_target))))%R. - unfold frame_choice_bijection, Frame. destruct b; simpl; - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location; - rewrite Hconf_target; - now simpl. - destruct b; rewrite (dist_compat _ _ (reflexivity _) _ _ H0); - rewrite <- frame_dist; - rewrite Hlight_on_round, Hlight_on_round_target; - now rewrite Hconf_round. - } - specialize (Hlight_close Hdist_dp_target). - specialize (Hlight_close ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g_false))), snd (conf (Good g_false)))). - assert ((((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_false))), snd (conf (Good g_false))) - ∈ obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali,bas)))), snd (pos, (ide,lig,ali,bas))))%set). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g_false). - reflexivity. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_false. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in Hdist_false. - rewrite dist_sym. - destruct b; simpl in *; unfold Dp in *; - generalize D_0, Dmax_7D; lra. - unfold get_alive in *; simpl in *; assumption. - unfold get_alive in *; rewrite Hconf in *; simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - specialize (Hlight_close H0); clear H0. - absurd (dist (get_location (pos, (ide, lig, ali, bas))) - (get_location (conf (Good g_false))) <= Dp)%R . - apply Rgt_not_le. - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin conf g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hlight_close. - - assert ( - (Datatypes.id - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_false))), - snd (conf (Good g_false))))) == - (frame_choice_bijection (r,t,b)) (get_location (conf (Good g_false))))%R. - unfold frame_choice_bijection, Frame. destruct b; simpl; - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location; - now simpl. - destruct Hlight_close as [Hlight_close|Hlight_close]. - destruct b; rewrite (dist_compat _ _ (reflexivity _) _ _ H0), <- frame_dist in Hlight_close; - assumption. - - - assert (g_false <> g). - { - destruct (Geq_dec g_false g). - rewrite e in Hident_false. - rewrite Hconf in Hident_false. - unfold get_ident in Hident_false; simpl in Hident_false; auto. - lia. - auto. - } - specialize (Hcol g_false g H1). - rewrite Hconf in Hcol. - unfold get_based in Hcol. - simpl (snd _) in Hcol. - assert (get_based (conf (Good g_false)) = false). - { - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (_,(_,Hbh)). - destruct (get_based (conf (Good g_false))) eqn : Hfalse_based. - specialize (Hbh g_false g Hfalse_based). - rewrite Hconf in Hbh. - unfold get_based in Hbh; simpl (snd _) in Hbh; specialize (Hbh (reflexivity _)). - lia. - reflexivity. - } - unfold get_based in H2; specialize (Hcol H2 (reflexivity _)). - revert Hcol; intro Hcol. - specialize (Hcol Halive_false); unfold get_alive in *; simpl (snd _) in Hcol; specialize (Hcol Halive). - rewrite Hconf in *; - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite (frame_dist _ _ (r,t,b)) in Hcol. - rewrite <- Hlight_close in Hcol. - destruct Hcol. - simpl (frame_choice_bijection _); simpl (get_location _). - unfold Datatypes.id. - assert (Htrue := dist_same (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good g_false))))). - rewrite <- Htrue. - destruct b; auto. - - apply Hdist_false. - apply Rge_le. - apply r0. - ++ specialize (Hlighted g_target). - absurd (get_light (round rbg_ila da conf (Good g_target)) = true). - intro. - now rewrite H0 in *. - apply Hlighted; try apply H; try (now unfold get_ident in *; simpl in *; assumption). - ** - destruct (move_to _) eqn : Hmove. - --- assert (Hmove_some := move_to_Some_zone Hmove). - simpl in Hsnd_round. - - assert (dist (get_location (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g_target))) <= D)%R. - { - rewrite <- Rle_bool_true_iff. - rewrite dist_round_max_d; auto. - rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_alive in *; rewrite Hconf_target; - simpl in *; auto. - } - - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := @choose_new_pos_correct _ _ new_pos (reflexivity _)). - destruct Hchoose_correct as (Hdist_choose_dp, Hdist_chose_d). - assert (Hdist_dp: (dist pos_round p_target <= Dp)%R). - { - assert ( Hdist_t_dp: (dist new_pos (fst (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target,b_target))), - snd (p_target, (i_target, l_target, a_target, b_target))))) = (dist new_pos - (fst - (choose_target - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide, lig, ali,bas))))))))%R). - f_equiv. - now rewrite (fst_compat Htarget_spec). - assert ((dist new_pos (fst (((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target,b_target))), - snd (p_target, (i_target, l_target, a_target, b_target))))) <= Dp)%R). - destruct b; rewrite Hbased in *; rewrite Hdist_t_dp. apply (Hdist_choose_dp). - apply Hdist_choose_dp. - rewrite <- Hpos_round. - rewrite (frame_dist _ _ (r,t,b)). - unfold frame_choice_bijection, Frame. - destruct b; rewrite section_retraction; simpl in *; lra. - } - rewrite Hconf_target in H0; - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - simpl (fst _) in *. - rewrite dist_sym. - assert (Htri := RealMetricSpace.triang_ineq pos_round p_target (fst (round rbg_ila da conf (Good g_target)))). - unfold Dp in *; generalize D_0, Dmax_7D. - intros. - transitivity (dist pos_round p_target + - dist p_target (fst (round rbg_ila da conf (Good g_target))))%R. - assumption. - transitivity (Dmax - D +dist p_target (fst (round rbg_ila da conf (Good g_target))))%R. - lra. - transitivity (Dmax - D + D)%R. - apply Rplus_le_compat_l. - assumption. - lra. - --- - assert (Hlight_close := light_close_first Htarget_spec Hlight_target). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - destruct (Rle_lt_dec (dist (0, 0)%R - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (p_target, (i_target, l_target, a_target,b_target))), - snd (p_target, (i_target, l_target, a_target,b_target))))) Dp)%R. - simpl (fst _) in Hpos_round. - assert (Hpos_round' : retraction - (comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (0%R, 0%R) == pos_round) by (now destruct b; rewrite Hpos_round). - rewrite <- Inversion in Hpos_round'. - rewrite <- Hpos_round' in r0. - assert (Hframe := frame_dist pos_round p_target (r,t,b)). - unfold frame_choice_bijection, Frame in Hframe. - assert (Htarg_d := dist_round_max_d g_target Hpred Hpath_backup). - assert (get_alive (conf (Good g_target)) == true). - rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_alive in *; rewrite Hconf_target; now simpl in *. - specialize (Htarg_d H0); clear H0. - unfold equiv, bool_Setoid, eq_setoid in Htarg_d. - rewrite Rle_bool_true_iff in Htarg_d. - rewrite dist_sym. - simpl (fst _). - assert (Htri := RealMetricSpace.triang_ineq pos_round p_target (fst (round rbg_ila da conf (Good g_target)))). - unfold Dp in *; generalize D_0, Dmax_7D. - intros. - transitivity (dist pos_round p_target + - dist p_target (fst (round rbg_ila da conf (Good g_target))))%R. - assumption. - transitivity (Dmax - D +dist p_target (fst (round rbg_ila da conf (Good g_target))))%R. - apply Rplus_le_compat. - destruct b; rewrite <- Hframe in r0; apply r0. - lra. - transitivity (Dmax - D + D)%R. - apply Rplus_le_compat_l. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - rewrite Hconf_target in *; simpl (fst _) in Htarg_d. - assumption. - lra. - apply Rlt_gt in r0. - specialize (Hlight_close r0). - assert (Hnear_exists : (forall g_near : G, - get_alive (conf (Good g_near)) = true -> - (dist (get_location (conf (Good g_near))) - (get_location (pos, (ide, lig, ali,bas))) <= Dmax)%R -> - get_ident (conf (Good g_near)) < get_ident (pos, (ide, lig, ali, bas)) -> - get_light (conf (Good g_near)) = true)). - intros g_near Halive_near Hdist_near Hident_near. - assert (Hlight_off_first := light_off_first Htarget_spec Hlight_target). - specialize (Hlight_off_first ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g_near))), snd (conf (Good g_near)))). - assert ((((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_near))), snd (conf (Good g_near))) - ∈ obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide,lig,ali,bas))))%set). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g_near). - reflexivity. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_near. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in Hdist_near. - destruct b; simpl in *; - apply Hdist_near. - unfold get_alive in *; simpl in *; assumption. - unfold get_alive in *; simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - specialize (Hlight_off_first H0); clear H0. - simpl in Hlight_off_first. - now unfold get_light in *; simpl in *. - rewrite <- Hconf_round, <- ident_preserved, Hconf in H_not_0; auto. - specialize (Hexists H_not_0 Hnear_exists). - revert Hexists; intros Hexists. - exfalso. - destruct Hexists as (g_false, (Halive_false, (Hident_false, Hdist_false))). - specialize (Hlight_close ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) (fst (conf (Good g_false))), snd (conf (Good g_false)))). - assert ((((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_false))), snd (conf (Good g_false))) - ∈ obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide, lig, ali, bas))))%set). - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g_false). - reflexivity. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite (frame_dist _ _ (r,t,b)) in Hdist_false. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in Hdist_false. - rewrite dist_sym. - destruct b; simpl in *; unfold Dp in *; - generalize D_0, Dmax_7D; lra. - unfold get_alive in *; simpl in *; assumption. - unfold get_alive in *; simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - specialize (Hlight_close H0); clear H0. - absurd (dist (get_location (pos, (ide, lig, ali, bas))) - (get_location (conf (Good g_false))) <= Dp)%R . - apply Rgt_not_le. - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin conf g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hlight_close. - - assert ( - (Datatypes.id - (fst - ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_false))), - snd (conf (Good g_false))))) == - (frame_choice_bijection (r,t,b)) (get_location (conf (Good g_false))))%R. - unfold frame_choice_bijection, Frame. destruct b; simpl; - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location; - now simpl. - destruct Hlight_close as [Hlight_close|Hlight_close]. - destruct b; - rewrite (dist_compat _ _ (reflexivity _) _ _ H0) - , <- frame_dist in Hlight_close - ; assumption. - - assert (g_false <> g). - { - destruct (Geq_dec g_false g). - rewrite e in Hident_false. - rewrite Hconf in Hident_false. - unfold get_ident in Hident_false; simpl in Hident_false; auto. - lia. - auto. - } - specialize (Hcol g_false g H1). - rewrite Hconf in Hcol. - unfold get_based in Hcol. - simpl (snd _) in Hcol. - assert (get_based (conf (Good g_false)) = false). - { - unfold based_higher in Hbased_higher. - destruct Hbased_higher as (_,(_,Hbh)). - destruct (get_based (conf (Good g_false))) eqn : Hfalse_based. - specialize (Hbh g_false g Hfalse_based). - rewrite Hconf in Hbh. - unfold get_based in Hbh; simpl (snd _) in Hbh; specialize (Hbh (reflexivity _)). - lia. - reflexivity. - } - unfold get_based in H2; specialize (Hcol H2 (reflexivity _)). - revert Hcol; intro Hcol. - specialize (Hcol Halive_false); unfold get_alive in *; simpl (snd _) in Hcol; specialize (Hcol Halive). - - unfold get_location, State_ILA, OnlyLocation, AddInfo, - get_location, Datatypes.id in *. - rewrite (frame_dist _ _ (r,t,b)) in Hcol. - rewrite <- Hlight_close in Hcol. - destruct Hcol. - simpl (frame_choice_bijection _); simpl (get_location _). - unfold Datatypes.id. - assert (Htrue := dist_same (comp (bij_rotation r) (comp (bij_translation t) (if b then reflexion else Similarity.id)) (fst (conf (Good g_false))))). - rewrite <- Htrue. - destruct b; auto. - - - apply Hdist_false. - ** - rewrite <- Hconf_round. - rewrite <- 2 ident_preserved. - - assert (Htarget_lower := @target_lower (@obs_from_config _ _ _ _ Obs_ILA - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (pos, (ide, lig, ali, bas)))), snd (pos, (ide, lig,ali, bas)))) - ( - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_target))), snd (conf (Good g_target)))) ( - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g))))). - rewrite Hconf_target, Hconf in *. - unfold get_ident in *. - simpl (fst (fst( snd _))) in *. - apply Htarget_lower. -unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g). - rewrite Hconf. - destruct b; simpl; - try now repeat split. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite dist_same. - generalize D_0, Dmax_7D; lra. - now simpl. - simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Hpred1, Hpred3). - unfold change_frame_origin in Hpred1. - rewrite <- (Hpred1 conf g (r,t,b)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id. - destruct b; simpl; - rewrite Hconf; - now simpl. - auto. - revert Htarget_spec; intro Htarget_spec. - rewrite <- Htarget_spec. - f_equiv. - apply Hpred. - apply Hpred. - * destruct ((round rbg_ila da conf (Good g))) as (pos_round, (((ident_round, light_round), alive_round), based_round)) eqn : Hconf_round. - assert (Hconf_round' : round rbg_ila da conf (Good g) == - (pos_round, (ident_round, light_round, alive_round, based_round))) by now rewrite Hconf_round. - rewrite round_good_g in Hconf_round'; try apply Hpred. - simpl in Hconf_round'. - unfold rbg_fnc in Hconf_round'. - destruct (move_to _) eqn : Hmove_to. - ** exists g_target. - repeat split. - ++ rewrite round_good_g. - destruct ((round rbg_ila da conf (Good g_target))) as - (pos_target_round, (((ident_target_round, light_target_round), alive_target_round), based_target_round)) eqn : Hconf_target_round. - assert (Hconf_target_round' : round rbg_ila da conf (Good g_target) == - (pos_target_round, - (ident_target_round, light_target_round, alive_target_round, based_target_round))) by now rewrite Hconf_target_round. - rewrite round_good_g in Hconf_target_round'. - simpl. - simpl in Hconf_target_round'. - unfold upt_aux, map_config in *. - unfold get_alive; destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. simpl. - simpl in Hconf_target_round'. - unfold get_based in Hbased_target; simpl in Hbased_target; rewrite Hbased_target in *. - - destruct (upt_robot_dies_b _ g_target) eqn : Hdies_target. - -- rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_light in Hlight_target; simpl in Hlight_target. - unfold get_alive in Halive_target; simpl in Halive_target. - simpl in Hconf_target_round'. - destruct Hconf_target_round' as (Hpos_target_round, (Hident_target_round, (Hlight_target_round, (Halive_target_round, Hbased_target_round)))). - revert Hexecuted; intro Hexecuted. - specialize (Hexecuted g_target da Hpred). - unfold get_alive in Hexecuted. - rewrite Hconf_target, Hconf_target_round in Hexecuted. - simpl in Hexecuted. - specialize (Hexecuted Halive_target (symmetry Halive_target_round)). - unfold get_light in *; simpl in Hexecuted. - rewrite Hexecuted in *; discriminate. - -- simpl. - rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_alive in *; simpl in Halive_target. - assumption. - -- apply Hpred. - -- apply Hpred. - ++ rewrite <- Hconf_round, <- 2 (ident_preserved conf _ Hpred). - assert (Htarget_lower := @target_lower - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g)))) - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_target))), snd (conf (Good g_target))) - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - unfold get_ident in *. - destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. - rewrite Hconf in *. - simpl (fst _) in Htarget_lower. - simpl. - apply Htarget_lower. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g). - rewrite Hconf. - destruct b; simpl; - try now repeat split. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite dist_same. - generalize D_0, Dmax_7D; lra. - now simpl. - unfold get_alive in *; simpl in *; assumption. - rewrite Nat.leb_le. - unfold get_ident in *; simpl in *; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Hpred1, Hpred3). - unfold change_frame_origin in Hpred1. - rewrite <- (Hpred1 conf g (r,t,b)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id. - destruct b; simpl; - rewrite Hconf; - now simpl. - auto. - revert Htarget_spec; intro Htarget_spec. - rewrite <- Htarget_spec. - f_equiv. - ++ rewrite <- Hconf_round, round_good_g; try apply Hpred. - destruct ((round rbg_ila da conf (Good g_target))) as - (pos_target_round, (((ident_target_round, light_target_round), alive_target_round), based_target_round)) eqn : Hconf_target_round. - assert (Hconf_target_round' : round rbg_ila da conf (Good g_target) == - (pos_target_round, - (ident_target_round, light_target_round, alive_target_round, based_target_round))) by now rewrite Hconf_target_round. - rewrite round_good_g in Hconf_target_round'; try apply Hpred. - simpl (get_location _). - simpl in Hconf_target_round'. - unfold upt_aux, map_config in *. - unfold get_alive; destruct (conf (Good g_target)) as (pos_target, (((ident_target, light_target), alive_target), based_target)) eqn : Hconf_target. simpl (Datatypes.id _). - simpl in Hconf_target_round'. - unfold get_based in Hbased_target; simpl in Hbased_target; rewrite Hbased_target in *. - destruct (upt_robot_dies_b _ g_target) eqn : Hdies_target. - -- rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_light in Hlight_target; simpl in Hlight_target. - unfold get_alive in Halive_target; simpl in Halive_target. - simpl in Hconf_target_round'. - destruct Hconf_target_round' as (Hpos_target_round, (Hident_target_round, (Hlight_target_round, (Halive_target_round, Hbased_target_round)))). - revert Hexecuted; intro Hexecuted. - specialize (Hexecuted g_target da Hpred). - unfold get_alive in Hexecuted. - rewrite Hconf_target, Hconf_target_round in Hexecuted. - simpl in Hexecuted. - specialize (Hexecuted Halive_target (symmetry Halive_target_round)). - unfold get_light in *; simpl in Hexecuted. - rewrite Hexecuted in *; discriminate. - -- - assert (Hmove_some := move_to_Some_zone Hmove_to). - destruct Hconf_target_round' as (Hpos_target_round, (Hident_target_round, (Hlight_target_round, (Halive_target_round, Hbased_target_round)))). - rewrite Hchange in *. - rewrite Hbool in *. - (* faire attention on a mixé round g_target et round g *) - unfold rbg_fnc. - unfold path_conf in Hpath_backup. - assert (get_light (round rbg_ila da conf (Good g_target)) = true). - { - apply (Hlighted g_target). - rewrite Hconf_target_round. - unfold get_alive; simpl. - rewrite <- Halive_target_round. - rewrite (get_alive_compat Htarget_spec) in Halive_target. - unfold get_alive in Halive_target; simpl in *; auto. - assert (Htri := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g_target))) (get_location (conf (Good g_target))) (get_location (round rbg_ila da conf (Good g)))). - rewrite <- Hconf_round. - assert (Hdist_round_target: (dist (get_location (round rbg_ila da conf (Good g_target))) (get_location (conf (Good g_target))) <= D)%R). - { - rewrite dist_sym, <- Rle_bool_true_iff; - apply dist_round_max_d; - auto. - assert (Ht_alive := choose_target_alive Htarget_spec). - rewrite Hconf_target in *; unfold get_alive in *; now simpl in *. - } - assert( Hdist_target_round_g : (dist (get_location (conf (Good g_target))) - (get_location (round rbg_ila da conf (Good g))) <= Dp)%R). - { rewrite round_good_g; auto. - simpl (get_location _). - unfold upt_aux, map_config; rewrite Hchange, Hbool. - unfold rbg_fnc. rewrite Hmove_to. - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := choose_new_pos_correct (reflexivity new_pos)). - rewrite (frame_dist _ _ (r,t,b)). - rewrite Hconf. - simpl (fst _). - unfold frame_choice_bijection, Frame, Datatypes.id. - rewrite section_retraction. - rewrite dist_sym. - destruct b; rewrite (fst_compat Htarget_spec) in Hchoose_correct; - rewrite Hconf_target in *; simpl in *; lra. - } - unfold Dp in *; generalize Dmax_7D, D_0. - lra. - rewrite <- Hconf_round, <- 2 ident_preserved. - assert (Htarget_lower := @target_lower - (obs_from_config - (fun id : ident => - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf id)), snd (conf id))) - (Datatypes.id - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g)))), snd (conf (Good g)))) - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g_target))), snd (conf (Good g_target))) - ((comp (bij_rotation r) - (comp (bij_translation t) - (if b then reflexion else Similarity.id))) - (fst (conf (Good g))), snd (conf (Good g)))). - unfold get_ident in *. - rewrite Hconf, Hconf_target in *. - simpl (fst _) in Htarget_lower. - simpl. - apply Htarget_lower. - unfold obs_from_config, Obs_ILA. - rewrite make_set_spec, filter_InA, config_list_InA. - split. - exists (Good g). - rewrite Hconf. - destruct b; simpl; - try now repeat split. - rewrite 3 andb_true_iff, Rle_bool_true_iff. - repeat split. - rewrite dist_same. - generalize D_0, Dmax_7D; lra. - now simpl. - simpl in *; assumption. - rewrite Nat.leb_le. unfold get_ident; simpl; lia. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - destruct Hpred as (Hpred1, Hpred3). - unfold change_frame_origin in Hpred1. - rewrite <- (Hpred1 conf g (r,t,b)). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id. - destruct b; simpl; - rewrite Hconf; - now simpl. - auto. - revert Htarget_spec; intro Htarget_spec. - rewrite <- Htarget_spec. - f_equiv. - apply Hpred. - apply Hpred. - } - rewrite Hconf in *. - rewrite Hmove_to. - specialize (Hlight_on_means conf g_target da Hpred Hpath_backup). - rewrite Hconf_target_round in Hlight_on_means. - unfold get_alive in Hlight_on_means; - simpl (snd _) in Hlight_on_means; - rewrite <- Halive_target_round in Hlight_on_means; - rewrite (get_alive_compat Htarget_spec) in Halive_target; - unfold get_alive in Halive_target; - simpl in Halive_target; - specialize (Hlight_on_means Halive_target). - rewrite Hconf_target_round in H; specialize (Hlight_on_means H). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in Hlight_on_means. - revert Hlight_on_means; intro Hlight_on_means. - simpl in Hlight_on_means. - rewrite <- Hlight_on_means. - (* en retravaillant Hmove_Some *) - simpl (fst _). - simpl (choose_new_pos _ _) in *. - revert Htarget_spec; intro Htarget_spec. - destruct Htarget_spec as (Htarget_spec_pos, _). - set (new_pos := choose_new_pos _ _) in *. - assert (Hchoose_correct := choose_new_pos_correct (reflexivity new_pos)). - rewrite (frame_dist _ _ (r,t,b)). - simpl (fst _). - unfold frame_choice_bijection, Frame, Datatypes.id. - rewrite section_retraction. - destruct b; - rewrite Htarget_spec_pos in Hchoose_correct; rewrite Hconf_target in *; simpl in *; lra. - ** assert (Hmove_None := move_to_None Hmove_to). - rewrite Hchange in Hmove_None. - assert (Hsame_pos : get_location (round rbg_ila da conf (Good g)) = get_location (conf (Good g))). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux, map_config in *; rewrite Hchange, Hbool in *. - unfold rbg_fnc; rewrite Hmove_to. - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in *. - simpl (fst _) . - unfold frame_choice_bijection, Frame in Horigin. - rewrite <- Horigin. - rewrite retraction_section. - now cbn. - } - destruct Hmove_None as (other,(Hother_obs, Hmove_other)). - -- unfold obs_from_config, Obs_ILA, map_config in Hother_obs. - rewrite make_set_spec, filter_InA, config_list_InA in Hother_obs. - rewrite 3 andb_true_iff, Rle_bool_true_iff in Hother_obs. - destruct Hother_obs as (([g_other|b_other],Hother_spec), - (((Hother_Dmax, Hother_alive), Hother_alive'), Hother_ident)). - destruct (get_alive (round rbg_ila da conf (Good g_other))) - eqn : Halive_other_round. - ++ exists g_other. - repeat split; try auto. - rewrite Hother_spec in Hother_ident. - rewrite <- Hconf_round. - rewrite <- 2 ident_preserved; try auto. - unfold get_ident in *; simpl in Hother_ident. - rewrite Nat.leb_le in Hother_ident. - destruct Hmove_other as (Hmove_other, Hdist_other_round_2D). - destruct Hmove_other as (other1, (Hother_in,(*(Hother_pos,*) Hother_ident'(*)*))). - revert Hcol; intro Hcol. - unfold no_collision_conf in Hcol. - unfold obs_from_config, Obs_ILA, map_config in Hother_in. - rewrite make_set_spec, filter_InA, config_list_InA in Hother_in. - rewrite 3 andb_true_iff in Hother_in. - destruct Hother_in as (([g_other'|byz], Hother1_spec), (((Hother1_dist,Hother1_alive),Hother1_aliveg), Hother1_ident)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - assert (Hident_g :get_ident (conf (Good g_other')) <= get_ident (conf (Good g))). - { rewrite Hother1_spec in Hother1_ident. - unfold get_ident in *; simpl in *; auto. - rewrite Nat.leb_le, Hconf in *. - simpl in *; auto. - } - assert (get_ident (other) < get_ident (other1)). - unfold get_ident in *; auto. - rewrite Hother_spec, Hother1_spec in H. - unfold get_ident in H; simpl in H. - apply (Nat.lt_le_trans (fst (fst (fst (snd (conf (Good g_other)))))) ( fst (fst (fst (snd (conf (Good g_other')))))) (fst (fst (fst (snd (conf (Good g))))))); auto. -(* now rewrite <- Hident_g.*) - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - rewrite (fst_compat Hother_spec) in Hother_Dmax. - destruct Hmove_other. - rewrite Hother_spec in H0. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - assert (dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))) <= 3 * D)%R. - unfold map_config in *. - rewrite Hchange in Hmove_to. - set (new_pos := choose_new_pos _ _) in *. - assert (Htri_new_pos := RealMetricSpace.triang_ineq (fst (conf (Good g_other))) (retraction (frame_choice_bijection (r,t,b)) new_pos) (fst (round rbg_ila da conf (Good g)))). - assert (Hnew_correct := choose_new_pos_correct (reflexivity new_pos)). - destruct Hnew_correct as (_,Hdist_new_D). - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hdist_new_D. - assert (Hdist_new_D_aux : (dist (retraction (frame_choice_bijection (r, t, b)) new_pos) - (fst (round rbg_ila da conf (Good g))) <= D)%R). - { assert (Hconf_round_aux : round rbg_ila da conf (Good g) == (pos_round, (ident_round, light_round, alive_round, based_round))). - unfold ILA in *. now rewrite <- Hconf_round. - unfold ILA in *. - rewrite (fst_compat Hconf_round_aux) in Hsame_pos. - rewrite (fst_compat Hconf_round_aux). - rewrite Hsame_pos. - rewrite Hconf. - rewrite (frame_dist _ _ (r,t,b)). - rewrite section_retraction. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in Hdist_new_D. - generalize D_0; simpl in *; lra. - } - assert (Hdist_other_new : (dist (fst (conf (Good g_other))) - (retraction (frame_choice_bijection (r, t, b)) new_pos) <= 2*D)%R). - { - rewrite (frame_dist _ _ (r,t,b)), section_retraction. - unfold frame_choice_bijection, Frame; - destruct b; simpl in *; lra. - } - generalize D_0; simpl in *; lra. - assert (Dp > 4*D)%R. - { - generalize Dmax_7D, D_0. unfold Dp. lra. - } - assert (Htri := RealMetricSpace.triang_ineq (fst (round rbg_ila da conf (Good g_other))) (fst (conf (Good g_other))) (fst (pos_round, (ident_round, light_round, alive_round, based_round)))). - transitivity (dist (fst (round rbg_ila da conf (Good g_other))) (fst (conf (Good g_other))) + - dist (fst (conf (Good g_other))) (fst (pos_round, (ident_round, light_round, alive_round, based_round))))%R. - auto. - rewrite dist_sym at 1; apply Htri. - assert (Hdist_round := dist_round_max_d g_other Hpred Hpath_backup). - unfold equiv, bool_Setoid, eq_setoid in *; - rewrite Rle_bool_true_iff in *. - rewrite Hother_spec in Hother_alive; unfold get_alive in *; simpl in Hother_alive. - specialize (Hdist_round Hother_alive). - rewrite dist_sym in Hdist_round. - unfold Dp. - Unset Printing All. - unfold ILA in *; rewrite <- Hconf_round. - fold Dp. - transitivity (D + dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))))%R. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - apply Rplus_le_compat_r. - apply Hdist_round. - transitivity (D + 3*D)%R. - apply Rplus_le_compat_l. - apply H1. - lra. - ++ assert (Halive_near := @executed_means_alive_near conf g_other da Hpred). - assert (Hother_alive_aux : get_alive (conf (Good g_other)) = true). - { - rewrite <- Hother_alive. - rewrite Hother_spec; - unfold get_alive; simpl; auto. - } - revert Hpath_backup; intro Hpath_backup. - destruct (nat_eq_eqdec (get_ident (conf (Good g_other))) 0). - rewrite (ident_preserved conf g_other Hpred) in e. - assert (Hfalse := ident_0_alive (round rbg_ila da conf) g_other e). - rewrite Hfalse in *; discriminate. - rename c into Hother_else. - specialize (Halive_near Hother_alive_aux Hother_else Halive_other_round Hpath_backup). - destruct (Hpath_backup g_other Hother_alive_aux). - destruct Hother_else. - now rewrite H. - destruct Halive_near as (g_near, (Hnear_ident, (Hnear_dist, (Hnear_based, Hnear_alive)))). - destruct H as (g_0, (Halive_0, (Hdist_0, (Hident_0, Hbased_0)))). - exists g_0; repeat split; auto. - - now rewrite dist_sym. - rename H into Hother_path. - (*si near marche aps, le bourreau de near marche, mais aussi near ne devrait aps mourrir: il est bourreau donc il bouge *) - - - -(* je voudrais un axiom qui dit que pour chaque configuration, soit c'est la première, où tout est bon, soit il existe une DA qui permet d'y acceder par un round. Si j'ai ça, alors je peux avoir light_off_means_alive_next car en vrai, la prop d'avant marche sur 3 round: si on est light_off, c'est qu'au round -1, on n'avait personne trop pret ni trop loins, et donc personne ne peut faire en sorte que l'on meurt au round . à tout round, sans avoir à m'inquiéter d'avoir 3 round dans mes lemmes de continuités. - - *) - - - - assert (get_alive (round rbg_ila da conf (Good g_near)) = true). - { - rewrite round_good_g; try auto. - simpl. - unfold upt_aux. - simpl. - destruct (conf (Good g_near)) as (?, (((?,?),?),?)) eqn : Hconf_near; simpl. - unfold get_based in Hnear_based; simpl in Hnear_based; rewrite Hnear_based. - simpl. - destruct (upt_robot_dies_b _ g_near) eqn : Hbool_near. - - assert (Hfalse : get_alive (round rbg_ila da conf (Good g_near)) = false). - rewrite round_good_g; try auto. - simpl. - unfold upt_aux. - rewrite Hbool_near. - unfold get_alive ; simpl. - rewrite Hconf_near, Hnear_based; simpl. - now simpl. - rewrite <- Hconf_near in *. - assert (Hlight_false : get_light (conf (Good g_near)) = true). - apply (Hexecuted g_near da Hpred Hnear_alive Hfalse). - assert (Hlight_true : get_light (conf (Good g_near)) = false). - apply (Hexecutioner g_near da Hpred Hnear_alive). - exists g_other. - repeat split; try auto. - destruct (get_based (conf (Good g_other))) eqn : Hfalse_based; try auto. - rewrite (get_ident_compat Hother_spec) in Hother_ident. - destruct (Hbased_higher) as (_,(_,Hhi)). - specialize (Hhi g_other g Hfalse_based). - unfold get_based in *; rewrite Hconf in Hhi. - simpl in Hhi; specialize (Hhi (reflexivity _)). - rewrite Nat.leb_le, Hconf in *; unfold get_ident in *; simpl in *; lia. - rewrite dist_sym. auto. - rewrite Hlight_true in *. - discriminate. - - unfold get_alive in *; - simpl; auto. - } - exists g_near. - repeat split; auto. - rewrite <- ident_preserved; auto. - transitivity (get_ident (conf (Good g_other))); auto. - - (* début *) - destruct Hmove_other as ((copy, (Hcopy_spec, (*(Hcopy_pos,*) Hcopy_ident(*)*))), _). - unfold obs_from_config, Obs_ILA in Hcopy_spec. - rewrite make_set_spec, filter_InA, config_list_InA in Hcopy_spec. - rewrite 3 andb_true_iff, Rle_bool_true_iff in Hcopy_spec. - destruct Hcopy_spec as (([g_inf|byz],Hinf_spec), ((Hinf_dist, Hinf_alive), Hinf)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - rewrite Nat.leb_le in Hinf. - rewrite <- Hconf_round, <- ident_preserved; try auto. - apply (Nat.lt_le_trans _ (get_ident copy) _). - rewrite Hother_spec in Hcopy_ident. - unfold get_ident in *; now simpl in *. - unfold get_ident in *; now simpl in *. - intros x y Hxy. - rewrite (RelationPairs.fst_compat _ _ _ _ Hxy). - rewrite (get_alive_compat Hxy). - rewrite (get_ident_compat Hxy). - reflexivity. - assert (Hdist_round_g_near := @dist_round_max_d g_near conf da Hpred Hpath_backup Hnear_alive). - unfold equiv, bool_Setoid, eq_setoid in Hdist_round_g_near; - rewrite Rle_bool_true_iff in Hdist_round_g_near. - destruct Hmove_other as (?,Hdist_other). - rewrite Hother_spec, dist_sym in Hdist_other. - revert Hdist_other; intro Hdist_other. - assert (dist (fst (conf (Good g_other))) (fst (round rbg_ila da conf (Good g))) <= 3 * D)%R. - unfold map_config in *. - rewrite Hchange in Hmove_to. - set (new_pos := choose_new_pos _ _) in *. - assert (Htri_new_pos := RealMetricSpace.triang_ineq (fst (conf (Good g_other))) (retraction (frame_choice_bijection (r,t,b)) new_pos) (fst (round rbg_ila da conf (Good g)))). - assert (Hnew_correct := choose_new_pos_correct (reflexivity new_pos)). - destruct Hnew_correct as (_,Hdist_new_D). - destruct Hpred as (Horigin,_). - revert Horigin; intro Horigin. - specialize (Horigin (conf) g (r,t,b) Hchange). - rewrite Hconf in Horigin. - rewrite <- Horigin in Hdist_new_D. - assert (Hdist_new_D_aux : (dist (retraction (frame_choice_bijection (r, t, b)) new_pos) - (fst (round rbg_ila da conf (Good g))) <= D)%R). - { assert (Hconf_round_aux : round rbg_ila da conf (Good g) == (pos_round, (ident_round, light_round, alive_round, based_round))). - unfold ILA in *. now rewrite <- Hconf_round. - unfold ILA in *. - rewrite (fst_compat Hconf_round_aux) in Hsame_pos. - rewrite (fst_compat Hconf_round_aux). - rewrite Hsame_pos. - rewrite Hconf. - rewrite (frame_dist _ _ (r,t,b)). - rewrite section_retraction. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - generalize D_0; simpl in *; lra. - } - assert (Hdist_other_new : (dist (fst (conf (Good g_other))) - (retraction (frame_choice_bijection (r, t, b)) new_pos) <= 2*D)%R). - { - rewrite (frame_dist _ _ (r,t,b)), section_retraction. - rewrite dist_sym. unfold frame_choice_bijection, Frame; - destruct b; simpl in *; lra. - } - generalize D_0; simpl in *; lra. - rewrite <- Hconf_round. - assert (Htri1 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_other))) (get_location (conf (Good g_near)))). - clear Hdist_other. - assert (Htri': (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) <= 4*D)%R). - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id in *. - replace (4*D)%R with (3*D + D)%R by lra. - transitivity ((dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g_other))) + D)%R); try (now generalize D_0; lra). - transitivity ( - dist (fst (round rbg_ila da conf (Good g))) (fst (conf (Good g_other))) + - dist (fst (conf (Good g_other))) (fst (conf (Good g_near))))%R ; - try auto. - apply Rplus_le_compat_l. - apply Hnear_dist. - apply Rplus_le_compat_r. - rewrite dist_sym. - apply H1. - assert (Htri2 := RealMetricSpace.triang_ineq (get_location (round rbg_ila da conf (Good g))) (get_location (conf (Good g_near))) - (get_location (round rbg_ila da conf (Good g_near)))). - unfold Dp. - transitivity (4*D + D)%R. - transitivity (dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) + - dist (get_location (conf (Good g_near))) - (get_location (round rbg_ila da conf (Good g_near))))%R. - apply Htri2. - transitivity ((dist (get_location (round rbg_ila da conf (Good g))) - (get_location (conf (Good g_near))) + D))%R. - apply Rplus_le_compat_l. - apply Hdist_round_g_near. - apply Rplus_le_compat_r. - apply Htri'. - generalize Dmax_7D, D_0; lra. - - (* fin *) - ++ - try (assert (Hfalse := In_Bnames b_other); - now simpl in *). - ++ intros x y Hxy. - rewrite (fst_compat Hxy). - rewrite (get_alive_compat Hxy), (get_ident_compat Hxy). - reflexivity. -Qed. - - - -Lemma executioner_means_light_off_init : forall conf, - config_init_pred conf -> - executioner_means_light_off conf. -Proof. - intros conf Hconf_pred. - intros g da Hpred Halive Hexecutioner. - destruct Hconf_pred as (?,(?,?)). - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - specialize (H0 (Good g)). - rewrite Hconf in H0. - now simpl. -Qed. - -Lemma executed_means_light_on_init : forall conf, - config_init_pred conf -> - executed_means_light_on conf. -Proof. - intros conf Hconf_pred. - intros g da Hpred Halive Halive_round. - assert (Halive_round' := Halive_round). - rewrite round_good_g in Halive_round; try auto. - simpl in *. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - unfold upt_aux, map_config in *. - unfold get_alive in *. - rewrite Hconf in *. - destruct (based) eqn : Hbased. - simpl in *. - unfold config_init_pred in *. - unfold config_init_based_0 in *. - destruct Hconf_pred as (_,(_,(_,(_,(_,Hap0))))). - destruct (upt_robot_too_far _) ; simpl in Halive_round; try (rewrite Halive in *; discriminate). - simpl in Halive_round. - destruct (upt_robot_dies_b _) eqn : Hbool; - try (simpl in *; rewrite Halive in *; discriminate). - simpl in *. - unfold upt_robot_dies_b in *. - destruct Hconf_pred as (?,(?,(?,?))). - unfold config_init_lower, config_init_not_close in *. - specialize (H (Good g)); specialize (H1(Good g)). - rewrite Hconf in *. - rewrite orb_true_iff in Hbool; destruct Hbool as [Hex|Hneg]. - + rewrite existsb_exists in Hex. - destruct Hex as (rila, (Hin,Hex)). - rewrite filter_In in Hin. - rewrite config_list_spec, 2 andb_true_iff in Hin. - destruct Hin as (Hin, ((Hident', Halive_in), Hbased')). - rewrite in_map_iff in Hin. - destruct Hin as ([g_other|byz], (Hother_spec, Hin)); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - specialize (H1 (reflexivity _) (Good g_other)). - rewrite <- Hother_spec in Hex. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - rewrite Rle_bool_true_iff in Hex. - destruct (conf (Good g_other)) as (p_other,(((i_other, l_other), a_other), b_other)) eqn : Hconf_other. - simpl (fst _) in Hex. - rewrite (frame_dist _ _ (r,t,b)) in H1. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - rewrite dist_sym in H1. - rewrite <- Hother_spec in Hbased'; rewrite negb_true_iff in Hbased'; unfold get_based in Hbased'; simpl in Hbased'. - specialize (H1 Hbased'). - destruct b; simpl in *; lra. - + destruct H as [H_0| Hnot_far]. - auto. - assert (get_alive (round rbg_ila da conf (Good g)) = true). - apply ident_0_alive. - rewrite <- ident_preserved; auto. - rewrite Hconf; unfold get_ident; now simpl. - unfold get_alive in *; rewrite H in *; discriminate. - rewrite forallb_existsb, forallb_forall in Hneg. - destruct Hnot_far as ([g_other|byz], Hnot_far); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - destruct (conf (Good g_other)) as (p_other,(((i_other, l_other), a_other), b_other)) eqn : Hconf_other. - destruct (change_frame da conf g) as ((r,t),b) eqn : Hchange. - specialize (Hneg ((comp (bij_rotation r) - (comp (bij_translation t) (if b then reflexion else Similarity.id))) - (fst (conf (Good g_other))), snd (conf (Good g_other)))). - rewrite negb_true_iff, Rle_bool_false_iff in Hneg. - destruct Hneg. - rewrite filter_In, config_list_spec, in_map_iff, 2 andb_true_iff. - repeat split. - exists (Good g_other); try auto. - split; try auto. - destruct b; rewrite Hconf_other; now simpl. - apply In_names. - unfold get_ident in *; simpl in *. - now rewrite Hconf_other, Nat.ltb_lt. - now rewrite Hconf_other; unfold get_alive; simpl in *; auto. - rewrite negb_true_iff. - destruct (b_other) eqn : Hbased_other. - destruct H2 as (?,(?,?)). - specialize (H2 g g_other). - rewrite Hconf_other, Hconf in *. - simpl in H2; specialize (H2 (reflexivity _) (reflexivity _)). - unfold get_ident in *; simpl in *; intuition. - unfold get_based; simpl; rewrite Hconf_other; auto. - simpl (fst _). - destruct Hnot_far as (?,(?,Hpos)). - rewrite (frame_dist _ _ (r,t,b)) in Hpos. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, frame_choice_bijection, Frame, Datatypes.id in *. - rewrite dist_sym. - destruct b; rewrite Hconf_other; simpl in *; lra. -Qed. - -Lemma exists_at_less_that_Dp_init : forall conf, - config_init_pred conf -> - exists_at_less_that_Dp conf. -Proof. - intros conf (Hlower, (Hoff, Hnot_close)) g Halive Hall_lighted Hall. - destruct (conf (Good g)) as (pos, (((ident, light), alive), based)) eqn : Hconf. - specialize (Hlower (Good g)). - rewrite Hconf in Hlower. - destruct Hlower as [Hfalse|([g_other|byz],Hlower)]; - try (unfold get_alive in *; now simpl in *); - try (assert (Hfalse := In_Bnames byz); - now simpl in *). - unfold get_ident in *; simpl in *; lia. - destruct (conf (Good g_other)) as (p_other, (((i_other, l_other), a_other), b_other)) eqn : Hconf_other. - specialize (Hall g_other). - assert (Hfalse : get_light (conf (Good g_other)) = true). - { - apply Hall; rewrite Hconf_other; unfold get_alive, get_ident; try now simpl. - unfold get_location, State_ILA, OnlyLocation, AddInfo, get_location, Datatypes.id. - rewrite dist_sym; simpl in *; lra. - } - specialize (Hoff (Good g_other)). - rewrite Hconf_other in *. - unfold get_light in *; simpl in *; rewrite Hoff in *; discriminate. -Qed. - -Definition NoCollAndPath e := - Stream.forever (Stream.instant (fun conf => no_collision_conf conf /\ path_conf conf) ) e. - -Instance no_collision_conf_compat : Proper (equiv ==> equiv) no_collision_conf. -Proof. - intros x y Hxy. - unfold no_collision_conf. - split; intros. - rewrite <- (get_alive_compat (Hxy (Good g))) in *. - rewrite <- (get_alive_compat (Hxy (Good g'))) in *. - rewrite <- (get_based_compat (Hxy (Good g))) in *. - rewrite <- (get_based_compat (Hxy (Good g'))) in *. - specialize (H g g' H0 H1 H2 H3 H4). - rewrite <- (get_location_compat _ _ (Hxy (Good g))). - rewrite <- (get_location_compat _ _ (Hxy (Good g'))). - apply H. - rewrite (get_alive_compat (Hxy (Good g))) in *. - rewrite (get_alive_compat (Hxy (Good g'))) in *. - rewrite (get_based_compat (Hxy (Good g))) in *. - rewrite (get_based_compat (Hxy (Good g'))) in *. - specialize (H g g' H0 H1 H2 H3 H4). - rewrite (get_location_compat _ _ (Hxy (Good g))). - rewrite (get_location_compat _ _ (Hxy (Good g'))). - apply H. -Qed. - -Instance path_conf_compat : Proper (equiv ==> equiv) path_conf. -Proof. - intros x y Hxy. - unfold path_conf. - split; intros. - rewrite <- (get_alive_compat (Hxy (Good g))) in *. - specialize (H g H0). - destruct H as [H|(g',(Halive,(Hdist,Hident)))]. - now left; rewrite (get_ident_compat (Hxy (Good g))) in *. - right. - exists g'. - rewrite (Hxy (Good g)) in *. - rewrite (Hxy (Good g')) in *. - auto. - rewrite (get_alive_compat (Hxy (Good g))) in *. - specialize (H g H0). - destruct H as [H|(g',(Halive,(Hdist,Hident)))]. - now left; rewrite (get_ident_compat (Hxy (Good g))) in *. - right. - exists g'. - rewrite (Hxy (Good g)) in *. - rewrite (Hxy (Good g')) in *. - auto. -Qed. - -Definition conf_pred conf := - no_collision_conf conf /\ based_higher conf /\ path_conf conf /\ - executed_means_light_on conf /\ executioner_means_light_off conf /\ - exists_at_less_that_Dp conf. - -Definition exists_at_base conf := exists g, get_based (conf (Good g)) = true. - -Definition exists_at_based e := Stream.forever (Stream.instant (exists_at_base)) e. - -Lemma validity: - forall(demon : demon_ila_type) conf, - conf_pred conf -> - demon_ILA demon -> - exists_at_based (execute rbg_ila demon conf) -> - NoCollAndPath (execute rbg_ila demon conf). -Proof. - cofix Hcoind. constructor. - simpl. - split. - now destruct H. - now destruct H. - assert (Hexec_tail := @execute_tail (R2*ILA)%type Loc State_ILA _ Obs_ILA). - rewrite Hexec_tail in *. - apply (Hcoind (Stream.tl demon)). - unfold conf_pred in *; split; - destruct H as (?,(?,(?,(?,(?,?))))); destruct H0. - apply no_collision_cont; try auto. - split. - apply based_higher_round; try auto. - destruct H1. - simpl in H1. - apply H1. - repeat split. - - apply path_round; try auto. - apply executed_means_light_on_round; try auto. - apply executioner_means_light_off_round; try auto. - apply exists_at_less_round; try auto. - destruct H0. - apply H2. - apply H1. -Qed. - -Parameter config_init: config. -Axiom config_init_pred_true : config_init_pred config_init. - -Lemma based_higher_init conf : config_init_pred conf -> based_higher conf. -Proof. - intros Hpred; split; simpl; intros; - unfold config_init_pred in *; - unfold config_init_based_0 in *; - repeat destruct Hpred as (?,Hpred). - specialize (Hpred g H). - intuition. - intuition. -Qed. - - - -Lemma validity_conf_init: - forall demon, demon_ILA demon -> exists_at_based (execute rbg_ila demon config_init) -> - NoCollAndPath (execute rbg_ila demon config_init). -Proof. - intros. - apply validity. - split; generalize config_init_pred_true; intros. - now apply no_collision_init. - split. - apply based_higher_init; auto. - repeat split. - now apply paht_conf_init. - now apply executed_means_light_on_init. - now apply executioner_means_light_off_init. - now apply exists_at_less_that_Dp_init. - auto. - auto. -Qed. diff --git a/Core/Configuration.v b/Core/Configuration.v index 49b95291ac92b0e253270f49b43821d4220ce0dc..0aae6b8cec081218c2ab6c2d0ba2dc3837aa598c 100644 --- a/Core/Configuration.v +++ b/Core/Configuration.v @@ -16,6 +16,7 @@ This file is distributed under the terms of the CeCILL-C licence *) +Require Import Utf8. Require Import SetoidList. Require Import SetoidDec. Require Import Decidable. @@ -38,61 +39,56 @@ Context `{State}. Context `{Names}. (** Equality of configurations is extensional. *) -Global Instance configuration_Setoid : Setoid configuration := fun_equiv ident _. +Global Instance configuration_Setoid : Setoid configuration := fun_Setoid ident _. Global Instance configuration_compat : forall config : configuration, Proper (Logic.eq ==> equiv) config. Proof using . repeat intro. now subst. Qed. +Open Scope program_scope. + (** The lists of positions for good, Byzantine, and all robots. *) -Definition Gpos := fun config : configuration => List.map (fun g => config (Good g)) Gnames. -Definition Bpos := fun config : configuration => List.map (fun b => config (Byz b)) Bnames. +Definition Gpos := fun config : configuration => List.map (config ∘ Good) Gnames. +Definition Bpos := fun config : configuration => List.map (config ∘ Byz) Bnames. Definition config_list := fun config => Gpos config ++ Bpos config. -Lemma Gpos_spec : forall config, Gpos config = List.map (fun g => config (Good g)) Gnames. +Lemma Gpos_spec : forall config, Gpos config = List.map (config ∘ Good) Gnames. Proof using . reflexivity. Qed. -Lemma Bpos_spec : forall config, Bpos config = List.map (fun g => config (Byz g)) Bnames. +Lemma Bpos_spec : forall config, Bpos config = List.map (config ∘ Byz) Bnames. Proof using . reflexivity. Qed. +Close Scope program_scope. + Lemma config_list_spec : forall config, config_list config = List.map config names. Proof using . intros. unfold config_list, names. rewrite map_app. now do 2 rewrite map_map. Qed. (** Compatilities with equivalences. *) Global Instance Gpos_compat : Proper (@equiv _ configuration_Setoid ==> eqlistA equiv) Gpos. -Proof using . -intros f g Hfg. eapply map_extensionalityA_compat; reflexivity || autoclass; []. -intros x y Hxy. cbn in Hxy. subst. apply Hfg. -Qed. +Proof using . intros f g Hfg. rewrite 2Gpos_spec, Hfg. reflexivity. Qed. Global Instance Bpos_compat : Proper (@equiv _ configuration_Setoid ==> eqlistA equiv) Bpos. -Proof using . -intros f g Hfg. eapply map_extensionalityA_compat; reflexivity || autoclass; []. -intros x y Hxy. cbn in Hxy. subst. apply Hfg. -Qed. +Proof using . intros f g Hfg. rewrite 2Bpos_spec, Hfg. reflexivity. Qed. Global Instance config_list_compat : Proper (@equiv _ configuration_Setoid ==> eqlistA equiv) config_list. -Proof using . -intros f g Hfg. rewrite 2 config_list_spec. f_equiv. -intros x y Hxy. cbn in Hxy. subst. apply Hfg. -Qed. +Proof using . intros f g Hfg. rewrite 2config_list_spec, Hfg. reflexivity. Qed. (** Properties w.r.t. [InA] and [length]. *) Lemma Gpos_InA : forall l config, InA equiv l (Gpos config) <-> exists g, equiv l (config (Good g)). Proof using . -intros. rewrite Gpos_spec, InA_map_iff; autoclass; [|]. -+ split; intros [g Hg]; exists g. +intros. rewrite Gpos_spec, InA_map_iff; autoclass; +try (now repeat intro; cbn in *; now subst); []. (* for 8.16 and 8.17 *) +split; intros [g Hg]; exists g. - now symmetry. - split; try (now symmetry); []. rewrite InA_Leibniz. apply In_Gnames. -+ repeat intro. cbn in *. now subst. Qed. Lemma Bpos_InA : forall l config, InA equiv l (Bpos config) <-> exists b, equiv l (config (Byz b)). Proof using . -intros. rewrite Bpos_spec, InA_map_iff; autoclass; [|]. -+ split; intros [b Hb]; exists b. +intros. rewrite Bpos_spec, InA_map_iff; autoclass; +try (now repeat intro; cbn in *; now subst); []. (* for 8.16 and 8.17 *) +split; intros [b Hb]; exists b. - now symmetry. - split; try (now symmetry); []. rewrite InA_Leibniz. apply In_Bnames. -+ repeat intro. cbn in *. now subst. Qed. Lemma config_list_InA : forall l config, InA equiv l (config_list config) <-> exists id, equiv l (config id). @@ -138,7 +134,7 @@ intros configâ‚ configâ‚‚. split; intro Hneq. - inversion Hin. - inversion_clear Habs. inversion_clear Hin; solve [now subst | now apply IHl]. } induction names as [| id l]. - - now elim Hlist. + - now contradiction Hlist. - cbn in Hlist. destruct (equiv_dec (configâ‚ id) (configâ‚‚ id)) as [Hid | Hid]. -- apply IHl. intro Heq. apply Hlist. now constructor. -- eauto. @@ -147,40 +143,6 @@ Qed. End Configuration. -(** Applying a function on all states of a configuration. *) - -Section MapConfig. - -Context `{Location}. -Context {info1 info2 : Type}. -Context {St1 : @State _ info1}. -Context {St2 : @State _ info2}. -Context `{Names}. - -Definition map_config (f : info1 -> info2) (config : @configuration _ _ St1 _) : configuration := - fun id => f (config id). - -Global Instance map_config_compat : - Proper ((equiv ==> equiv) ==> @equiv _ configuration_Setoid ==> @equiv _ configuration_Setoid) map_config. -Proof using . intros f g Hfg ? ? Hconfig id. unfold map. apply Hfg, Hconfig. Qed. - -Lemma config_list_map : forall f, Proper (equiv ==> equiv) f -> - forall config, config_list (map_config f config) == List.map f (config_list config). -Proof using . intros. now rewrite 2 config_list_spec, map_map. Qed. - -End MapConfig. - -Arguments map_config {_} {info1} {info2} {_} {_} {_} f config id /. - -Lemma map_config_id `{State} `{Names} : forall config, - map_config Datatypes.id config == config. -Proof using . now repeat intro. Qed. - -Lemma map_config_merge `{Location} {T U V : Type} `{@State _ T} `{@State _ U} `{@State _ V} `{Names} : - forall (f : T -> U) (g : U -> V), Proper (equiv ==> equiv) f -> Proper (equiv ==> equiv) g -> - forall config : configuration, map_config g (map_config f config) == map_config (fun x => g (f x)) config. -Proof using . now repeat intro. Qed. - (** Injective configurations *) Definition config_injective `{State} `{Names} := Util.Preliminary.injective (@eq ident) (@equiv _ state_Setoid). @@ -231,3 +193,46 @@ intros config. split; intro Hinj. rewrite <- NoDupA_Leibniz, Hnames in Hnodup. inversion_clear Hnodup. subst. tauto. + destruct Hinj as [id [id' [Hid Heq]]]. intro Habs. apply Habs in Heq. contradiction. Qed. + +(** Applying a function on all states of a configuration. *) + +Section MapConfig. + +Context {L : Location} {info1 info2 : Type}. +Context {St1 : @State _ info1} {St2 : @State _ info2}. +Context {N : Names}. + +Definition map_config (f : info1 -> info2) (config : @configuration _ _ St1 _) : configuration := + fun id => f (config id). + +Global Instance map_config_compat : + Proper ((equiv ==> equiv) ==> @equiv _ configuration_Setoid ==> @equiv _ configuration_Setoid) map_config. +Proof using . intros f g Hfg ? ? Hconfig id. unfold map. apply Hfg, Hconfig. Qed. + +Lemma config_list_map : forall f config, + config_list (map_config f config) == List.map f (config_list config). +Proof using . intros. now rewrite 2 config_list_spec, map_map. Qed. + +Lemma map_config_inj' : ∀ (f : info1 -> info2), + Preliminary.injective equiv equiv f -> Preliminary.injective equiv equiv (map_config f). +Proof using . intros f H config1 config2 Hc id. apply H, Hc. Qed. + +Lemma map_config_inj : ∀ (f : info1 -> info2) (config : @configuration _ _ St1 _), + Preliminary.injective equiv equiv f -> config_injective config + -> config_injective (map_config f config). +Proof using . + unfold map_config. intros * Hf Hc id1 id2 Hm. apply Hc, Hf, Hm. +Qed. + +End MapConfig. + +Arguments map_config {_} {info1} {info2} {_} {_} {_} f config id /. + +Lemma map_config_id `{State} `{Names} : forall config, + map_config Datatypes.id config == config. +Proof using . now repeat intro. Qed. + +Lemma map_config_merge `{Location} {T U V : Type} `{@State _ T} `{@State _ U} `{@State _ V} `{Names} : + forall (f : T -> U) (g : U -> V), Proper (equiv ==> equiv) g -> + forall config : configuration, map_config g (map_config f config) == map_config (fun x => g (f x)) config. +Proof using . now repeat intro. Qed. diff --git a/Core/Formalism.v b/Core/Formalism.v index 9ecb74f34b49dd3bf206d436db90b5388d5f33d1..19cce23fd7c789375f88e0815c69a8d298aaa2b9 100644 --- a/Core/Formalism.v +++ b/Core/Formalism.v @@ -29,6 +29,8 @@ Require Import Pactole.Core.State. Require Import Pactole.Core.Configuration. Require Import Pactole.Observations.Definition. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Typeclasses eauto := 5. Remove Hints eq_setoid : typeclass_instances. @@ -51,11 +53,11 @@ Definition execution := Stream.t configuration. (** Good robots have a common program, which we call a [robogram]. It returns some piece of information (e.g. target location) which must form a setoid. *) -Class robot_choice := { robot_choice_Setoid :> Setoid Trobot }. +Class robot_choice := { #[global] robot_choice_Setoid :: Setoid Trobot }. Record robogram `{robot_choice} := { pgm :> observation -> Trobot; - pgm_compat :> Proper (equiv ==> equiv) pgm}. + #[global] pgm_compat :: Proper (equiv ==> equiv) pgm}. Global Instance robogram_Setoid `{robot_choice} : Setoid robogram. simple refine {| equiv := fun r1 r2 => forall s, pgm r1 s == pgm r2 s |}; @@ -87,28 +89,28 @@ Qed. It must at least contain a bijection to compute the change of frame of reference. *) Class frame_choice := { frame_choice_bijection : Tframe -> bijection location; - frame_choice_Setoid :> Setoid Tframe; - frame_choice_bijection_compat :> Proper (equiv ==> equiv) frame_choice_bijection }. + #[global] frame_choice_Setoid :: Setoid Tframe; + #[global] frame_choice_bijection_compat :: Proper (equiv ==> equiv) frame_choice_bijection }. (** An [update_choice] represents the choices the demon makes after a robot computation. *) Class update_choice := { - update_choice_Setoid :> Setoid Tactive; - update_choice_EqDec :> EqDec update_choice_Setoid }. + #[global] update_choice_Setoid :: Setoid Tactive; + #[global] update_choice_EqDec :: EqDec update_choice_Setoid }. (** An [inactive_choice] represents the choices the demon makes when a robot is not activated. *) Class inactive_choice := { - inactive_choice_Setoid :> Setoid Tinactive; - inactive_choice_EqDec :> EqDec inactive_choice_Setoid }. + #[global] inactive_choice_Setoid :: Setoid Tinactive; + #[global] inactive_choice_EqDec :: EqDec inactive_choice_Setoid }. (** These choices are then used by update functions that depend on the model. *) (* RMK: we cannot combine them toghether otherwise we get dependencies on the other parameter. *) Class update_function `{robot_choice} `{frame_choice} `{update_choice} := { update : configuration -> G -> Tframe -> Trobot -> Tactive -> info; - update_compat :> Proper (equiv ==> Logic.eq ==> equiv ==> equiv ==> equiv ==> equiv) update }. + #[global] update_compat :: Proper (equiv ==> Logic.eq ==> equiv ==> equiv ==> equiv ==> equiv) update }. Class inactive_function `{inactive_choice} := { inactive : configuration -> ident -> Tinactive -> info; - inactive_compat :> Proper (equiv ==> Logic.eq ==> equiv ==> equiv) inactive }. + #[global] inactive_compat :: Proper (equiv ==> Logic.eq ==> equiv ==> equiv) inactive }. Context {RC : robot_choice}. Context {FC : frame_choice}. @@ -311,10 +313,10 @@ induction names as [| id l]; simpl. * destruct (round r1 da1 c1 id =?= c1 id) as [Heq1 | Heq1], (round r2 da2 c2 id =?= c2 id) as [Heq2 | Heq2]. + apply IHl. - + elim Heq2. transitivity (round r1 da1 c1 id). + + contradiction Heq2. transitivity (round r1 da1 c1 id). - symmetry. now apply round_compat. - rewrite Heq1. apply Hc. - + elim Heq1. transitivity (round r2 da2 c2 id). + + contradiction Heq1. transitivity (round r2 da2 c2 id). - now apply round_compat. - rewrite Heq2. symmetry. apply Hc. + f_equal. apply IHl. @@ -476,13 +478,8 @@ Qed. (** ** Fairness **) (** A [demon] is [Fair] if at any time it will later activate any robot. *) -(* RMK: This is a stronger version of eventually because P is negated in the Later clause *) -Inductive LocallyFairForOne id (d : demon) : Prop := - | NowFair : activate (Stream.hd d) id == true -> LocallyFairForOne id d - | LaterFair : activate (Stream.hd d) id = false -> - LocallyFairForOne id (Stream.tl d) -> LocallyFairForOne id d. - -Definition Fair : demon -> Prop := Stream.forever (fun d => ∀ id, LocallyFairForOne id d). +Definition Fair : demon -> Prop := + Stream.forever (fun d => ∀ id, Stream.eventually (Stream.instant (fun da => activate da id = true)) d). (** [Between id id' d] means that [id] will be activated before at most [k] steps of [id'] in demon [d]. *) @@ -496,26 +493,17 @@ Inductive Between id id' (d : demon) : nat -> Prop := Between id id' (Stream.tl d) k -> Between id id' d k. (** k-fairnes: Every robot is activated within at most k activation of any other robot. *) -Definition kFair k : demon -> Prop := Stream.forever (fun d => forall id id', Between id id' d k). +Definition kFair k : demon -> Prop := Stream.forever (fun d => ∀ id id', Between id id' d k). (** Compatibility properties *) -Local Lemma LocallyFairForOne_compat_aux : forall id d1 d2, - d1 == d2 -> LocallyFairForOne id d1 -> LocallyFairForOne id d2. -Proof using . -intros id da1 da2 Hda Hfair. revert da2 Hda. induction Hfair; intros da2 Hda; -constructor; solve [now rewrite <- Hda | apply IHHfair; now f_equiv]. -Qed. - -Global Instance LocallyFairForOne_compat : Proper (Logic.eq ==> equiv ==> iff) LocallyFairForOne. +Global Instance Fair_compat : Proper (equiv ==> iff) Fair. Proof using . -intros ? ? ? ? ? Heq. subst. split; intro. -- eapply LocallyFairForOne_compat_aux; eauto. -- symmetry in Heq. eapply LocallyFairForOne_compat_aux; eauto. +apply Stream.forever_compat. intros d1 d2 Hd. +split; intros Hlater id; +(generalize (Hlater id); eapply Stream.eventually_compat; auto; []; +apply Stream.instant_compat; intros ? ? Heq; now rewrite Heq). Qed. -Global Instance Fair_compat : Proper (equiv ==> iff) Fair. -Proof using . apply Stream.forever_compat. intros ? ? Heq. now setoid_rewrite Heq. Qed. - Local Lemma Between_compat_aux : forall id id' k d1 d2, d1 == d2 -> Between id id' d1 k -> Between id id' d2 k. Proof using . @@ -533,21 +521,16 @@ Qed. Global Instance kFair_compat : Proper (Logic.eq ==> equiv ==> iff) kFair. Proof using . intros k ??. subst. apply Stream.forever_compat. intros ?? Heq. now setoid_rewrite Heq. Qed. -(** A robot is never activated before itself with a fair demon! - The fairness hypothesis is necessary, otherwise the robot may never be activated. *) -Lemma Between_same : - forall id (d : demon) k, LocallyFairForOne id d -> Between id id d k. -Proof using . intros id d k Hd. induction Hd; now econstructor; eauto. Qed. - (** A k-fair demon is fair. *) -Lemma Between_LocallyFair : forall id (d : demon) id' k, - Between id id' d k -> LocallyFairForOne id d. -Proof using . intros * Hg. induction Hg; now constructor; trivial; firstorder. Qed. +Lemma Between_eventually_activated : forall id (d : demon) id' k, + Between id id' d k -> Stream.eventually (Stream.instant (fun da => activate da id = true)) d. +Proof using . intros * Hg. induction Hg; constructor; assumption. Qed. Theorem kFair_Fair : forall k (d : demon), kFair k d -> Fair d. Proof using . intro. apply Stream.forever_impl_compat. -intros ? ? id. eauto using (@Between_LocallyFair id _ id). +intros ? ? id. +eapply (@Between_eventually_activated id _ id);eauto. Qed. (** [Between g h d k] is monotonic on [k]. *) @@ -573,7 +556,7 @@ Qed. Lemma FSYNC_implies_0Fair: ∀ d, FSYNC d → kFair 0 d. Proof using . apply Stream.forever_impl_compat. intros s Hs id id'. constructor. apply Hs. Qed. -Corollary FSYNC_implies_fair: ∀ d, FSYNC d → Fair d. +Corollary FSYNC_implies_Fair: ∀ d, FSYNC d → Fair d. Proof using . intros. now eapply kFair_Fair, FSYNC_implies_0Fair. Qed. (** If a demon is 0-fair, then the activation states of all robots are the same: @@ -612,6 +595,44 @@ intros r1 r2 Hr d1 d2 Hd c1 c2 Hc. split; intro Hfirst. - destruct Hd. apply IHHfirst; trivial. now apply round_compat; f_equiv. Qed. +(** Change a demonic action to activate all robots. *) +Definition da_with_all_activated da := {| + activate := fun _ => true; + relocate_byz := da.(relocate_byz); + change_frame := da.(change_frame); + choose_update := da.(choose_update); + choose_inactive := da.(choose_inactive); + precondition_satisfied := da.(precondition_satisfied); + precondition_satisfied_inv := da.(precondition_satisfied_inv); + activate_compat := fun _ _ _ => reflexivity true; + relocate_byz_compat := da.(relocate_byz_compat); + change_frame_compat := da.(change_frame_compat); + choose_update_compat := da.(choose_update_compat); + choose_inactive_compat := da.(choose_inactive_compat) |}. + +Lemma da_with_all_activated_FSYNC_da : forall da, FSYNC_da (da_with_all_activated da). +Proof using . now intros da id. Qed. + +(** An unfair demon activates at least one activable robot if such a robot exists. *) +Definition unfair_da r da config := + (* If assuming all robots are activated, a robot can move *) + moving r (da_with_all_activated da) config <> nil -> + (* then at least a robot moves *) + moving r da config <> nil. + +Definition unfair r d config := + Stream.forever2 (Stream.instant2 (unfair_da r)) d (execute r d config). + +(* Conditional unfair demon. *) +Definition unfair_da_cond P r da := + forall config, + P config -> + (* There exists a way to make some robot move *) + (exists da', moving r da' config <> nil) -> + moving r da config <> nil. + +Definition unfair_cond P r d := Stream.forever (Stream.instant (unfair_da_cond P r)) d. + End Formalism. Arguments robogram {info} {_} {_} {_} {_} {Trobot} {_}. diff --git a/Core/Identifiers.v b/Core/Identifiers.v index 614fb3d8bfee97ac32ff6845320bcebcc909af51..d6a61b02be35b598c7ca4a47e95addcf00234b1b 100644 --- a/Core/Identifiers.v +++ b/Core/Identifiers.v @@ -14,204 +14,11 @@ Require Import SetoidDec SetoidList. Require Import Arith_base. Require Import Lia. Require Import Pactole.Util.Coqlib. +Require Import Pactole.Util.Fin. +Require Import Pactole.Util.Enum. Set Implicit Arguments. -Typeclasses eauto := (bfs). - - -(* TODO: should we add a fold operator? *) -(* FIXME: change the equalities to use equiv and the Setoid class *) - -(** Finite sets as a prefix of natural numbers. *) -Notation "'fin' N" := {n : nat | n < N} (at level 10). - -Lemma subset_dec : forall N (x y : fin N), {x = y} + {x <> y}. -Proof using . -intros N [x Hx] [y Hy]. destruct (Nat.eq_dec x y). -+ subst. left. f_equal. apply le_unique. -+ right. intro Habs. inv Habs. auto. -Qed. - -Lemma eq_proj1 : forall N (x y : fin N), proj1_sig x = proj1_sig y -> x = y. -Proof using . intros N [x Hx] [y Hy] ?. simpl in *. subst. f_equal. apply le_unique. Qed. - -Program Fixpoint build_enum N k (Hle : k <= N) acc : list (fin N) := - match k with - | 0 => acc - | S m => @build_enum N m _ (exist (fun x => x < N) m _ :: acc) - end. -Next Obligation. -lia. -Qed. - -(** A list containing all elements of [fin N]. *) -Definition enum N : list (fin N) := build_enum (Nat.le_refl N) nil. - -(** Specification of [enum]. *) -Lemma In_build_enum : forall N k (Hle : k <= N) l x, In x (build_enum Hle l) <-> In x l \/ proj1_sig x < k. -Proof using . -intros N k. induction k; intros Hle l x; simpl. -+ intuition. -+ rewrite IHk. simpl. split; intro Hin. - - destruct Hin as [[Hin | Hin] | Hin]; intuition; []. - subst. simpl. right. lia. - - destruct Hin as [Hin | Hin]; intuition; []. - assert (Hcase : proj1_sig x < k \/ proj1_sig x = k) by lia. - destruct Hcase as [Hcase | Hcase]; intuition; []. - subst. do 2 left. destruct x; f_equal; simpl in *. apply le_unique. -Qed. - -Lemma In_enum : forall N x, In x (enum N) <-> proj1_sig x < N. -Proof using . intros. unfold enum. rewrite In_build_enum. simpl. intuition. Qed. - -(** Length of [enum]. *) -Lemma build_enum_length : forall N k (Hle : k <= N) l, length (build_enum Hle l) = k + length l. -Proof using . -intros N k. induction k; intros Hle l; simpl. -+ reflexivity. -+ rewrite IHk. simpl. lia. -Qed. - -Lemma enum_length : forall N, length (enum N) = N. -Proof using . intro. unfold enum. now rewrite build_enum_length. Qed. - -(** [enum] does not contain duplicates. *) -Lemma build_enum_NoDup : forall N k (Hle : k <= N) l, - (forall x, In x l -> k <= proj1_sig x) -> NoDup l -> NoDup (build_enum Hle l). -Proof using . -intros N k. induction k; intros Hle l Hin Hl; simpl; auto; []. -apply IHk. -+ intros x [Hx | Hx]. - - now subst. - - apply Hin in Hx. lia. -+ constructor; trivial; []. - intro Habs. apply Hin in Habs. simpl in Habs. lia. -Qed. - -Lemma enum_NoDup : forall N, NoDup (enum N). -Proof using . intro. unfold enum. apply build_enum_NoDup; simpl; intuition; constructor. Qed. - -(** [enum] is sorted in increasing order. *) -Notation Flt := (fun x y => lt (proj1_sig x) (proj1_sig y)). - -Lemma build_enum_Sorted : forall N k (Hle : k <= N) l, - (forall x, In x l -> k <= proj1_sig x) -> Sorted Flt l -> Sorted Flt (build_enum Hle l). -Proof using . -intros N k. induction k; intros Hle l Hin Hl; simpl; auto; []. -apply IHk. -+ intros x [Hx | Hx]. - - now subst. - - apply Hin in Hx. lia. -+ constructor; trivial; []. - destruct l; constructor; []. simpl. apply Hin. now left. -Qed. - -Lemma enum_Sorted : forall N, Sorted Flt (enum N). -Proof using . intro. unfold enum. apply build_enum_Sorted; simpl; intuition. Qed. - -(** Extensional equality of functions is decidable over finite domains. *) -Lemma build_enum_app_nil : forall N k (Hle : k <= N) l, - build_enum Hle l = build_enum Hle nil ++ l. -Proof using . -intros N k. induction k; intros Hle l; simpl. -+ reflexivity. -+ now rewrite (IHk _ (_ :: nil)), IHk, <- app_assoc. -Qed. - -Theorem build_enum_eq : forall {A} eqA N (f g : fin N -> A) k (Hle : k <= N) l, - eqlistA eqA (List.map f (build_enum Hle l)) (List.map g (build_enum Hle l)) -> - forall x, proj1_sig x < k -> eqA (f x) (g x). -Proof using . -intros A eqA N f g k. induction k; intros Hle l Heq x Hx; simpl. -* destruct x; simpl in *; lia. -* assert (Hlt : k <= N) by lia. - assert (Hcase : proj1_sig x < k \/ proj1_sig x = k) by lia. - destruct Hcase as [Hcase | Hcase]. - + apply IHk with (x := x) in Heq; auto. - + subst k. simpl in Heq. rewrite build_enum_app_nil, map_app, map_app in Heq. - destruct (eqlistA_app_split _ _ _ _ Heq) as [_ Heq']. - - now do 2 rewrite map_length, build_enum_length. - - simpl in Heq'. inv Heq'. - assert (Heqx : x = exist (fun x => x < N) (proj1_sig x) Hle). - { clear. destruct x; simpl. f_equal. apply le_unique. } - now rewrite Heqx. -Qed. - -Corollary enum_eq : forall {A} eqA N (f g : fin N -> A), - eqlistA eqA (List.map f (enum N)) (List.map g (enum N)) -> forall x, eqA (f x) (g x). -Proof using . -unfold enum. intros A eqA N f g Heq x. -apply build_enum_eq with (x := x) in Heq; auto; []. apply proj2_sig. -Qed. - -(** Cutting [enum] after some number of elements. *) -Lemma firstn_build_enum_le : forall N k (Hle : k <= N) l k' (Hk : k' <= N), k' <= k -> - firstn k' (build_enum Hle l) = @build_enum N k' Hk nil. -Proof using . -intros N k. induction k; intros Hk l k' Hk' Hle. -* assert (k' = 0) by lia. now subst. -* rewrite build_enum_app_nil, firstn_app, build_enum_length. - replace (k' - (S k + length (@nil (fin N)))) with 0 by lia. - rewrite app_nil_r. - destruct (Nat.eq_dec k' (S k)) as [Heq | Heq]. - + subst k'. rewrite firstn_all2. - - f_equal. apply le_unique. - - rewrite build_enum_length. simpl. lia. - + simpl build_enum. erewrite IHk. - - f_equal. - - lia. -Qed. - -Lemma firstn_build_enum_lt : forall N k (Hle : k <= N) l k', k <= k' -> - firstn k' (build_enum Hle l) = build_enum Hle (firstn (k' - k) l). -Proof using . -intros N k. induction k; intros Hle l k' Hk. -+ now rewrite Nat.sub_0_r. -+ rewrite build_enum_app_nil, firstn_app, build_enum_length, Nat.add_0_r. - rewrite firstn_all2, <- build_enum_app_nil; trivial; []. - rewrite build_enum_length. simpl. lia. -Qed. - -Lemma firstn_enum_le : forall N k (Hle : k <= N), firstn k (enum N) = build_enum Hle nil. -Proof using . intros. unfold enum. now apply firstn_build_enum_le. Qed. - -Lemma firstn_enum_lt : forall N k, N <= k -> firstn k (enum N) = enum N. -Proof using . intros. unfold enum. now rewrite firstn_build_enum_lt, firstn_nil. Qed. - -Lemma firstn_enum_spec : forall N k x, In x (firstn k (enum N)) <-> proj1_sig x < k. -Proof using . -intros N k x. destruct (le_lt_dec k N) as [Hle | Hlt]. -+ rewrite (firstn_enum_le Hle), In_build_enum. simpl. intuition. -+ rewrite (firstn_enum_lt (lt_le_weak _ _ Hlt)). - split; intro Hin. - - transitivity N; trivial; []. apply proj2_sig. - - apply In_enum, proj2_sig. -Qed. - -(** Removing some number of elements from the head of [enum]. *) -Lemma skipn_build_enum_lt : forall N k (Hle : k <= N) l k', k <= k' -> - skipn k' (build_enum Hle l) = skipn (k' - k) l. -Proof using . -intros N k Hle l k' Hk'. apply app_inv_head with (firstn k' (build_enum Hle l)). -rewrite firstn_skipn, firstn_build_enum_lt; trivial; []. -rewrite (build_enum_app_nil Hle (firstn _ _)). -now rewrite build_enum_app_nil, <- app_assoc, firstn_skipn. -Qed. - -Lemma skipn_enum_lt : forall N k, N <= k -> skipn k (enum N) = nil. -Proof using . intros. unfold enum. now rewrite skipn_build_enum_lt, skipn_nil. Qed. - -Lemma skipn_enum_spec : forall N k x, In x (skipn k (enum N)) <-> k <= proj1_sig x < N. -Proof using . -intros N k x. split; intro Hin. -+ assert (Hin' : ~In x (firstn k (enum N))). - { intro Habs. rewrite <- InA_Leibniz in *. revert x Habs Hin. apply NoDupA_app_iff; autoclass; []. - rewrite firstn_skipn. rewrite NoDupA_Leibniz. apply enum_NoDup. } - rewrite firstn_enum_spec in Hin'. split; auto with zarith; []. apply proj2_sig. -+ assert (Hin' : In x (enum N)) by apply In_enum, proj2_sig. - rewrite <- (firstn_skipn k), in_app_iff, firstn_enum_spec in Hin'. intuition lia. -Qed. (** ** Byzantine Robots *) @@ -295,15 +102,15 @@ destruct id as [g | b], id' as [g' | b']; try (now right; discriminate); [|]. - right; intro Habs. now injection Habs. Qed. -Instance ident_Setoid `{Names} : Setoid ident := { equiv := eq; setoid_equiv := eq_equivalence }. -Instance ident_EqDec `{Names} : EqDec ident_Setoid := names_eq_dec. +Global Instance ident_Setoid `{Names} : Setoid ident := { equiv := eq; setoid_equiv := eq_equivalence }. +Global Instance ident_EqDec `{Names} : EqDec ident_Setoid := names_eq_dec. -Instance fun_refl `{Names} : forall A (f : ident -> A) R, +Global Instance fun_refl `{Names} : forall A (f : ident -> A) R, Reflexive R -> Proper (@SetoidClass.equiv ident _ ==> R) f. Proof using . intros A f R HR ? ? Heq. simpl in Heq. now subst. Qed. -Instance list_ident_Setoid `{Names} : Setoid (list ident) := { equiv := eq; setoid_equiv := eq_equivalence }. -Instance list_ident_Eqdec `{Names} : EqDec list_ident_Setoid := list_eq_dec ident_EqDec. +Global Instance list_ident_Setoid `{Names} : Setoid (list ident) := { equiv := eq; setoid_equiv := eq_equivalence }. +Global Instance list_ident_Eqdec `{Names} : EqDec list_ident_Setoid := list_eq_dec ident_EqDec. Lemma fun_names_eq `{Names} : forall {A : Type} eqA f g, @eqlistA A eqA (List.map f names) (List.map g names) -> forall x, eqA (f x) (g x). @@ -316,27 +123,69 @@ unfold names in Heq. repeat rewrite ?map_app, map_map in Heq. apply eqlistA_app_ + now do 2 rewrite map_length. Qed. -(** Given a number of good and byzntine robots, we can build canonical names. +Section Robots. + +(** Given a number of correct and Byzantine robots, we can build canonical names. It is not declared as a global instance to avoid creating spurious settings. *) Definition Robots (n m : nat) : Names. -Proof. -refine {| - nG := n; - nB := m; - G := fin n; - B := fin m; - Gnames := enum n; - Bnames := enum m |}. -+ abstract (intro g; apply In_enum, proj2_sig). -+ abstract (intro b; apply In_enum, proj2_sig). -+ apply enum_NoDup. -+ apply enum_NoDup. -+ apply enum_length. -+ apply enum_length. -+ apply subset_dec. -+ apply subset_dec. -+ intros ? ?. apply enum_eq. -+ intros ? ?. apply enum_eq. +Proof using . + refine {| + nG := n; + nB := m; + G := fin n; + B := fin m; + Gnames := enum n; + Bnames := enum m |}. + + abstract (intro g; apply In_enum, fin_lt). + + abstract (intro b; apply In_enum, fin_lt). + + apply enum_NoDup. + + apply enum_NoDup. + + apply enum_length. + + apply enum_length. + + apply fin_dec. + + apply fin_dec. + + intros ? ?. apply enum_eq. + + intros ? ?. apply enum_eq. Defined. Global Opaque G B. +(* TODO: discuss this +Section NM. + +Variables n m: nat. + +Notation GRob := (@G (Robots n m)). +Notation BRob := (@B (Robots n m)). + +Lemma GRob_Robots : GRob = fin n. +Proof using . reflexivity. Qed. + +Lemma BRob_Robots : BRob = fin m. +Proof using . reflexivity. Qed. + +Lemma GRob_Robots_eq_iff : forall g1 g2 : GRob, g1 = g2 :> GRob <-> g1 = g2 :> fin n. +Proof using . reflexivity. Qed. + +Lemma BRob_Robots_eq_iff : forall b1 b2 : BRob, b1 = b2 :> BRob <-> b1 = b2 :> fin m. +Proof using . reflexivity. Qed. + +Definition good0 : GRob := fin0. + +Definition byz0 : BRob := fin0. + +Lemma all_good0 : forall g : GRob, n = 1 -> g = good0. +Proof using . intros * H. rewrite GRob_Robots_eq_iff. apply all_fin0, H. Qed. + +Lemma all_good_eq : forall g1 g2 : GRob, n = 1 -> g1 = g2. +Proof using ltc_l_n. intros * H. rewrite GRob_Robots_eq_iff. apply all_eq, H. Qed. + +Lemma all_byz0 : forall b : BRob, m = 1 -> b = byz0. +Proof using . intros * H. rewrite BRob_Robots_eq_iff. apply all_fin0, H. Qed. + +Lemma all_byz_eq : forall b1 b2 : BRob, m = 1 -> b1 = b2. +Proof using ltc_l_m. intros * H. rewrite BRob_Robots_eq_iff. apply all_eq, H. Qed. +End NM. + +*) + +End Robots. diff --git a/Core/State.v b/Core/State.v index e687468389c853df61616bccffb2b8166eacb251..b8cec83788aeec79b85c22804772baaa5f6390db 100644 --- a/Core/State.v +++ b/Core/State.v @@ -34,8 +34,8 @@ Require Import Pactole.Core.Identifiers. Instead, the user must explicitely provide the instance. *) Class Location := { location : Type; - location_Setoid :> Setoid location; - location_EqDec :> EqDec location_Setoid }. + #[global] location_Setoid :: Setoid location; + #[global] location_EqDec :: EqDec location_Setoid }. Definition make_Location (T : Type) `{EqDec T} := {| location := T |}. Arguments make_Location T {_} {_}. @@ -54,18 +54,15 @@ Arguments make_Location T {_} {_}. Class State `{Location} info := { get_location : info -> location; (** States are equipped with a decidable equality *) - state_Setoid :> Setoid info; - state_EqDec :> EqDec state_Setoid; + #[global] state_Setoid :: Setoid info; + #[global] state_EqDec :: EqDec state_Setoid; (** Lifting a change of frame from a location to a full state, under some precondition *) precondition : (location -> location) -> Type; lift : sigT precondition -> info -> info; - lift_id : forall Pid, lift (existT precondition id Pid) == id; -(* lift_compose : forall f g state, Proper (equiv ==> equiv) f -> Proper (equiv ==> equiv) g -> - lift f (lift g state) == lift (fun x => f (g x)) state;*) get_location_lift : forall f state, get_location (lift f state) == projT1 f (get_location state); (** Compatibility properties *) - get_location_compat :> Proper (equiv ==> equiv) get_location; - lift_compat :> + #[global] get_location_compat :: Proper (equiv ==> equiv) get_location; + #[global] lift_compat :: Proper ((equiv ==> equiv)%signature @@ (@projT1 _ precondition) ==> equiv ==> equiv) lift }. (* We cannot have [lift_compat :> Proper (equiv ==> equiv ==> equiv) lift] because we also need extensionality in the input function, which function's [equiv] has not. *) @@ -92,11 +89,9 @@ refine {| precondition := precondition |}. Proof. + apply prod_EqDec; apply state_EqDec || apply location_EqDec. -+ intros Pid []. simpl. split; try reflexivity; []. apply lift_id. -(* + intros f g [] **. simpl. split; try reflexivity; []. now apply lift_compose. *) -+ intros f []. simpl. apply get_location_lift. -+ intros [] [] []. simpl. now apply get_location_compat. -+ intros f g Hfg []. simpl. split. ++ intros f []. cbn. apply get_location_lift. ++ intros [] [] []. cbn. now apply get_location_compat. ++ intros f g Hfg []. cbn. split. - now apply lift_compat. - now apply Hfg. Defined. @@ -108,11 +103,9 @@ refine {| get_location := fun x => get_location (fst x); precondition := precondition |}. Proof. + apply prod_EqDec; apply state_EqDec || auto. -+ intros Pid []. simpl. split; try reflexivity; []. apply lift_id. -(* + intros f g [] **. simpl. split; try reflexivity; []. now apply lift_compose. *) -+ intros f []. simpl. apply get_location_lift. -+ intros [] [] []. simpl. now apply get_location_compat. -+ intros f g Hfg [] [] []. simpl in *. split; trivial; []. now apply lift_compat. ++ intros f []. cbn. apply get_location_lift. ++ intros [] [] []. cbn. now apply get_location_compat. ++ intros f g Hfg [] [] []. cbn in *. split; trivial; []. now apply lift_compat. Defined. (* RMK: As [AddLocation] has less parameters than [AddInfo], its priority is higher, @@ -129,10 +122,9 @@ Proof. + apply sig_Setoid, state_Setoid. + autoclass. + apply (snd (projT2 f)), proj2_sig. -+ intros ? ?. simpl. apply lift_id. -+ intros f x. simpl. apply get_location_lift. ++ intros f x. cbn. apply get_location_lift. + repeat intro. now apply get_location_compat. -+ intros f g Hfg x y Hxy. simpl. now apply lift_compat. ++ intros f g Hfg x y Hxy. cbn. now apply lift_compat. Defined. (** A more general version restricting a state to have a dependent witness of some type. *) @@ -146,9 +138,8 @@ Proof. + apply sigT_Setoid, state_Setoid. + autoclass. + apply (snd (projT2 f)), projT2. -+ intros ? ?. simpl. apply lift_id. -+ intros f x. simpl. apply get_location_lift. ++ intros f x. cbn. apply get_location_lift. + repeat intro. now apply get_location_compat. -+ intros f g Hfg x y Hxy. simpl. now apply lift_compat. ++ intros f g Hfg x y Hxy. cbn. now apply lift_compat. Defined. (* end show *) diff --git a/DEV.md b/DEV.md new file mode 100644 index 0000000000000000000000000000000000000000..9597e8b9db1f558fefbd28d14538788f8d384d46 --- /dev/null +++ b/DEV.md @@ -0,0 +1,13 @@ +# Contributing + +The golden rule is that **MASTER SHOULD COMPILE AT ALL TIME** and does not +contain in-progress developments. + +More precisely, development never happens on master: each topic should have +its own branch (feel free to create one if needed) and work happens there, +even when several people are involved. + +Once a set of changes (e.g. a new case study) is complete and polished enough +(comments, adequate identation, no use of generated hypothesis names, etc.), +you may merge it to master by squashing its commits into meaningful pieces +(only a few, usually one may be enough) or submit a pull request. diff --git a/INSTALL.md b/INSTALL.md new file mode 100644 index 0000000000000000000000000000000000000000..1d7367e4b0694976bd6ebbc2c51f603f6dacffaf --- /dev/null +++ b/INSTALL.md @@ -0,0 +1,117 @@ + +# Requirements + +- Coq 8.19 or 8.20 (including the executable `coqc`, `codep`, `coq_makefile`) +- GNU `make` + +# Configuration + +You should perform once the following command to generate the Makefile: + +```bash +coq_makefile -f _CoqProject -o Makefile +``` + +# Compilation + +To compile the whole projet including case studies: + +``` bash +make +``` + +# Use in your coq development + +Suppose you have compiled Pactole in a directory calle `pactole` and +you want to use it in your own development called `myproject`. You +should use a `_CoqProject` file containing: + +``` +-Q pactole Pactole +.. # your own project options +``` + +and use `coq_makefile -f _CoqProject -o Makefile` to generate your +makefile. Note that ProofGeneral and other IDE also read the +`_CoqProject`. + +From now on you can use Pactole in your coq files by importing the +relevant modules, e.g.: + +``` coq +Require Import Pactole.Core.Formalism. +``` + +See the file README.md for the directory structure of the library. + +# NOTES FOR DEVELOPPERS + +Pactole provides fast (and unsafe) compilation target. This is +explained in this section. + + +## build (slow) + +The default makefile target may take more than 10mn to compile due to +some case studies. You may want to comment them in the _CoqProject +file before the configure step above if you are not interested in +these case studies. Another solution is to follow the fast compilation +recipes below. + +Doing this once a week when developing is good practice. This is the +only way to make a really safe compilation including all +xxx_Assumption.v. You should always do this once in a while to make +sure some universe constraints aren't failing and to check if you did +not miss any remaining axiom. + +But during the development process this makes the compilation when +switching between files too slow. Hence the following less safe +compilation processes: + +## unsafe builds when developing + +During development you can benefit from coq's "vos/vok" generation to +speed up compilation. The only specificity of Pactole concerns +compilation of the files named xxx_Assumptions.v. This files print the +sets of assumptions used in the proofs of final lemmas in Case +studies. Compiling these files for vos/vok target would raise errors. +We provide adapted targets below. + +## build (slow, almost safe, but very parallelisable) + +``` +make [-j] vok +``` +or better when developing +``` +make [-j] vok-nocheck +``` +Doing this once a day when developing is good practice. + +## build (Very fast, unsafe) + +``` +make [-j] vos +``` +or better when developing +``` +make [-j] vos-nocheck +``` + +This should be your prefered way of compiling when developing. It is +much faster. It is unsafe but in most situations no bad surprise is to +be expected. + +You should do real compilations from time to time as explained above. + +## Proofgeneral + +For easy use of this feature (vos) you can use the auto compilation +feature of proofgeneral. menu: + +menu: Coq / Auto Compilation / Compile Before Require +and then: Coq / Auto Compilation / vos compilation + +Now you can transparently script in any buffer, all needed file will +be compiled quickly. Don't forget to make a big fat "make" from time +to time. diff --git a/Makefile.local b/Makefile.local index af1a23b9adc724abe32667d8eb30153e961233e0..060b3d563d3f0159cf4d3dbcd34f7f7223355165 100644 --- a/Makefile.local +++ b/Makefile.local @@ -26,5 +26,31 @@ vok-nocheck: $(VOKFILESNOASSUMPTION) # specified in _coqProject because spaces in args are not supported by # coq_makefile. # Remove to see where we should add "Proof using annotations". -COQEXTRAFLAGS+= -set "Suggest Proof Using=true" +# Comment this before pushing, until we drop support of coq <=8.13.2, +# because these older versions of coq need "true" instead of "yes" +# COQEXTRAFLAGS+= -set "Suggest Proof Using=yes" +# developpers could use this. +# COQEXTRAFLAGS+= -set "Suggest Proof Using=yes" -w -deprecated-instance-without-locality + +COQEXTRAFLAGS+= -set "Suggest Proof Using=yes" +#COQEXTRAFLAGS+= -w -deprecated-instance-without-locality +#COQEXTRAFLAGS+= -w -intuition-auto-with-star +#COQEXTRAFLAGS+= -w -argument-scope-delimiter +COQEXTRAFLAGS+= -w -deprecated-syntactic-definition +#COQEXTRAFLAGS+= -w -future-coercion-class-field + + +core: Pactole_all.vo + +# make ../foo.check Ask to compile foo.v id necessary and then coqchk foo.vo +%.check: %.vo + $(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS_NOML) $< + +# make ../foo.v displays the dependencies of foo.v +%.dep: %.v + $(TIMER) $(COQDEP) $(COQLIBS) $< + +recore: + dev/rebuild_pactole_all.sh > Pactole_all.v + make Pactole_all.vo diff --git a/Models/ContinuousGraph.v b/Models/ContinuousGraph.v index 986662f39d84c5e8d225ba78362bfba18ebfac9b..dd141bbc2c21b315755ab3a47a13a9217e9eb63e 100644 --- a/Models/ContinuousGraph.v +++ b/Models/ContinuousGraph.v @@ -17,15 +17,10 @@ *) (**************************************************************************) -Require Import Utf8_core. -Require Import Arith_base. -Require Import Reals. -Require Import Lia. -Require Import Psatz. -Require Import SetoidList. -Require Import Pactole.Setting. -Require Import Pactole.Spaces.Graph. -Require Import Pactole.Spaces.Isomorphism. +Require Import Utf8_core Arith_base Reals Lia Psatz SetoidList. +From Pactole Require Import Setting Util.SetoidDefs Util.Bijection. +From Pactole Require Import Spaces.Graph Spaces.Isomorphism + Spaces.ThresholdIsomorphism. Require Import Pactole.Models.Flexible. @@ -36,8 +31,8 @@ Remove Hints eq_setoid : typeclass_instances. Section CGF. Context {V E : Type}. -Context `{Names}. -Context {G : Graph V E}. +Context {N : Names}. +Context {TG : ThresholdGraph V E}. Instance LocationV : Location := { location := V }. @@ -53,6 +48,9 @@ Global Instance E_src_tgt_thd_EqDec : EqDec E_src_tgt_thd_Setoid := Global Instance E_subrelation : subrelation (@equiv E E_Setoid) (@equiv E E_src_tgt_thd_Setoid). Proof using . intros ? ? Heq. split; simpl; now rewrite Heq. Qed. +Global Instance E_subrelation' : subrelation (@equiv E E_src_tgt_thd_Setoid) (@equiv E E_Setoid). +Proof using . intros ?? [H _]. apply simple_graph, H. Qed. + Global Instance src_compat : Proper (equiv ==> equiv) src. Proof using . intros ? ? Heq. apply Heq. Qed. @@ -64,14 +62,14 @@ Proof using . intros ? ? Heq. apply Heq. Qed. (* Since iso_E gives a bijection that comes with its setoid, we need to be lower level to change it from [E_Setoid] to [E_src_tgt_thd_Setoid]. *) -Global Instance iso_E_compat : forall iso, - Proper (equiv ==> equiv) (iso_E iso). +Global Instance iso_E_compat : + ∀ iso : threshold_isomorphism TG, Proper (equiv ==> equiv) (iso_E iso). Proof using . -intros iso ? ? [[Hsrc Htgt] Hthd]. -repeat split; unfold equiv in *; cbn -[equiv] in *. -- now rewrite <- 2 (proj1 (iso_morphism _ _)), Hsrc. -- now rewrite <- 2 (proj2 (iso_morphism _ _)), Htgt. -- rewrite <- 2 iso_threshold. now f_equiv. + intros iso ? ? [[Hsrc Htgt] Hthd]. + repeat split; unfold equiv in *; cbn -[equiv] in *. + - now rewrite <- 2 (proj1 (iso_morphism _ _)), Hsrc. + - now rewrite <- 2 (proj2 (iso_morphism _ _)), Htgt. + - rewrite <- iso_threshold, Hthd. apply iso_threshold. Qed. @@ -87,11 +85,11 @@ simple refine {| equiv := fun l l' => | OnEdge e p, OnEdge e' p' => e == e' /\ p == p' | _, _ => False end |}; autoclass; []. -Proof. split. -+ now intros []. -+ intros [] [] Heq; simpl in *; decompose [False and] Heq; repeat split; now symmetry. -+ intros [] [] [] Heq1 Heq2; simpl in *; - decompose [False and] Heq1; decompose [False and] Heq2; repeat split; etransitivity; eauto. +Proof using . split. + + now intros []. + + intros [] [] Heq; simpl in *; decompose [False and] Heq; repeat split; now symmetry. + + intros [] [] [] Heq1 Heq2; simpl in *; + decompose [False and] Heq1; decompose [False and] Heq2; repeat split; etransitivity; eauto. Defined. Global Instance locG_EqDec: EqDec locG_Setoid. @@ -115,53 +113,32 @@ Global Instance OnEdge_compat : Proper (equiv ==> equiv ==> equiv) OnEdge. Proof using . repeat intro. auto. Qed. (** We can use an isomorphism to build a bijection on a continuous graph. *) -Definition bijectionG (iso : isomorphism G) : Bijection.bijection loc. -simple refine {| Bijection.section := fun pt => match pt with +Definition bijectionG (iso : threshold_isomorphism TG) : bijection loc. +simple refine {| section := fun pt => match pt with | OnVertex v => OnVertex (iso.(iso_V) v) | OnEdge e p => OnEdge (iso.(iso_E) e) p end; - Bijection.retraction := fun pt => match pt with - | OnVertex v => OnVertex (Bijection.retraction iso.(iso_V) v) - | OnEdge e p => OnEdge (Bijection.retraction iso.(iso_E) e) p + retraction := fun pt => match pt with + | OnVertex v => OnVertex (retraction iso.(iso_V) v) + | OnEdge e p => OnEdge (retraction iso.(iso_E) e) p end |}. -Proof. -* intros [| e1 p1] [| e2 p2] Heq; simpl in Heq; trivial. - + now repeat f_equiv. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split; simpl. - - rewrite <- (proj1 (iso_morphism iso e1)), Hsrc. apply iso_morphism. - - rewrite <- (proj2 (iso_morphism iso e1)), Htgt. apply iso_morphism. - - now rewrite <- 2 iso_threshold, Hthd. - - assumption. -* intros [| e1 p1] [| e2 p2] ; simpl in *; try tauto; [|]; split; intro Heq. - + rewrite <- Heq. apply Bijection.retraction_section. - + rewrite <- Heq. apply Bijection.section_retraction. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split. - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (proj1 (iso_morphism _ e2)). simpl. - now rewrite <- (Bijection.Inversion iso), (proj1 (iso_morphism _ e1)). - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (proj2 (iso_morphism _ e2)). simpl. - now rewrite <- (Bijection.Inversion iso), (proj2 (iso_morphism _ e1)). - - change (Bijection.retraction (iso_E iso)) with (Bijection.section (iso_E (inverse iso))). - rewrite <- (iso_threshold _ e2), <- Hthd, iso_threshold. - simpl. now rewrite Bijection.retraction_section. - - auto. - + destruct Heq as [[[Hsrc Htgt] Hthd] Hp]. repeat split. - - now rewrite <- (proj1 (iso_morphism _ e1)), <- Hsrc, (proj1 (iso_morphism _ _)), - Bijection.section_retraction. - - now rewrite <- (proj2 (iso_morphism _ e1)), <- Htgt, (proj2 (iso_morphism _ _)), - Bijection.section_retraction. - - now rewrite <- (iso_threshold _ e1), <- Hthd, (iso_threshold _ _), - Bijection.section_retraction. - - auto. +Proof using . + * intros [l1 | e1 p1] [l2 | e2 p2] H. 2,3: inversion H. + + cbn in H. rewrite H. reflexivity. + + destruct H as [H1 H2]. rewrite H1, H2. reflexivity. + * intros [l1 | e1 p1] [l2 | e2 p2]. all: split. all: intros H. 3-6: inversion H. + + cbn in H. rewrite <- H, retraction_section. reflexivity. + + cbn in H. rewrite <- H, section_retraction. reflexivity. + + destruct H as [H1 H2]. erewrite (proj1 (Inversion _ _ _)). + rewrite H2. reflexivity. apply E_subrelation', H1. + + destruct H as [H1 H2]. erewrite (proj2 (Inversion _ _ _)). + rewrite H2. reflexivity. apply E_subrelation', H1. Defined. Global Instance bijectionG_compat : Proper (equiv ==> equiv) bijectionG. Proof using . -intros iso1 iso2 Hiso []; simpl. -+ apply Hiso. -+ repeat split; apply Graph.src_compat || apply Graph.tgt_compat - || apply Graph.threshold_compat; apply Hiso. + intros iso1 iso2 Hiso [l | e p]. apply Hiso. split. + apply E_subrelation, Hiso. reflexivity. Qed. (** ** Translation of locations *) @@ -169,23 +146,16 @@ Qed. Definition location_G2V (loc : locG) : locV := match loc with | OnVertex l => l - | OnEdge e p => if Rle_dec (threshold e) p then Graph.tgt e else Graph.src e + | OnEdge e p => if Rle_dec (threshold e) p then tgt e else src e end. Global Instance location_G2V_compat : Proper (equiv ==> equiv) location_G2V. Proof using . unfold location_G2V. intros [l1 | e1 p1] [l2 | e2 p2] Hxy; try tauto; []. -destruct Hxy as [Hexy Hpxy], - (Rle_dec (threshold e1) p1) eqn:Hx, - (Rle_dec (threshold e2) p2) eqn:Hy. -+ apply Hexy. -+ assert (Ht := proj2 Hexy). - assert (Hr : (threshold e1 <= p1)%R) by assumption. - now rewrite Ht, Hpxy in Hr. -+ assert (Hr : (threshold e2 <= p2)%R) by assumption. - assert (Ht := proj2 Hexy). - now rewrite <- Ht, <- Hpxy in Hr. -+ apply Hexy. +destruct Hxy as [[[Hsrc Htgt] Ht] Hpxy]. cbn in Ht. +destruct (Rle_dec (threshold e1) p1) as [Hx | Hx], (threshold e1), p1, + (Rle_dec (threshold e2) p2) as [Hy | Hy], (threshold e2), p2; +cbn in *; subst; tauto. Qed. Definition location_V2G : locV -> locG := OnVertex. @@ -227,8 +197,8 @@ Lemma valid_stateV_iso' : forall v e iso pt, pt == iso.(iso_V) v -> valid_stateV (v, e) -> valid_stateV (pt, iso.(iso_E) e). Proof using . intros v e iso pt Hpt [Hcase | Hcase]. -+ left. simpl in *. rewrite Hpt, Hcase. apply iso_morphism. -+ right. simpl in *. rewrite Hpt, Hcase. apply iso_morphism. ++ left. cbn in *. rewrite Hpt, Hcase. apply iso_morphism. ++ right. cbn in *. rewrite Hpt, Hcase. apply iso_morphism. Qed. (** In the continuous case, the state must also contain the destination of the robot. @@ -286,7 +256,7 @@ Proof using . intros ? ? []. unfold state_V2G. now split. Qed. Definition state_G2V (state : stateG) : stateV := match state with | SOnVertex v e p => exist valid_stateV (v, e) p - | SOnEdge e p => if Rle_dec (@threshold locV E G e) p + | SOnEdge e p => if Rle_dec (@threshold locV E TG e) p then exist valid_stateV (Graph.tgt e, e) ltac:(now right) else exist valid_stateV (Graph.src e, e) ltac:(now left) end. @@ -296,11 +266,11 @@ Proof using . intros [v e p | e p] [v' e' p' | e' p'] Hstate; auto; []. destruct Hstate as [[[Hsrc Htgt] Hthd] Hp]. simpl. destruct (Rle_dec (threshold e) p), (Rle_dec (threshold e') p'); -repeat split; simpl in *; rewrite ?Hsrc, ?Htgt, ?Hthd in *; try reflexivity; [|]; -destruct p, p'; simpl in *; subst; contradiction. +repeat split; cbn in *; rewrite ?Hsrc, ?Htgt, ?Hthd in *; try reflexivity; [|]; +destruct (threshold e), p, (threshold e'), p'; cbn in *; subst; tauto. Qed. -Lemma state_V2G2V : forall state, state_G2V (state_V2G state) == state. +Lemma state_V2G2V : ∀ state, state_G2V (state_V2G state) == state. Proof using . intro. simpl. repeat (split; try reflexivity). Qed. (** ** On configurations *) @@ -308,38 +278,32 @@ Proof using . intro. simpl. repeat (split; try reflexivity). Qed. (** The precondition for liftable changes of frame is that they must come from isomorphisms (which must not change the thresholds). *) Local Instance InfoV : @State LocationV stateV. -simple refine {| +Proof using . + simple refine {| get_location := fun state => fst (proj1_sig state); - state_Setoid := stateV_Setoid; - precondition := fun f => sigT (fun iso => f == iso.(iso_V) /\ iso_T iso == @Bijection.id R _); - lift := fun f state => exist _ (projT1 f (fst (proj1_sig state)), - iso_E (projT1 (projT2 f)) (snd (proj1_sig state))) _ |}; autoclass. -Proof. -+ abstract (destruct f as [f [iso [Hiso ?]]], state as [state [Hcase | Hcase]]; - cbn; left + right; rewrite Hiso, Hcase; cbn; apply iso_morphism). -+ (* lift_id *) - intros [iso [Hiso Ht]] [state Hstate]. simpl. split; try reflexivity; []. - change (src (snd state)) with (Datatypes.id (src (snd state))). - change (tgt (snd state)) with (Datatypes.id (tgt (snd state))). - rewrite (Hiso (src (snd state))), (Hiso (tgt (snd state))). - repeat split; symmetry; try apply iso_morphism; []. - now rewrite <- iso_threshold, Ht. -+ intros ? ? Heq. apply Heq. -+ (* lift_compat *) - intros [f [iso1 [Hiso1 Ht1]]] [g [iso2 [Hiso2 Ht2]]] Heq [] [] [Heq1 [Heq2 Heq3]]. - cbn in *. repeat split. - - now apply Heq. - - rewrite <- (proj1 (iso_morphism iso1 _)), <- Hiso1, - <- (proj1 (iso_morphism iso2 _)), <- Hiso2. - now apply Heq. - - rewrite <- (proj2 (iso_morphism iso1 _)), <- Hiso1, - <- (proj2 (iso_morphism iso2 _)), <- Hiso2. - now apply Heq. - - now rewrite <- 2 iso_threshold, Ht1, Ht2. + state_Setoid := stateV_Setoid; + precondition := λ f, sigT (λ iso : threshold_isomorphism TG, + f == iso.(iso_V) /\ iso_T iso == Bijection.id); + lift := λ f state, exist _ (projT1 f (fst (proj1_sig state)), + iso_E (projT1 (projT2 f)) (snd (proj1_sig state))) _ |}; autoclass. + + abstract (destruct f as [f [iso [Hiso ?]]], state as [state [Hcase | Hcase]]; + cbn; rewrite Hiso, Hcase; left + right; apply iso_morphism). + + intros ? ? Heq. apply Heq. + + (* lift_compat *) + intros [f [iso1 [Hiso1 Ht1]]] [g [iso2 [Hiso2 Ht2]]] Heq [] [] [Heq1 [Heq2 Heq3]]. + cbn in *. repeat split. + - now apply Heq. + - rewrite <- (proj1 (iso_morphism iso1 _)), <- Hiso1, + <- (proj1 (iso_morphism iso2 _)), <- Hiso2. + now apply Heq. + - rewrite <- (proj2 (iso_morphism iso1 _)), <- Hiso1, + <- (proj2 (iso_morphism iso2 _)), <- Hiso2. + now apply Heq. + - now rewrite <- 2 iso_threshold, Ht1, Ht2. Defined. Definition good_iso_of f iso := f == Bijection.section (bijectionG iso) - /\ iso_T iso == @Bijection.id R _. + /\ iso_T iso == Bijection.id. Definition preconditionG := fun f => sigT (good_iso_of f). Definition liftG (f : sigT preconditionG) state := @@ -408,21 +372,6 @@ simple refine {| precondition := preconditionG; lift := liftG |}. Proof. -* (* lift_id *) - intros [iso [Hiso Ht]] [v e proof | e p]; simpl. - + assert (Hv := Hiso (OnVertex v)). - assert (Hsrc := Hiso (OnVertex (src e))). - assert (Htgt := Hiso (OnVertex (tgt e))). - specialize (Ht (threshold e)). - simpl in Hv, Hsrc, Htgt, Ht. - rewrite <- (proj1 (iso_morphism iso e)), - <- (proj2 (iso_morphism iso e)), <- iso_threshold; auto. - + assert (Hsrc := Hiso (OnVertex (src e))). - assert (Htgt := Hiso (OnVertex (tgt e))). - specialize (Ht (threshold e)). - simpl in Hsrc, Htgt, Ht. - rewrite <- (proj1 (iso_morphism iso e)), - <- (proj2 (iso_morphism iso e)), <- iso_threshold; auto. * (* get_location_lift *) intros [f [iso [Hiso Ht]]] [v e proof | e p]; simpl; destruct_match_eq Hf; simpl in Hf; @@ -493,25 +442,24 @@ Proof using . intro. simpl. repeat split; reflexivity. Qed. (** ** Demonic schedulers **) (** Acceptable frame changes must not change the thresholds. *) -Definition stable_threshold iso := iso_T iso == @Bijection.id R _. +Definition stable_threshold iso := iso_T iso == Bijection.id. -Definition stable_threshold_inverse : forall iso, +Definition stable_threshold_inverse : ∀ iso, stable_threshold iso -> stable_threshold (inverse iso). Proof using . -intros iso Hstable x. unfold stable_threshold in *. simpl in *. -now rewrite <- (Hstable x), Bijection.retraction_section. + unfold stable_threshold. intros iso Hstable x. apply (injective (iso_T iso)). + rewrite Hstable at 2. apply section_retraction. Qed. (** Frame choice: graph isomorphisms not changing thresholds *) Global Instance FrameChoiceIsomorphismV : @frame_choice LocationV (sig stable_threshold) := {| - frame_choice_bijection := fun f => @iso_V locV E G (proj1_sig f); - frame_choice_Setoid := sig_Setoid (@isomorphism_Setoid locV E G); - frame_choice_bijection_compat := - fun f g => @iso_V_compat locV E G (proj1_sig f) (proj1_sig g) |}. + frame_choice_bijection := λ f : sig stable_threshold, @iso_V locV E TG (proj1_sig f); + frame_choice_Setoid := sig_Setoid (@threshold_isomorphism_Setoid locV E TG); + frame_choice_bijection_compat := ltac:(intros ?? H; apply H) |}. Global Instance FrameChoiceIsomorphismG : @frame_choice LocationG (sig stable_threshold) := {| frame_choice_bijection := fun f => bijectionG (proj1_sig f); - frame_choice_Setoid := sig_Setoid (@isomorphism_Setoid locV E G); + frame_choice_Setoid := sig_Setoid (@threshold_isomorphism_Setoid locV E TG); frame_choice_bijection_compat := fun f g => bijectionG_compat (proj1_sig f) (proj1_sig g) |}. (** The demon update choice only contains the movement ratio, either a boolean or a ratio. *) diff --git a/Models/Flexible.v b/Models/Flexible.v index 05794447abdc89ae3d1339b349142957b1e1ad1b..b5f48b35f9a6f4789ea7db9ea728cef29c3b0430 100644 --- a/Models/Flexible.v +++ b/Models/Flexible.v @@ -48,7 +48,7 @@ Instance Frame : frame_choice (similarity location) := FrameChoiceSimilarity. Class FlexibleChoice `{update_choice Tactive} := { move_ratio : Tactive -> ratio; - move_ratio_compat :> Proper (@equiv Tactive update_choice_Setoid ==> @equiv _ (sig_Setoid _)) move_ratio }. + #[global] move_ratio_compat :: Proper (@equiv Tactive update_choice_Setoid ==> @equiv _ (sig_Setoid _)) move_ratio }. (** Flexible moves are parametrized by a minimum distance [delta] that robots must move when they are activated. *) Class FlexibleSetting `{FlexibleChoice} @@ -101,11 +101,15 @@ Proof. intros config1 config2 Hconfig gg g ? sim1 sim2 Hsim traj1 traj2 Htraj Ï1 Ï2 HÏ. subst gg. assert (Heq : get_location (traj1 Ï1) == get_location (traj2 Ï2)). { apply get_location_compat. now f_equiv. } -destruct_match_eq Hle; destruct_match_eq Hle'; rewrite Heq, ?Hsim in *; -solve [ reflexivity - | now f_equiv - | rewrite Hconfig, Htraj, HÏ in *; now rewrite Hle in Hle' ]. -Unshelve. all:autoclass. +destruct_match_eq Hle; destruct_match_eq Hle'. ++ now f_equiv. ++ exfalso. apply Bool.diff_false_true. rewrite <- Hle, <- Hle'. f_equal. + - f_equal. now f_equiv. + - apply dist_compat; apply get_location_compat; auto. ++ exfalso. apply Bool.diff_false_true. rewrite <- Hle, <- Hle'. f_equal. + - f_equal. now f_equiv. + - apply dist_compat; apply get_location_compat; auto. ++ apply Htraj. Defined. Global Instance FlexibleChoiceFlexibleUpdate delta : FlexibleSetting (Update := FlexibleUpdate delta) delta. diff --git a/Models/GraphEquivalence.v b/Models/GraphEquivalence.v index ed7bed142770ae8be14f9ff107f0741d9c12d553..7236135f150823c3ff07ab604c51ba49f3c833d9 100644 --- a/Models/GraphEquivalence.v +++ b/Models/GraphEquivalence.v @@ -17,14 +17,62 @@ Require Import Reals. Require Import Psatz. Require Import SetoidList. +Require Import Pactole.Util.Bijection. Require Import Pactole.Setting. Require Import Pactole.Spaces.Graph. Require Import Pactole.Spaces.Isomorphism. +Require Import Pactole.Spaces.ThresholdIsomorphism. Require Import Pactole.Models.ContinuousGraph. Require Import Pactole.Models.DiscreteGraph. - -Typeclasses eauto := (bfs). +#[local] +Hint Extern 9 (_ = _ :>nat) => lia: zarith. +#[local] +Hint Extern 9 (_ <= _) => lia: zarith. +#[local] +Hint Extern 9 (_ < _) => lia: zarith. +#[local] +Hint Extern 9 (_ >= _) => lia: zarith. +#[local] +Hint Extern 9 (_ > _) => lia: zarith. + +#[local] +Hint Extern 9 (_ <> _ :>nat) => lia: zarith. +#[local] +Hint Extern 9 (~ _ <= _) => lia: zarith. +#[local] +Hint Extern 9 (~ _ < _) => lia: zarith. +#[local] +Hint Extern 9 (~ _ >= _) => lia: zarith. +#[local] +Hint Extern 9 (~ _ > _) => lia: zarith. + +#[local] +Hint Extern 9 (_ = _ :>Z) => lia: zarith. +#[local] +Hint Extern 9 (_ <= _)%Z => lia: zarith. +#[local] +Hint Extern 9 (_ < _)%Z => lia: zarith. +#[local] +Hint Extern 9 (_ >= _)%Z => lia: zarith. +#[local] +Hint Extern 9 (_ > _)%Z => lia: zarith. + +#[local] +Hint Extern 9 (_ <> _ :>Z) => lia: zarith. +#[local] +Hint Extern 9 (~ (_ <= _)%Z) => lia: zarith. +#[local] +Hint Extern 9 (~ (_ < _)%Z) => lia: zarith. +#[local] +Hint Extern 9 (~ (_ >= _)%Z) => lia: zarith. +#[local] +Hint Extern 9 (~ (_ > _)%Z) => lia: zarith. + +#[local] +Hint Extern 9 False => lia: zarith. + +(* Typeclasses eauto := (bfs). *) Open Scope R_scope. @@ -32,7 +80,7 @@ Section GraphEquivalence. Context (V E : Type). Context {NN : Names}. -Context {G : Graph V E}. +Context {TG : ThresholdGraph V E}. (** We assume that the graph contains loops from each node to itself. *) Variable self_loop : V -> E. @@ -43,17 +91,17 @@ Ltac Rdec := repeat match goal with | |- context[Rdec ?x ?x] => let Heq := fresh "Heq" in destruct (Rdec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] + [clear Heq | exfalso; contradiction Heq; reflexivity] | |- context[Rdec 1 0] => let Heq := fresh "Heq" in destruct (Rdec 1 0) as [Heq | Heq]; - [now elim R1_neq_R0 | clear Heq] + [now contradiction R1_neq_R0 | clear Heq] | |- context[Rdec 0 1] => let Heq := fresh "Heq" in destruct (Rdec 0 1) as [Heq | Heq]; - [now symmetry in Heq; elim R1_neq_R0 | clear Heq] + [now symmetry in Heq; contradiction R1_neq_R0 | clear Heq] | H : context[Rdec ?x ?x] |- _ => let Heq := fresh "Heq" in destruct (Rdec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] - | H : ?x <> ?x |- _ => elim H; reflexivity + [clear Heq | exfalso; contradiction Heq; reflexivity] + | H : ?x <> ?x |- _ => contradiction H; reflexivity end. Existing Instance InfoV. @@ -115,16 +163,28 @@ Proof. exists (inverse (proj1_sig (change_frame da (config_G2V config) g))). split; try reflexivity; []. cbn -[equiv]. apply stable_threshold_inverse, proj2_sig. + intros config1 config2 Hconfig gg g ?. subst gg. now rewrite Hconfig. -+ intros config1 config2 Hconfig gg g ? traj1 traj2 Htraj. subst gg. now rewrite Hconfig, Htraj. ++ intros config1 config2 Hconfig gg g ? traj1 traj2 Htraj. subst gg. + destruct_match_eq Heq1; destruct_match_eq Heq2; try reflexivity; [|]; + exfalso; apply Bool.diff_false_true; rewrite <- Heq1, <- Heq2; [ symmetry |]; + apply (choose_update_compat da (config_G2V_compat _ _ Hconfig) (eq_refl g) _ _ Htraj). + intros config1 config2 Hconfig id1 id2 Hid. simpl in Hid. subst id1. - now rewrite Hconfig. + destruct_match_eq Heq1; destruct_match_eq Heq2; try reflexivity; [|]; + exfalso; apply Bool.diff_false_true; rewrite <- Heq1, <- Heq2; [ symmetry |]; + apply (choose_inactive_compat da (config_G2V_compat _ _ Hconfig) (eq_refl id2)). Defined. Instance da_D2C_compat : Proper (equiv ==> equiv) da_D2C. Proof using . intros da1 da2 Hda. split; [| split; [| split; [| split]]]; cbn -[equiv]. + intro. now apply activate_da_compat. -+ intros. apply relocate_byz_D2C_compat; reflexivity || now apply relocate_byz_da_compat. ++ intros. + apply relocate_byz_D2C_compat. + * red;intros. + red;intros. + rewrite Hda,H,H0. + reflexivity. + * reflexivity. + * reflexivity. + intros config g. apply (change_frame_da_compat Hda); auto. + intros id rc traj. erewrite (choose_update_da_compat Hda); auto. + intros config id. erewrite (choose_inactive_da_compat Hda); auto. @@ -147,24 +207,24 @@ simple refine {| end in let e := snd (proj1_sig (config id)) in Rle_bool (threshold e) (current_ratio + da.(choose_inactive) (config_V2G config) id) |}; -try exact G; autoclass. -Proof. -+ intros config g. exists (proj1_sig (change_frame da (config_V2G config) g)). - split; try reflexivity; []. apply proj2_sig. -+ intros config g. exists (inverse (proj1_sig (change_frame da (config_V2G config) g))). - split; try reflexivity; []. apply stable_threshold_inverse, proj2_sig. -+ intros config1 config2 Hconfig gg g ?. subst gg. now rewrite Hconfig. -+ intros config1 config2 Hconfig gg g ? pt1 pt2 Hpt. - f_equiv; try apply Hpt; []. - f_equiv. now apply (choose_update_compat da); f_equiv. -+ intros config1 config2 Hconfig id1 id2 Hid. simpl in Hid. subst id1. - assert (Hpt := Hconfig id2). - destruct Hpt as [Hpt [[Hsrc Htgt] Hthd]], - (proj1_sig (config1 id2)) as [pt1 e1], - (proj1_sig (config2 id2)) as [pt2 e2]. - simpl in Hpt, Hsrc, Htgt, Hthd. - destruct_match; simpl; rewrite Hthd; do 3 f_equiv; - apply (choose_inactive_compat da); trivial; reflexivity. +try exact TG; autoclass. +Proof using . + + intros config g. exists (proj1_sig (change_frame da (config_V2G config) g)). + split; try reflexivity; []. apply proj2_sig. + + intros config g. exists (inverse (proj1_sig (change_frame da (config_V2G config) g))). + split; try reflexivity; []. apply stable_threshold_inverse, proj2_sig. + + intros config1 config2 Hconfig gg g ?. subst gg. now rewrite Hconfig. + + intros config1 config2 Hconfig gg g ? pt1 pt2 Hpt. + subst. + f_equiv. + * f_equiv. + now rewrite Hpt. + * f_equiv. + rewrite Hconfig. + now setoid_rewrite Hpt. + + intros config1 config2 Hconfig id1 id2 Hid. simpl in Hid. subst id1. + destruct (Cconfig id2) as [v1 e1 p1 | e1 p1]; cbn -[ratio_0]; + now do 3 f_equiv; rewrite Hconfig. Defined. Instance da_C2D_compat : Proper (equiv ==> equiv ==> equiv) da_C2D. @@ -219,7 +279,7 @@ intros ? ? Hconfig ? ? ? ? ? Hframe ? ? Htarget ? ? Hchoice. simpl in Hchoice. subst. cbn zeta. repeat destruct_match; solve [ simpl; repeat split; apply Htarget - | match goal with | H : complement _ _ _ |- _ => elim H end; + | match goal with | H : complement _ _ _ |- _ => contradiction H end; simpl in Htarget; now (rewrite <- Hconfig, <- (proj1 (proj1 Htarget))) || (rewrite Hconfig, (proj1 (proj1 Htarget))) | apply Hconfig ]. @@ -245,9 +305,9 @@ Defined. Instance add_edge_compat : Proper (equiv ==> equiv ==> equiv ==> equiv) add_edge. Proof using . intros ? ? He Ï1 Ï1' HÏ1 Ï2 Ï2' HÏ2. unfold add_edge. -Time repeat destruct_match; solve [ rewrite HÏ1, HÏ2 in *; contradiction - | rewrite <- HÏ1, <- HÏ2 in *; contradiction - | simpl; rewrite ?HÏ1, ?HÏ2; repeat split; apply He ]. +repeat destruct_match; +try solve [hnf; split; apply He | apply proj_ratio_compat in HÏ1, HÏ2; congruence]; []. +split; trivial; []. cbn. apply proj_ratio_compat in HÏ1, HÏ2. congruence. Qed. (** Move by a ratio [Ï] from the state [state]. *) @@ -264,9 +324,9 @@ intros [v1 e1 proof1 | e1 p1] [v2 e2 proof2 | e2 p2] Heq Ï1 Ï2 HÏ; simpl in H + unfold move. destruct Heq as [Hv [[Hsrc Htgt] Hthd]]. do 2 destruct_match. - now f_equiv. - - match goal with H : complement _ _ _ |- _ => elim H end. + - match goal with H : complement _ _ _ |- _ => contradiction H end. now rewrite <- Hv, <- Hsrc. - - match goal with H : complement _ _ _ |- _ => elim H end. + - match goal with H : complement _ _ _ |- _ => contradiction H end. now rewrite Hv, Hsrc. - simpl. tauto. + unfold move. f_equiv; auto; []. now destruct p1, p2. @@ -278,14 +338,24 @@ Proof. repeat intro. subst. now f_equiv. Defined. Instance UpdateG : @update_function _ _ _ _ _ (sig stable_threshold) ratio _ FrameChoiceIsomorphismG _. -refine {| - update := fun (config : CGF_config) g frame target Ï => - match config (Good g) with - | SOnVertex v e proof => - if v =?= src target then move (SOnVertex v target ltac:(now left)) Ï else config (Good g) - | SOnEdge e p => config (Good g) - end |}. Proof. + simpl in *. + unshelve refine {| + update := fun (config : CGF_config) g frame target Ï => + match config (Good g) with + | SOnVertex v e proof => if v =?= src target then _ else config (Good g) + | SOnEdge e p => config (Good g) + end + |}; try typeclasses eauto. + + { eapply move. + 2:exact Ï. + apply (SOnVertex v target). + left. + cbn. + simpl in *. + assumption. } + intros config1 config2 Hconfig gg g ? iso1 iso2 Hframe target1 target2 Htarget Ï1 Ï2 HÏ. subst gg. assert (Hsrc := src_compat _ _ Htarget). assert (Htgt := tgt_compat _ _ Htarget). @@ -296,9 +366,9 @@ try destruct (v1 =?= src target1) as [Hsrc1 | Hsrc1], (v2 =?= src target2) as [Hsrc2 | Hsrc2]. + f_equiv; trivial; []. simpl. rewrite Hsrc1, Hsrc2. repeat split; auto; apply Htarget. -+ match goal with H : complement _ _ _ |- _ => elim H end. ++ match goal with H : complement _ _ _ |- _ => contradiction H end. now rewrite <- Hsrc, <- (proj1 Hconfig). -+ match goal with H : complement _ _ _ |- _ => elim H end. ++ match goal with H : complement _ _ _ |- _ => contradiction H end. now rewrite Hsrc, (proj1 Hconfig). + assumption. + tauto. @@ -371,25 +441,27 @@ simpl activate. destruct_match. * cbn -[equiv precondition_satisfied]. rewrite Hconfig. simpl snd. transitivity (iso_E Ciso e); [| transitivity (iso_E Diso e)]. + f_equiv. - + apply E_subrelation, (Isomorphism.iso_E_compat Hiso e). + + apply E_subrelation, Hiso. + repeat split. - unfold equiv. cbn -[equiv]. now rewrite <- (proj1 (iso_morphism Diso e)), <- HDiso, (proj1 (iso_morphism _ e)). - unfold equiv. cbn -[equiv]. now rewrite <- (proj2 (iso_morphism Diso e)), <- HDiso, (proj2 (iso_morphism _ e)). - - unfold equiv. cbn -[equiv]. - rewrite <- 2 iso_threshold. unfold Diso. rewrite (proj2_sig Dframe_choice). - destruct (precondition_satisfied da config g) as [? [? Ht]]. simpl. now rewrite Ht. } + - cbn. rewrite <- 2 iso_threshold. unfold Diso. rewrite (proj2_sig Dframe_choice). + destruct (precondition_satisfied da config g) as [? [? Ht]]. simpl. now rewrite Ht. } assert (Hlocal_state : Clocal_state == state_V2G Dlocal_state). { unfold Clocal_state. rewrite Hlocal_config. reflexivity. } assert (Hobs : Cobs == Dobs). { unfold Cobs, Dobs. unfold obs_from_config at 1. unfold obs_V2G. rewrite Hlocal_config, Hlocal_state. reflexivity. } assert (Hlocal_robot_decision : Clocal_robot_decision == Dlocal_robot_decision). - { unfold Dlocal_robot_decision. cbn -[equiv]. rewrite Hobs. reflexivity. } + { unfold Dlocal_robot_decision. cbn -[equiv]. now apply (pgm_compat rbg). } assert (Hchoice : Cchoice == if Dchoice then ratio_1 else ratio_0). - { cbn -[equiv]. unfold Dchoice. - rewrite Hlocal_config, config_V2G2V, Hobs. reflexivity. } + { cbn -[equiv]. unfold Dchoice, Dlocal_robot_decision. + apply config_G2V_compat in Hlocal_config. rewrite config_V2G2V in Hlocal_config. + destruct_match_eq Heq1; destruct_match_eq Heq2; reflexivity || exfalso; + apply Bool.diff_false_true; rewrite <- Heq1, <- Heq2; [ symmetry |]; + apply (choose_update_compat da Hlocal_config (eq_refl g) _ _ (pgm_compat rbg _ _ Hobs)). } assert (Hnew_local_state : Cnew_local_state == state_V2G Dnew_local_state). { unfold Cnew_local_state, Dnew_local_state. unfold update, UpdateG, UpdateV. assert (Hlocal_g := Hlocal_config (Good g)). unfold config_V2G in Hlocal_g. @@ -402,7 +474,8 @@ simpl activate. destruct_match. (v =?= src Dlocal_robot_decision) as [Hv | Hv]. + (* valid case: the robot chooses an adjacent edge *) unfold move. destruct_match; try contradiction; []. - rewrite Hchoice, Hlocal_robot_decision. destruct Dchoice. + rewrite (add_edge_compat Hlocal_robot_decision (reflexivity ratio_0) Hchoice). + destruct Dchoice. - (* the demon lets the robot move *) unfold add_edge. simpl equiv_dec. destruct ((0 + 1)%R =?= 0%R); try (simpl in *; lra); []. @@ -412,13 +485,13 @@ simpl activate. destruct_match. - (* the demon does not let the robot move *) unfold add_edge. simpl equiv_dec. destruct ((0 + 0)%R =?= 0%R); - try (match goal with H : complement _ _ _ |- _ => elim H; simpl in *; lra end); []. + try (match goal with H : complement _ _ _ |- _ => contradiction H; simpl in *; lra end); []. simpl. repeat split; reflexivity. + (* absurd case: the robot does not make the same choice *) - match goal with | H : complement _ _ _ |- _ => elim H end. + match goal with | H : complement _ _ _ |- _ => contradiction H end. rewrite <- Heqv, Hv'. apply Hlocal_robot_decision. + (* absurd case: the robot does not make the same choice *) - match goal with | H : complement _ _ _ |- _ => elim H end. + match goal with | H : complement _ _ _ |- _ => contradiction H end. rewrite Heqv, Hv. symmetry. apply Hlocal_robot_decision. + (* invalid case: the robot does not choose an adjacent edge *) simpl. repeat split; apply Heqv || apply Heqe. } @@ -443,13 +516,13 @@ simpl activate. destruct_match. unfold liftG. cbn [projT2]. repeat split. - rewrite HCiso'. cbn. f_equiv. symmetry. apply Hiso. - unfold equiv. cbn -[equiv precondition_satisfied_inv]. - Time setoid_rewrite <- (proj1 (iso_morphism _ e)). - Time setoid_rewrite HCiso'. + setoid_rewrite <- (proj1 (iso_morphism _ e)). + setoid_rewrite HCiso'. transitivity (inverse Diso (src e)); try apply HDiso'; []. f_equiv. apply inverse_compat. now symmetry. - unfold equiv. cbn -[equiv precondition_satisfied_inv]. - Time setoid_rewrite <- (proj2 (iso_morphism _ e)). - Time setoid_rewrite HCiso'. + setoid_rewrite <- (proj2 (iso_morphism _ e)). + setoid_rewrite HCiso'. transitivity (inverse Diso (tgt e)); try apply HDiso'; []. f_equiv. apply inverse_compat. now symmetry. - hnf. rewrite <- 2 iso_threshold. @@ -479,7 +552,7 @@ simpl activate. destruct_match. - (* the demon chooses not to let the robot move *) unfold add_edge. simpl. destruct ((0 + 0)%R =?= 0%R); - try (match goal with H : complement _ _ _ |- _ => elim H; simpl in *; lra end); []. + try (match goal with H : complement _ _ _ |- _ => contradiction H; simpl in *; lra end); []. repeat split; assumption || reflexivity. + unfold valid_stateV in *. simpl in Hvalid. destruct Hvalid as [ | Hvalid]; try contradiction; []. @@ -551,9 +624,9 @@ simpl activate. destruct_match_eq Hactive. assert (HDiso : projT1 (precondition_satisfied (da_C2D da config) (config_G2V config) g) == Diso). { destruct (precondition_satisfied (da_C2D da config) (config_G2V config) g) as [Diso' [HDf HDt]] eqn:HDiso. simpl projT1. - pose (iso_OKV := fun f (iso : @isomorphism (@location LocationV) E G) => + pose (iso_OKV := fun f (iso : @threshold_isomorphism (@location LocationV) E TG) => frame_choice_bijection f == iso.(iso_V) - /\ iso_T iso == @Bijection.id R R_Setoid). + /\ iso_T iso == Bijection.id). change (projT1 (existT (iso_OKV _) Diso' (conj HDf HDt)) == proj1_sig Dframe_choice). fold (iso_OKV (change_frame (da_C2D da config) (config_G2V config) g)) in HDiso. change (precondition_satisfied (da_C2D da config) (config_G2V config) g = @@ -573,25 +646,21 @@ simpl activate. destruct_match_eq Hactive. simpl state_G2V. repeat split. + simpl fst. rewrite Hframe. unfold config_G2V. transitivity (Ciso v); reflexivity || now symmetry; apply (HCiso (OnVertex v)). - + cbn -[precondition_satisfied]. - rewrite <- 2 (proj1 (iso_morphism _ _)), - (iso_V_compat HDiso (src e)), - (iso_V_compat Hiso (src e)). - symmetry. apply (HCiso (OnVertex (src e))). - + cbn -[precondition_satisfied]. - rewrite <- 2 (proj2 (iso_morphism _ _)), - (iso_V_compat HDiso (tgt e)), - (iso_V_compat Hiso (tgt e)). - symmetry. apply (HCiso (OnVertex (tgt e))). + + cbn-[precondition_satisfied]. rewrite <- 2 (proj1 (iso_morphism _ _)). + etransitivity. 2: symmetry; apply (HCiso (OnVertex (src e))). + etransitivity. 2: apply Hiso. apply HDiso. + + cbn-[precondition_satisfied]. rewrite <- 2 (proj2 (iso_morphism _ _)). + etransitivity. 2: symmetry; apply (HCiso (OnVertex (tgt e))). + etransitivity. 2: apply Hiso. apply HDiso. + cbn -[precondition_satisfied]. rewrite <- 2 iso_threshold. now rewrite (proj2 (projT2 (precondition_satisfied da config g))), (proj2 (projT2 (precondition_satisfied (da_C2D da config) (config_G2V config) g))). * (* OnEdge *) simpl liftG. simpl state_G2V. - assert (Heq : threshold ((iso_E (projT1 (precondition_satisfied da config g))) e) = threshold e). + assert (Heq : threshold ((iso_E (projT1 (precondition_satisfied da config g))) e) == threshold e). { now rewrite <- iso_threshold, (proj2 (projT2 (precondition_satisfied da config g))). } - destruct (Rle_dec (threshold e) p); - destruct_match; try (rewrite <- Heq in *; contradiction); [|]. + destruct (Rle_dec (threshold e) p); destruct_match; repeat rewrite proj_ratio_strict_ratio in *; + try (hnf in Heq; rewrite <- Heq in *; contradiction); [|]. + (* we are after the threshold, g is seen at the target of the edge *) cbn -[precondition_satisfied]. repeat split. - rewrite <- (proj2 (iso_morphism _ _)). @@ -634,15 +703,17 @@ simpl activate. destruct_match_eq Hactive. { unfold Cobs, Dobs. unfold obs_from_config at 2. unfold obs_V2G. rewrite Hlocal_config, Hlocal_state. reflexivity. } assert (Hlocal_robot_decision : Dlocal_robot_decision == Clocal_robot_decision). - { unfold Dlocal_robot_decision. cbn -[equiv]. rewrite Hobs. reflexivity. } + { unfold Clocal_robot_decision, Dlocal_robot_decision. now apply (pgm_compat rbg). } assert (Hchoice : Dchoice == if Rle_dec (threshold Clocal_robot_decision) Cchoice then true else false). - { unfold Dchoice, choose_update, da_C2D. - rewrite Hlocal_config, Hchoose_update. rewrite Hlocal_robot_decision at 2. - assert (Hthd := proj2 Hlocal_robot_decision). hnf in Hthd. rewrite Hthd. - destruct (Rle_dec (threshold Clocal_robot_decision) Cchoice) as [Hle | Hlt]. - - rewrite <- Rle_bool_true_iff in Hle. apply Hle. - - rewrite <- Rle_bool_false_iff in Hlt. apply Hlt. } + { unfold Dchoice, choose_update, da_C2D, Rle_bool. + do 2 destruct_match; try reflexivity; [|]; exfalso; + revert_one not; revert_one Rle; repeat rewrite proj_ratio_strict_ratio; + assert (Heq := proj1_sig_compat equiv_dec _ _ (threshold_compat _ _ Hlocal_robot_decision)); + hnf in Heq; rewrite Heq, Hlocal_config; + rewrite (proj_ratio_compat _ _ (Hchoose_update Clocal_config g Dlocal_robot_decision)), + (proj_ratio_compat _ _ (choose_update_compat da (reflexivity _) (eq_refl g) _ _ Hlocal_robot_decision)); + unfold Cchoice; congruence. } assert (HCiso' : forall v, projT1 (precondition_satisfied_inv da config g) v == inverse Ciso v). { destruct (precondition_satisfied_inv da config g)as [Ciso' [HCf HCt]]. simpl projT1. intro v. @@ -652,9 +723,17 @@ simpl activate. destruct_match_eq Hactive. == iso_E (inverse Ciso) e). { destruct (precondition_satisfied_inv da config g)as [Ciso' [HCf HCt]]. simpl projT1. intro e. - cut (bijectionG Ciso' (OnEdge e (1 /sr 2)) == bijectionG (inverse Ciso) (OnEdge e (1 /sr 2))); - try (now intros [Heq _]); []. - rewrite <- (HCf (OnEdge e (1 /sr 2))). reflexivity. } + (* why does this stopped working between coq-8.16 and 8.19? *) + (* cut ((bijectionG Ciso' (OnEdge e (1 /sr 2))) == (bijectionG (inverse Ciso) (OnEdge e (1 /sr 2)))). *) + pose (hsim := @bijectionG V E TG (inverse Ciso)). + + assert (@equiv (@loc V E TG) locG_Setoid + (@bijectionG V E TG Ciso' (@OnEdge V E TG e (1 /sr 2))) + (@hsim (@OnEdge V E TG e (1 /sr 2)) )). + { rewrite <- (HCf (@OnEdge V E TG e (1 /sr 2))). + reflexivity. } + destruct H. + assumption. } assert (HnotOnEdge : forall e p, Clocal_config (Good g) =/= SOnEdge e p). { intros e0 p0 H0. destruct (Clocal_config (Good g)) as [| e p] eqn:Habs; try tauto; []. @@ -680,8 +759,8 @@ simpl activate. destruct_match_eq Hactive. assert (Hnew_local_state : Dnew_local_state == state_G2V Cnew_local_state). { unfold Cnew_local_state, Dnew_local_state. unfold update, UpdateG, UpdateV. assert (Hlocal_g := Hlocal_config (Good g)). unfold config_G2V in Hlocal_g. - destruct (Clocal_config (Good g)) as [v e proof | e p] eqn:Hg; - try (exfalso; apply (HnotOnEdge _ _ (reflexivity _))); []. + destruct (Clocal_config (Good g)) as [v e proof | e p] eqn:Hg. + 2:(exfalso; apply (HnotOnEdge _ _ (reflexivity _))). (* the robot is on a vertex *) apply get_location_compat in Hlocal_g. unfold state_G2V in Hlocal_g. @@ -689,32 +768,43 @@ simpl activate. destruct_match_eq Hactive. unfold move. simpl fst. symmetry. do 2 destruct_match. * (* valid case: the robot chooses an adjacent edge *) - rewrite Hchoice. unfold add_edge. - destruct_match; [| destruct_match]; simpl state_G2V. - + (* the robot will not move so will end up in its starting position *) + revert Hchoice. destruct_match; intro Hchoice; hnf in Hchoice; rewrite Hchoice; + unfold add_edge; (destruct_match; [| destruct_match]); simpl state_G2V. + + (* absurd case *) assert (H0 : proj_ratio Cchoice = 0%R). - { assert (Hbounds := ratio_bounds Cchoice). simpl in *; lra. } + { assert (Hbounds := ratio_bounds Cchoice). cbn in *; lra. } assert (proj_ratio Cchoice < threshold Clocal_robot_decision)%R. - { rewrite H0. apply threshold_pos. } - destruct_match; try lra; []. symmetry. split; apply Hlocal_robot_decision. + { rewrite H0. apply strict_ratio_bounds. } + lra. + (* the robot reaches its destination *) change (proj_ratio ratio_0) with 0%R in *. rewrite Rplus_0_l in *. - assert (threshold Clocal_robot_decision <= proj_ratio Cchoice)%R. - { transitivity 1; trivial; []. apply Rlt_le. apply threshold_pos. } - destruct_match; try lra; []. symmetry. hnf. simpl fst. simpl snd. split; apply Hlocal_robot_decision. + (* the robot moves and ends up on the edge *) rewrite Rplus_0_l in *. - destruct_match. - - (* the ending point is after the edge threshold *) - symmetry. repeat split; simpl; apply Hlocal_robot_decision. + destruct_match; try contradiction; []. + (* the ending point is after the edge threshold *) + symmetry. repeat split; simpl; apply Hlocal_robot_decision. + + (* the robot will not move so will end up in its starting position *) + assert (H0 : proj_ratio Cchoice = 0%R). + { assert (Hbounds := ratio_bounds Cchoice). cbn in *; lra. } + assert (proj_ratio Cchoice < threshold Clocal_robot_decision)%R. + { rewrite H0. apply strict_ratio_bounds. } + symmetry. split; cbn -[equiv]; apply Hlocal_robot_decision. + + (* absurd case *) + exfalso. change (proj_ratio ratio_0) with 0%R in *. rewrite Rplus_0_l in *. + assert (threshold Clocal_robot_decision <= proj_ratio Cchoice)%R. + { transitivity 1; trivial; []. apply Rlt_le. apply strict_ratio_bounds. } + contradiction. + + (* the robot moves and ends up on the edge *) + rewrite Rplus_0_l in *. + destruct_match; try contradiction; []. - (* the ending point is before the edge threshold *) symmetry. repeat split; simpl; apply Hlocal_robot_decision. * (* absurd case: the robot does not make the same choice *) - match goal with | H : complement _ _ _ |- _ => elim H end. + match goal with | H : complement _ _ _ |- _ => contradiction H end. rewrite Hlocal_g. etransitivity; eauto. symmetry. apply Hlocal_robot_decision. * (* absurd case: the robot does not make the same choice *) - match goal with | H : complement _ _ _ |- _ => elim H end. + match goal with | H : complement _ _ _ |- _ => contradiction H end. rewrite <- Hlocal_g. etransitivity; eauto. apply Hlocal_robot_decision. * (* invalid case: the robot does not choose an adjacent edge *) rewrite Hlocal_config, <- Hg. reflexivity. } @@ -735,15 +825,13 @@ simpl activate. destruct_match_eq Hactive. change (Bijection.retraction (iso_E Diso)) with (Bijection.section (iso_E (inverse Diso))). symmetry. etransitivity; try apply (iso_E_compat (inverse Ciso)), He; []. apply E_subrelation. f_equiv. - apply Isomorphism.iso_E_compat. clear -Hiso. now apply inverse_compat. + apply Isomorphism.iso_E_compat. clear -Hiso. apply inverse_compat, Hiso. -- simpl snd. etransitivity; try apply HCisoE; []. change (Bijection.retraction (iso_E Diso)) with (Bijection.section (iso_E (inverse Diso))). symmetry. etransitivity; try apply (iso_E_compat (inverse Ciso)), He; []. - apply E_subrelation. f_equiv. - now apply Isomorphism.iso_E_compat, inverse_compat. - -- unfold equiv. cbn -[equiv Dnew_local_state inverse]. - change (iso_E Diso â»Â¹) with (iso_E (Diso â»Â¹)). - rewrite <- 2 iso_threshold. + apply E_subrelation. f_equiv. apply Isomorphism.iso_E_compat, inverse_compat, Hiso. + -- cbn-[Dnew_local_state IsoInverse]. change ((iso_VE Diso)â»Â¹) with (iso_VE (Diso â»Â¹)). + rewrite <-2 iso_threshold. rewrite (proj2 (projT2 (precondition_satisfied_inv da config g))), (proj2 (projT2 (precondition_satisfied_inv (da_C2D da config) (config_G2V config) g))). now rewrite He. @@ -751,22 +839,15 @@ simpl activate. destruct_match_eq Hactive. destruct Dnew_local_state as [[v' e'] Hvalid]. unfold state_G2V in *. simpl liftG. cbn iota beta. simpl snd. assert (Htest : threshold - (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) = threshold e) + (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) == threshold e) by now rewrite <- iso_threshold, (proj2 (projT2 (precondition_satisfied_inv da config g))). Time destruct (Rle_dec (threshold e) p) as [Hle | Hlt]. ++ destruct Hnew_local_state as [Hv He]. simpl fst in Hv. simpl snd in He. (* destruct takes too long... *) assert (threshold (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) - <= proj_ratio (proj_strict_ratio p)) by now rewrite Htest. - - (* too slow, case is faster *) - (* Time destruct_match; try contradiction; []. (* 230 sec!!!! *) *) - Time match goal with - | |- (match ?x with | _ => _ end) == _ => case x - end. - all:swap 1 2. - { intros notH. - apply (absurd _ H);assumption. } + <= proj_ratio (proj_strict_ratio p)). + { hnf in Htest. rewrite proj_ratio_strict_ratio in *. now rewrite Htest. } + destruct_match; try contradiction; []. split; simpl fst; simpl snd. -- transitivity (tgt (Bijection.section (iso_E (inverse Ciso)) e)); [apply HCisoE |]. rewrite Hv, <- (proj2 (iso_morphism _ e)). cbn -[equiv]. @@ -778,15 +859,9 @@ simpl activate. destruct_match_eq Hactive. symmetry. apply E_subrelation, Hiso. ++ destruct Hnew_local_state as [Hv He]. simpl fst in Hv. simpl snd in He. assert (¬ threshold (Bijection.section (iso_E (projT1 (precondition_satisfied_inv da config g))) e) - <= proj_ratio (proj_strict_ratio p)) by now rewrite Htest. - - (* too slow, case is faster *) - (* Time destruct_match; try contradiction; []. (* 230 sec!!!! *) *) - Time match goal with - | |- (match ?x with | _ => _ end) == _ => case x - end. - { intros notH. - apply (absurd _ notH);assumption. } + <= proj_ratio (proj_strict_ratio p)). + { hnf in Htest. rewrite proj_ratio_strict_ratio in *. now rewrite Htest. } + destruct_match; try contradiction; []. split; simpl fst; simpl snd. -- rewrite <- (proj1 (iso_morphism _ _)), Hv. transitivity (Bijection.section (iso_V (inverse Ciso)) (src e)); [apply HCiso' |]. @@ -805,9 +880,8 @@ simpl activate. destruct_match_eq Hactive. cbn -[equiv equiv_dec]. destruct_match. - (* the robot is at the edge src *) unfold add_edge, state_G2V. - assert (He := threshold_pos e). rewrite Hchoose_inactive. - change (proj_ratio ratio_0) with 0. rewrite Rplus_0_l at 2. - destruct (Rle_bool (threshold e) (proj_ratio (choose_inactive da config id))) eqn:Htest; + assert (He := strict_ratio_bounds (threshold e)). rewrite (proj_ratio_compat _ _ (Hchoose_inactive id)). + change (proj_ratio ratio_0) with 0. symmetry. destruct_match_eq Htest; rewrite Rle_bool_true_iff in Htest || rewrite Rle_bool_false_iff in Htest; repeat destruct_match; try (now split; auto); simpl in * |-; try rewrite Rplus_0_l in *; contradiction || lra. @@ -817,7 +891,7 @@ simpl activate. destruct_match_eq Hactive. + (* the robot is on an edge *) cbn -[equiv]. destruct_match; cbn -[equiv]. - (* the robot is already past the edge threshold *) - assert (He := threshold_pos e). + assert (He := strict_ratio_bounds (threshold e)). assert (Hp := ratio_bounds (choose_inactive da config id)). transitivity (exist valid_stateV (tgt e, e) (or_intror (reflexivity (tgt e)))); [| now repeat destruct_match]. @@ -830,11 +904,11 @@ simpl activate. destruct_match_eq Hactive. assert (Hp := strict_ratio_bounds p). intro Habs. simpl in Habs. lra. } destruct_match; try contradiction; []. - rewrite Hchoose_inactive. + rewrite (proj_ratio_compat _ _ (Hchoose_inactive id)). destruct (Rle_bool (threshold e) (p + choose_inactive da config id)) eqn:Hcase; rewrite Rle_bool_true_iff in Hcase || rewrite Rle_bool_false_iff in Hcase; repeat destruct_match; try solve [split; reflexivity | simpl in *; contradiction]; []. - elim Hcase. transitivity 1; trivial; []. apply Rlt_le. apply threshold_pos. + contradiction Hcase. transitivity 1; trivial; []. apply Rlt_le. apply strict_ratio_bounds. Qed. End GraphEquivalence. diff --git a/Models/Isometry.v b/Models/Isometry.v new file mode 100644 index 0000000000000000000000000000000000000000..d9da3e8908c2cd64a2e51f790ffd00ab1ebf0a4e --- /dev/null +++ b/Models/Isometry.v @@ -0,0 +1,125 @@ +(**************************************************************************) +(* Mechanised Framework for Local Interactions & Distributed Algorithms *) +(* T. Balabonski, P. Courtieu, L. Rieg, X. Urbain *) +(* PACTOLE project *) +(* *) +(* This file is distributed under the terms of the CeCILL-C licence. *) +(* *) +(**************************************************************************) + +(**************************************************************************) +(** Mechanised Framework for Local Interactions & Distributed Algorithms + T. Balabonski, P. Courtieu, L. Rieg, X. Urbain + PACTOLE project + + This file is distributed under the terms of the CeCILL-C licence. + *) +(**************************************************************************) + +Set Implicit Arguments. +Require Import Utf8. +Require Import SetoidDec. +Require Import Reals. +Require Import Pactole.Setting. +Require Import Pactole.Spaces.RealMetricSpace. +Require Import Pactole.Spaces.Isometry. + + +Typeclasses eauto := (bfs). + + +Section IsometryCenter. + +Context `{Observation}. +Context {VS : RealVectorSpace location}. +Context {RMS : RealMetricSpace location}. +Context `{robot_choice}. +Context `{update_choice}. +Context `{inactive_choice}. + +(** Similarities as a frame choice, inside real metric spaces *) +Global Instance FrameChoiceIsometry : frame_choice (isometry location) := {| + frame_choice_bijection := @iso_f location _ _ _ _; + frame_choice_Setoid := isometry_Setoid; + frame_choice_bijection_compat := f_compat |}. + +Definition isometry_da_prop da := + forall config g, center (change_frame da config g) == get_location (config (Good g)). + +(** When using isometries to change the frame of reference, we need to ensure that + the center of the isometry is the location of the robot. *) +Definition isometry_da := { iso_da : demonic_action | isometry_da_prop iso_da }. + +Instance isometry_da_Setoid : Setoid isometry_da := sig_Setoid _. + +Definition proj_iso_da : isometry_da -> demonic_action := @proj1_sig _ _. + +Global Instance proj_iso_da_compat : Proper (equiv ==> equiv) proj_iso_da. +Proof using . intros ? ? Heq. apply Heq. Qed. + +Definition isometry_center : forall da config g, + center (change_frame (proj_iso_da da) config g) == get_location (config (Good g)) + := @proj2_sig _ _. + + +(** Demons are now stream of [isometry_da]s.*) +Definition isometry_demon := Stream.t isometry_da. + + +(** From a [isometry_demon], we can recover the demon] and its property [isometry_da_prop]. *) +Definition isometry_demon2demon : isometry_demon -> demon := Stream.map proj_iso_da. + +Definition isometry_demon2prop : forall d, + Stream.forever (Stream.instant isometry_da_prop) (isometry_demon2demon d). +Proof. coinduction Hcorec. simpl. unfold proj_iso_da. apply proj2_sig. Defined. + +Global Instance isometry_demon2demon_compat : + Proper (@equiv _ Stream.stream_Setoid ==> @equiv _ Stream.stream_Setoid) isometry_demon2demon. +Proof using . unfold isometry_demon2demon. autoclass. Qed. + +Lemma isometry2demon_hd : forall d, Stream.hd (isometry_demon2demon d) = proj_iso_da (Stream.hd d). +Proof using . now intros []. Qed. + +Lemma isometry2demon_tl : forall d, Stream.tl (isometry_demon2demon d) = isometry_demon2demon (Stream.tl d). +Proof using . now intros []. Qed. + +Lemma isometry2demon_cons : forall da d, + @equiv _ Stream.stream_Setoid + (isometry_demon2demon (Stream.cons da d)) + (Stream.cons (proj_iso_da da) (isometry_demon2demon d)). +Proof using . intros. unfold isometry_demon2demon. now rewrite Stream.map_cons. Qed. + +(** Conversely, from a [demon] and a proof that [isometry_da_prop] is always true, + we can build a [isometry_demon]. *) + +CoFixpoint demon2isometry_demon : forall (d : demon) (Hd : Stream.forever (Stream.instant isometry_da_prop) d), + isometry_demon := fun d Hd => + Stream.cons + (exist _ (Stream.hd d) match Hd with | Stream.Always x _ => x end) + (@demon2isometry_demon (Stream.tl d) (match Hd with | Stream.Always _ x => x end)). + +Arguments demon2isometry_demon d Hd : clear implicits. + +Lemma demon2isometry_hd : forall d Hd, + Stream.hd (demon2isometry_demon d Hd) == exist _ (Stream.hd d) (match Hd with Stream.Always x _ => x end). +Proof using . now intros [] []. Qed. + +Lemma demon2isometry_tl : forall d Hd, + Stream.tl (demon2isometry_demon d Hd) == demon2isometry_demon (Stream.tl d) (match Hd with Stream.Always _ x => x end). +Proof using . now intros [] []. Qed. + +Theorem demon2demon : forall d Hd, + @equiv _ Stream.stream_Setoid (isometry_demon2demon (demon2isometry_demon d Hd)) d. +Proof using . coinduction Hcorec. unfold Stream.instant2. now rewrite isometry2demon_hd, demon2isometry_hd. Qed. + +End IsometryCenter. + +Coercion proj_iso_da : isometry_da >-> demonic_action. +Coercion isometry_demon2demon : isometry_demon >-> demon. + +(* + *** Local Variables: *** + *** coq-prog-name: "coqtop" *** + *** fill-column: 117 *** + *** End: *** + *) diff --git a/Models/NoByzantine.v b/Models/NoByzantine.v index 977fbbc9edd6ae9c62cd30a6ad38d1fcb8e9c83d..5e99c0b73ca084ce49970cac62683a55bc49cd08 100644 --- a/Models/NoByzantine.v +++ b/Models/NoByzantine.v @@ -1,32 +1,54 @@ -Require Import SetoidClass. -Require Import Pactole.Core.Identifiers. -Require Import Pactole.Core.State. -Require Import Pactole.Core.Configuration. +Require Import Utf8 SetoidClass. +From Pactole Require Import Core.Identifiers Core.State Core.Configuration. (** To use these results, just provide an instance of the [NoByzantine] class. *) Section NoByzantine. -Context `{Names}. -Context `{St : State}. +Context {N : Names}. +Context {L : Location} {info : Type} {St : State info}. Class NoByzantine := { nB_eq_0 : nB = 0 }. Context {NoByz : NoByzantine}. + (** There is no byzantine robot so we can simplify properties about identifiers and configurations. *) -Lemma no_byz : forall (id : ident) P, (forall g, P (Good g)) -> P id. + +Lemma Bnames_nil : Bnames = nil. +Proof using NoByz. + now rewrite <- List.length_zero_iff_nil, Bnames_length, nB_eq_0. +Qed. + +Lemma no_byz : ∀ (id : ident) (P : ident -> Type), (∀ g, P (Good g)) -> P id. +Proof using NoByz. + intros [g | b] P HP. { apply HP. } exfalso. apply (@List.in_nil _ b). + rewrite <- Bnames_nil. apply In_Bnames. +Qed. + +Lemma no_byz_inv : ∀ (P : ident -> Type), + (∀ id : ident, P id) -> ∀ g, P (Good g). +Proof using . intros * H *. apply H. Qed. + +Lemma no_byz_eq : ∀ config1 config2 : configuration, + (∀ g, config1 (Good g) == config2 (Good g)) -> config1 == config2. +Proof using NoByz. + intros config1 config2 Heq id. apply (no_byz id). intro g. apply Heq. +Qed. + +Lemma no_byz_fun {T : Type} : B -> T. Proof using NoByz. -intros [g | b] P HP. -+ apply HP. -+ assert (Hnil : Bnames = nil) by now rewrite <- List.length_zero_iff_nil, Bnames_length, nB_eq_0. - elim (@List.in_nil _ b). rewrite <- Hnil. apply In_Bnames. + intros b. exfalso. apply (@List.in_nil B b). rewrite <- Bnames_nil. + apply In_Bnames. Qed. -Lemma no_byz_eq : forall config1 config2 : configuration, - (forall g, config1 (Good g) == config2 (Good g)) -> config1 == config2. +Lemma no_byz_ex : ∀ (P : ident → Prop), + (∃ id : ident, P id) <-> (∃ g : G, P (Good g)). Proof using NoByz. -intros config1 config2 Heq id. apply (no_byz id). intro g. apply Heq. + intros. split. + - intros [[g | b] H]. { exists g. apply H. } exfalso. + apply (@List.in_nil B b). rewrite <- Bnames_nil. apply In_Bnames. + - intros [g H]. exists (Good g). apply H. Qed. End NoByzantine. diff --git a/Models/RigidFlexibleEquivalence.v b/Models/RigidFlexibleEquivalence.v index 98077dfc007a5408192e0eb40fb067ca5344c780..282b2cee0a9ba43d9016bd55b9649ea7b742c837 100644 --- a/Models/RigidFlexibleEquivalence.v +++ b/Models/RigidFlexibleEquivalence.v @@ -24,6 +24,7 @@ Require Import Reals Psatz. Require Import Pactole.Setting. Require Import Pactole.Spaces.RealMetricSpace. Require Import Pactole.Spaces.Similarity. +Require Import Pactole.Util.Bijection. Require Pactole.Models.Rigid. Require Pactole.Models.Flexible. diff --git a/Models/RingSSync.v b/Models/RingSSync.v new file mode 100644 index 0000000000000000000000000000000000000000..d3752a1540185656436c7de65550987b6c03be6c --- /dev/null +++ b/Models/RingSSync.v @@ -0,0 +1,37 @@ +Require Import Utf8. +From Pactole Require Export Setting Spaces.Graph Util.Bijection + Spaces.Isomorphism Util.Fin Spaces.Ring Observations.MultisetObservation. + +Set Implicit Arguments. + +Section RingSSync. + +Context {n : nat} {ltc_2_n : 2 <c n} {Robots : Names}. + +Global Existing Instance NodesLoc. + +Global Instance St : State location := OnlyLocation (fun _ => True). + +Global Existing Instance multiset_observation. + +Global Instance RC : robot_choice direction := + { robot_choice_Setoid := direction_Setoid }. + +(** Demon's frame choice: + we move back the robot to the origin with a translation + and we can choose the orientation of the ring. *) +Global Instance FC : frame_choice (location * bool) := { + frame_choice_bijection := + λ nb, if snd nb then (trans (fst nb)) ∘ (sym (fst nb)) + else trans (fst nb); + frame_choice_Setoid := eq_setoid _ }. + +Global Existing Instance NoChoice. +Global Existing Instance NoChoiceIna. +Global Existing Instance NoChoiceInaFun. + +Global Instance UpdFun : update_function direction (location * bool) unit := { + update := λ config g _ dir _, move_along (config (Good g)) dir; + update_compat := ltac:(repeat intro; subst; now apply move_along_compat) }. + +End RingSSync. diff --git a/Observations/Definition.v b/Observations/Definition.v index 69151d2e93a4ea3f0636596e6318cf84efafdb9e..424f97353726bee8fee2ac46dd54ba7836904cec 100644 --- a/Observations/Definition.v +++ b/Observations/Definition.v @@ -30,7 +30,7 @@ Class Observation {info} `{State info} `{Names} := { obs_is_ok : observation -> configuration -> info -> Prop; obs_from_config_spec : forall config st, obs_is_ok (obs_from_config config st) config st }. -Existing Instance observation_Setoid. -Existing Instance observation_EqDec. -Existing Instance obs_from_config_compat. +Global Existing Instance observation_Setoid. +Global Existing Instance observation_EqDec. +Global Existing Instance obs_from_config_compat. Arguments obs_from_config : simpl never. diff --git a/Observations/LimitedMultisetObservation.v b/Observations/LimitedMultisetObservation.v index ee35ad67c2aa44532d5ff1d3185e78c0b8d1bbae..2fc98cc07d6b32a2ca64ab901160daad01972440 100644 --- a/Observations/LimitedMultisetObservation.v +++ b/Observations/LimitedMultisetObservation.v @@ -23,6 +23,7 @@ Require Import Lia. Require Import SetoidList. Require Import SetoidDec. Require Import Rbase. +Require Import Pactole.Util.Bijection. Require Import Pactole.Util.FMaps.FMapList. Require Import Pactole.Util.MMultiset.MMultisetWMap. Require Export Pactole.Util.MMultiset.MMultisetInterface. @@ -41,8 +42,6 @@ Require Import Pactole.Spaces.Similarity. Close Scope R_scope. Set Implicit Arguments. -Coercion Bijection.section : Bijection.bijection >-> Funclass. - Section MultisetObservation. diff --git a/Observations/LimitedSetObservation.v b/Observations/LimitedSetObservation.v index d530d54229535d5f68e4898e2d7740b297ff1145..d147a81de8f3742089e2b852481e86df171531da 100644 --- a/Observations/LimitedSetObservation.v +++ b/Observations/LimitedSetObservation.v @@ -22,6 +22,7 @@ Require Import SetoidDec. Require Import Rbase. Require Import Pactole.Util.FSets.FSetInterface. Require Import Pactole.Util.FSets.FSetFacts. +Require Import Pactole.Util.Bijection. Require Import Pactole.Util.Coqlib. Require Import Pactole.Core.Identifiers. Require Import Pactole.Core.State. @@ -30,7 +31,6 @@ Require Import Pactole.Observations.Definition. Require Import Pactole.Spaces.RealMetricSpace. Require Import Pactole.Spaces.Similarity. Require Pactole.Observations.SetObservation. -Coercion Bijection.section : Bijection.bijection >-> Funclass. Section LimitedSetObservation. @@ -75,7 +75,7 @@ Lemma obs_from_config_map : forall radius (sim : similarity location), Proof using . repeat intro. unfold obs_from_config, limited_set_observation. rewrite config_list_map; try (now apply lift_compat; simpl; apply Bijection.section_compat); []. -rewrite map_map, 2 filter_map, <- SetObservation.make_set_map, map_map; autoclass; []. +rewrite map_map, 2 filter_map, <- SetObservation.make_set_map, map_map; autoclass. apply SetObservation.make_set_compat, eqlistA_PermutationA_subrelation. assert (Hequiv : (@equiv _ state_Setoid ==> @equiv _ location_Setoid)%signature (fun x => sim (get_location x)) (fun x => get_location (lift (existT precondition sim Psim) x))). diff --git a/Observations/MultisetObservation.v b/Observations/MultisetObservation.v index d7fecb23396b8d4af8c596a6cc48b56adda817d2..0eb756a587d23b5cbe28bf27c38a11ca2d942ac9 100644 --- a/Observations/MultisetObservation.v +++ b/Observations/MultisetObservation.v @@ -79,7 +79,7 @@ rewrite make_multiset_cons in Hin. destruct (equiv_dec x a) as [Heq | Heq]. + rewrite equiv_dec_refl in Hin. lia. + rewrite equiv_dec_refl in Hin. - rewrite plus_comm in Hin. cbn in Hin. apply eq_add_S, IHl in Hin. destruct Hin as [l' [Hl1 Hl2]]. + rewrite Nat.add_comm in Hin. cbn in Hin. apply eq_add_S, IHl in Hin. destruct Hin as [l' [Hl1 Hl2]]. exists l'. split. assumption. simpl. now constructor. - rewrite add_other in Hin; auto. apply IHl in Hin. destruct Hin as [l' [Hl1 Hl2]]. exists (a :: l'). split. intro Hin; inversion_clear Hin; contradiction. @@ -93,7 +93,7 @@ intros x n. induction n. + simpl alls. rewrite make_multiset_cons. rewrite IHn. intro y. rewrite singleton_spec. destruct (equiv_dec y x) as [Heq | Heq]. - rewrite Heq, add_spec, singleton_spec. - destruct (equiv_dec x x) as [_ | Helim]. lia. now elim Helim. + destruct (equiv_dec x x) as [_ | Helim]. lia. now contradiction Helim. - rewrite add_other; auto. rewrite singleton_spec. destruct (equiv_dec y x); trivial; []. contradiction. Qed. @@ -210,7 +210,6 @@ rewrite config_list_map, map_map, <- make_multiset_map, map_map. { intros pt1 pt2 Heq. now rewrite get_location_lift, Heq. } now apply (map_extensionalityA_compat _ Hequiv). + assumption. -+ now apply lift_compat. Qed. Theorem cardinal_obs_from_config : forall config pt, cardinal (obs_from_config config pt) = nG + nB. diff --git a/Observations/SetObservation.v b/Observations/SetObservation.v index 21052c66e32e99f8c9d3ffa0309bcf4a49328a08..dc3bcf01b91d31adde53e30415861f1ccce0290f 100644 --- a/Observations/SetObservation.v +++ b/Observations/SetObservation.v @@ -101,7 +101,7 @@ intros x l. induction l as [| e l]. * destruct IHl as [l' [n [Hin Hperm]]]. destruct (e =?= x) as [Heq | Heq]. + exists l', (S n). split; trivial. simpl. apply PermutationA_cons; assumption. + exists (e :: l'), n. split. - - intro Habs. inversion_clear Habs. elim Heq. now symmetry. contradiction. + - intro Habs. inversion_clear Habs. apply Heq. now symmetry. contradiction. - rewrite Hperm. apply (PermutationA_middle _). Qed. @@ -204,7 +204,6 @@ rewrite config_list_map, map_map, <- make_set_map, map_map. { intros pt1 pt2 Heq. now rewrite get_location_lift, Heq. } now apply (map_extensionalityA_compat _ Hequiv). + assumption. -+ now apply lift_compat. Qed. Theorem cardinal_obs_from_config : forall config state, diff --git a/Pactole_all.v b/Pactole_all.v new file mode 100644 index 0000000000000000000000000000000000000000..78c1c8ef8701e73e91d51bbe0120ae94d86e25e1 --- /dev/null +++ b/Pactole_all.v @@ -0,0 +1,53 @@ +Require Import Pactole.Models.NoByzantine. +Require Import Pactole.Models.RigidFlexibleEquivalence. +Require Import Pactole.Models.ContinuousGraph. +Require Import Pactole.Models.Rigid. +Require Import Pactole.Models.RigidFlexibleEquivalence_Assumptions. +Require Import Pactole.Models.DiscreteGraph. +Require Import Pactole.Models.GraphEquivalence. +Require Import Pactole.Models.Flexible. +Require Import Pactole.Models.Similarity. +Require Import Pactole.Spaces.Isomorphism. +Require Import Pactole.Spaces.R. +Require Import Pactole.Spaces.Ring. +Require Import Pactole.Spaces.RealNormedSpace. +Require Import Pactole.Spaces.RealMetricSpace. +Require Import Pactole.Spaces.RealVectorSpace. +Require Import Pactole.Spaces.EuclideanSpace. +Require Import Pactole.Spaces.Grid. +Require Import Pactole.Spaces.Similarity. +Require Import Pactole.Spaces.Graph. +Require Import Pactole.Spaces.R2. +Require Import Pactole.Setting. +Require Import Pactole.Observations.Definition. +Require Import Pactole.Observations.MultisetObservation. +Require Import Pactole.Observations.LimitedMultisetObservation. +Require Import Pactole.Observations.PreCompositionObservation. +Require Import Pactole.Observations.SetObservation. +Require Import Pactole.Observations.PointedObservation. +Require Import Pactole.Observations.LimitedSetObservation. +Require Import Pactole.minipactole.minipactole. +Require Import Pactole.Core.State. +Require Import Pactole.Core.Configuration. +Require Import Pactole.Core.Formalism. +Require Import Pactole.Core.Identifiers. +Require Import Pactole.Util.NumberComplements. +Require Import Pactole.Util.ListComplements. +Require Import Pactole.Util.SetoidDefs. +Require Import Pactole.Util.Lexprod. +Require Import Pactole.Util.FSets.FSetList. +Require Import Pactole.Util.FSets.FSetInterface. +Require Import Pactole.Util.FSets.FSetFacts. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Util.FMaps.FMapInterface. +Require Import Pactole.Util.FMaps.FMapList. +Require Import Pactole.Util.FMaps.FMapFacts. +Require Import Pactole.Util.Coqlib. +Require Import Pactole.Util.Stream. +Require Import Pactole.Util.MMultiset.MMultisetWMap. +Require Import Pactole.Util.MMultiset.MMultisetFacts. +Require Import Pactole.Util.MMultiset.MMultisetInterface. +Require Import Pactole.Util.MMultiset.MMultisetExtraOps. +Require Import Pactole.Util.MMultiset.Preliminary. +Require Import Pactole.Util.Ratio. +Require Import Pactole.Util.Preliminary. diff --git a/README.md b/README.md index d30a3e8798ae897337e0e866e440e70340aa3303..7b6fd626d2a0fe84631858cd64561ba69507b336 100644 --- a/README.md +++ b/README.md @@ -1,72 +1,106 @@ -This repository stores the Coq code of the Pactole project, -dedicated to formal verification of mobile robotic swarm protocols. +# Content -# Overall Structure - -- *Setting.v*: All you need to setup a working framework. A good starting point. -- *Util/*: Extension the to the Coq standard library that are not specific to Pactole -- *Core/*: The core the the Pactole framework, implementing the Look/Compute/Move cycle -- *Spaces/*: Spaces in which robots evolve -- *Observations/*: Types of robot views of the configuration -- *Models/*: Additional properties of some models -- *CasesStudies/* - - *Convergence/*: Convergence of robots in the 2D Euclidean plane - - *Gathering/*: Gathering in R or R² for various models - - *Exploration/*: Exploration of a ring with stop - - *Volume/*: Life line connection in the 2D Euclidean plane - -# Contributing - -The golden rule is that **MASTER SHOULD COMPILE AT ALL TIME** and does not -contains in-progress developments. - -More precisely, development never happens on master: each topic should have -its own branch (feel free to create one if needed) and work happens there, -even when several people are involved. - -Once a set of changes (e.g. a new case study) is complete and polished enough -(comments, adequate identation, no use of generated hypothesis names, etc.), -you may merge it to master by squashing its commits into meaningful pieces -(only a few, usually one may be enough) or submit a pull request. - -# Fast Compiling when developing - -During development you can benefit from coq's "vos/vok" generation to -speed up compilation. The only specificity of Pactole concerns -compilation of the files named xxx_Assumptions.v. This files print the -sets of assumptions used in the proofs of final lemmas in Case -studies. Compiling these files for vos/vok target would raise errors. -We provide adapted targets: +This repository stores the Coq code of the Pactole +project (https://pactole.liris.cnrs.fr/), dedicated to formal +verification of mobile robotic swarm protocols in many variants of the +model initially proposed by Suzuki and Yamashita [1] and sometimes +called the "look-compute-move model". -## Very fast but unsafe +It contains an abstract and parametrized formal model and a few +case studies. The structure of the repository is described below. -make [-j] vos-nocheck +# Support -Replaces "make vos". Prefer this when developing. +Pactole was financially supported by the following projects: -### proofgeneral +- [Pactole](https://pactole.liris.cnrs.fr/) started as + Digiteo Project #2009-38HD. +- [SAPPORO](https://sapporo.liris.cnrs.fr/) funded by the French + National Research Agency (ANR) under the reference 2019-CE25-0005 -For easy use of this feature you can use the auto compilation feature -of proofgeneral. menu: +[1] I. Suzuki and M. Yamashita. Distributed Anonymous Mobile Robots: Formation + of Geometric Patterns. SIAM Journal of Computing, 28(4):1347–1363, 1999. -menu: Coq / Auto Compilation / Compile Before Require -and then: Coq / Auto Compilation / vos compilation - -Now you can transparently script in any buffer, all needed file will -be compiled quickly. Don't forget to make a big fat "make" from time -to time. - -## slow, almost safe, very parallelisable - -make [-j] vok-nocheck - -replaces "make vok". Prefer this once a day. - -## Completely safe - -make [-j] +# Overall Structure -This is the only way to make a really safe compilation including all -xxx_Assumption.v. You should always do this once in a while to make -sure some universe constraints aren't failing and to check if you did -not miss an remaining axiom. +- *Setting.v*: All you need to setup a working framework. A good starting point. +- *Util/*: Extension to the Coq standard library that are not specific to Pactole. +- *Core/*: The core of the Pactole framework, implementing the Look/Compute/Move cycle. +- *Spaces/*: Spaces in which robots evolve. +- *Observations/*: Types of robot views of the configuration. +- *Models/*: Additional properties of some models. +- *CaseStudies/* : Case studies. + +# Case Studies + +The directory `CaseStudies` contains the following case +studies. Each case study (say `casestudy.v`) has a companion file +called `casestudy_Assuptions.v` the only purpose of which is to call `Print +Assumption` on the main theorem(s) of the case study. +This command is not included in the case study itself to allow for +fast compilation. + +Here is a list of the current case studies: + +- [Convergence/](Casestudy/Convergence): + - [Algorithm_noB.v](CaseStudies/Convergence/Algorithm_noB.v): + Convergence without Byzantine robots on the Euclidean plane. + - [Impossibility_2G_1B.v](CaseStudies/Convergence/Impossibility_2G_1B.v): + Impossibility of convergence on the real line when 1/3 of robots are + Byzantine. Auger, Bouzid, Courtieu, Tixeuil, Urbain. Certified + Impossibility Results for Byzantine-Tolerant Mobile Robots. SSS + 2013. + +- [Gathering/](CaseStudies/Gathering): Gathering in R or R² for + various models + - [Impossibility.v](CaseStudies/Gathering/Impossibility.v): + Impossibility of gathering in SSYNC. + - [Definitions.v](CaseStudies/Gathering/Definitions.v): + Common definitions about the gathering problem. + - [WithMultiplicity.v](CaseStudies/Gathering/WithMultiplicity.v): + Common definition on gathering when robots enjoy strong + multiplicity detection. + - [InR/](CaseStudies/Gathering/InR) case studies for the gathering + on the Euclidean line + - [Impossibility.v](CaseStudies/Gathering/InR/Impossibility.v): + Impossibility of gathering on the line in SSYNC. + Courtieu, Rieg, Tixeuil, Urbain. Impossibility of gathering, a certification. IPL 115. + - [Algorithm.v](CaseStudies/Gathering/InR/Algorithm.v): Gathering + one the line in SSYNC with strong multiplicity detection, from + non bivalent configurations. + - [InR2/](CaseStudies/Gathering/InR2) case studies for the gathering + on the Euclidean plane + - [Peleg.v](CaseStudies/Gathering/InR2/Peleg.v): + Gathering in FSYNC and non rigid moves with weak mutliplicity detection, due to Peleg. + Cohen, Peleg. Convergence Properties of the Gravitational Algorithm in Asynchronous Robot Systems. SIAM Journal of Computing, 34(6):1516–1528, 2005. + - [Viglietta.v](CaseStudies/Gathering/InR2/Viglietta.v): + Formalization of a protocol for gathering with lights due to Viglietta. + Viglietta. Rendezvous of two robots with visible bits. ALGOSENSORS 2013. + - [FSyncFlexNoMultAlgorithm.v](CaseStudies/Gathering/InR2/FSyncFlexNoMultAlgorithm.v): + Gathering in FSYNC and non rigid moves with no mutliplicity detection. + Balabonski, Delga, Rieg, Tixeuil, Urbain. Synchronous Gathering Without Multiplicity Detection: A Certified Algorithm. Theory Comput. Syst. 63(2): 200-218 (2019) + - [Algorithm.v](CaseStudies/Gathering/InR2/Algorithm.v): + SSYNC Gathering in R² with strong multiplicity detection, from + non-bivalent configurations. + Courtieu, Rieg, Tixeuil, Urbain. Certified Universal Gathering in R² for Oblivious Mobile Robots. DISC 2016. +- [Exploration/](CaseStudies/Exploration): Exploration of a ring with stop + - [ImpossibilityKDividesN.v](CaseStudies/Exploration/ImpossibilityKDividesN.v): + Impossibility of exploration of a ring when the number of robots + divides the number of nodes. + Balabonski, Pelle, Rieg, Tixeuil. A Foundational Framework for Certified Impossibility Results with Mobile Robots on Graphs. ICDCN 2018. + - [ExplorationDefs.v](CaseStudies/Exploration/ExplorationDefs.v): + Common definitions on exploration. + - [Tower.v](CaseStudies/Exploration/Tower.v): + Exploration with stop on a ring requires forming a tower, in particular a single robot is no enough. +- [LifeLine/](CaseStudies/LifeLine): Life line connection in the 2D Euclidean plane + - [Algorithm.v](CaseStudies/LifeLineAlgorithm.v): + Connection maintenance protocol on R2. + Balabonski, Courtieu, Pelle, Rieg, Tixeuil, Urbain. Computer Aided Formal Design of Swarm Robotics Algorithms. SSS 2021. + +# Other Related Ressources + +A general description of the Pactole library and its use: + + Courtieu, L. Rieg, S. Tixeuil, and X. Urbain. Swarms of Mobile Robots: Towards + Versatility with Safety. Leibniz Transactions on Embedded Systems, 8(2):02:1– + 02:36, 2022. diff --git a/Setting.v b/Setting.v index 8146c286d2f025822b9c76887d15ea14a9cde81f..d2caa4e46a69a4208428606399ac42f72a09a2f1 100644 --- a/Setting.v +++ b/Setting.v @@ -1,8 +1,13 @@ Require Export SetoidDec. Require Export Pactole.Util.Coqlib. +Require Export Pactole.Util.SetoidDefs. +Require Export Pactole.Util.NumberComplements. +Require Export Pactole.Util.ListComplements. Require Export Pactole.Util.Ratio. Require Pactole.Util.Stream. Require Pactole.Util.Lexprod. +Require Pactole.Util.Fin. +Require Pactole.Util.Enum. Require Export Pactole.Core.Identifiers. Require Export Pactole.Core.State. Require Export Pactole.Core.Configuration. @@ -14,26 +19,26 @@ Remove Hints eq_setoid : Setoid. Coercion Bijection.section : Bijection.bijection >-> Funclass. Coercion Similarity.sim_f : Similarity.similarity >-> Bijection.bijection. -Existing Instance Stream.stream_Setoid. -Existing Instance Stream.hd_compat. -Existing Instance Stream.tl_compat. -Existing Instance Stream.constant_compat. -Existing Instance Stream.aternate_compat. -Existing Instance Stream.instant_compat. -Existing Instance Stream.forever_compat. -Existing Instance Stream.eventually_compat. -Existing Instance Stream.instant2_compat. -Existing Instance Stream.forever2_compat. -Existing Instance Stream.eventually2_compat. -Existing Instance Stream.instant2_refl. -Existing Instance Stream.instant2_sym. -Existing Instance Stream.instant2_trans. -Existing Instance Stream.forever2_refl. -Existing Instance Stream.forever2_sym. -Existing Instance Stream.forever2_trans. -Existing Instance Stream.eventually2_refl. -Existing Instance Stream.eventually2_sym. -Existing Instance Stream.map_compat. +Global Existing Instance Stream.stream_Setoid. +Global Existing Instance Stream.hd_compat. +Global Existing Instance Stream.tl_compat. +Global Existing Instance Stream.constant_compat. +Global Existing Instance Stream.aternate_compat. +Global Existing Instance Stream.instant_compat. +Global Existing Instance Stream.forever_compat. +Global Existing Instance Stream.eventually_compat. +Global Existing Instance Stream.instant2_compat. +Global Existing Instance Stream.forever2_compat. +Global Existing Instance Stream.eventually2_compat. +Global Existing Instance Stream.instant2_refl. +Global Existing Instance Stream.instant2_sym. +Global Existing Instance Stream.instant2_trans. +Global Existing Instance Stream.forever2_refl. +Global Existing Instance Stream.forever2_sym. +Global Existing Instance Stream.forever2_trans. +Global Existing Instance Stream.eventually2_refl. +Global Existing Instance Stream.eventually2_sym. +Global Existing Instance Stream.map_compat. (* By experience, this is not very useful (** For simplicity, we gather into one definition all the classes that must be instanciated @@ -67,4 +72,4 @@ Class GlobalDefinitions := { glob_robot_choice glob_frame_choice glob_update_choice; glob_inactive_function :> @inactive_function _ _ glob_State glob_Names _ glob_inactive_choice }. -*) \ No newline at end of file +*) diff --git a/Spaces/EuclideanSpace.v b/Spaces/EuclideanSpace.v index a1cb84fbabbea91b04c2ee5af67ba704fbee10b0..8ceadb0e0b1e52ded20f15d6f91a787dca62a3cb 100644 --- a/Spaces/EuclideanSpace.v +++ b/Spaces/EuclideanSpace.v @@ -28,8 +28,8 @@ Class EuclideanSpace (T : Type) {S : Setoid T} {EQ : @EqDec T S} {VS : RealVecto inner_product_nonneg : forall u, 0 <= inner_product u u; inner_product_defined : forall u, inner_product u u = 0%R <-> u == origin}. -Existing Instance inner_product_compat. -Arguments inner_product {T%type} {_} {_} {_} {_} u%VS V%VS. +Global Existing Instance inner_product_compat. +Arguments inner_product {T%_type} {_} {_} {_} {_} u%_VS V%_VS. Notation "〈 u , v 〉" := (inner_product u v) (format "'〈' u , v '〉'"): VectorSpace_scope. (* Open Scope VectorSpace_scope. *) @@ -156,7 +156,7 @@ Section PerpendicularResults. Qed. End PerpendicularResults. -Arguments perpendicular {T%type} {_} {_} {_} {_} u%VS v%VS. +Arguments perpendicular {T%_type} {_} {_} {_} {_} u%_VS v%_VS. Notation "u ⟂ v" := (perpendicular u v) (at level 50, no associativity). (** *** The norm induced by the [inner_product] **) diff --git a/Spaces/Graph.v b/Spaces/Graph.v index f4b39d64bdede77fcc076d62e75b8ca793ac05a8..9ce11aa9a542f8dcc5dbd6043c801c013ebcedcc 100644 --- a/Spaces/Graph.v +++ b/Spaces/Graph.v @@ -8,43 +8,80 @@ *) (**************************************************************************) - -Require Import Rbase SetoidDec. +Require Import Utf8 Rbase SetoidDec. Require Import Pactole.Util.Coqlib. -Require Pactole.Core.Identifiers. +Require Import Pactole.Util.Fin. +Require Import Pactole.Util.Ratio. +Require Import Pactole.Core.State. Set Implicit Arguments. Class Graph (V E : Type) := { - V_Setoid :> Setoid V; - E_Setoid :> Setoid E; - V_EqDec :> EqDec V_Setoid; - E_EqDec :> EqDec E_Setoid; + #[global] V_Setoid :: Setoid V; + #[global] E_Setoid :: Setoid E; + #[global] V_EqDec :: EqDec V_Setoid; + #[global] E_EqDec :: EqDec E_Setoid; src : E -> V; (* source and target of an edge *) tgt : E -> V; - threshold : E -> R; (* TODO: use [strict_ratio] instead? *) - threshold_pos : forall e, (0 < threshold e < 1)%R; - src_compat :> Proper (equiv ==> equiv) src; - tgt_compat :> Proper (equiv ==> equiv) tgt; - threshold_compat :> Proper (equiv ==> Logic.eq) threshold; + #[global] src_compat :: Proper (equiv ==> equiv) src; + #[global] tgt_compat :: Proper (equiv ==> equiv) tgt; find_edge : V -> V -> option E; - find_edge_compat :> Proper (equiv ==> equiv ==> opt_eq equiv) find_edge; - find_edge_None : forall a b : V, find_edge a b == None <-> forall e : E, ~(src e == a /\ tgt e == b); - find_edge_Some : forall v1 v2 e, find_edge v1 v2 == Some e <-> v1 == src e /\ v2 == tgt e }. + #[global] find_edge_compat :: Proper (equiv ==> equiv ==> opt_eq equiv) find_edge; + find_edge_Some : ∀ e v1 v2, find_edge v1 v2 == Some e <-> v1 == src e /\ v2 == tgt e}. -Global Opaque threshold_pos src_compat tgt_compat threshold_compat find_edge_compat find_edge_None find_edge_Some. +Class ThresholdGraph (V E : Type) := { + nothreshold_graph : Graph V E; + threshold : E -> strict_ratio; + #[global] threshold_compat :: Proper (equiv ==> equiv) threshold}. -(** A finite graph ia a graph where the set [V] of vertices is a prefix of N. *) -(* FIXME: nothing prevents E from being infinite! *) -(* TODO: Maybe we should reuse the type used for robot names *) -Definition finite_node n := {m : nat | m < n}. +Coercion nothreshold_graph : ThresholdGraph >-> Graph. +Global Existing Instance nothreshold_graph. +Global Opaque src_compat tgt_compat threshold_compat find_edge_compat find_edge_Some. + +Section some_lemmas. + +Context {V E : Type} {G : Graph V E}. + +Lemma find_edge_None : ∀ v1 v2 : V, + @find_edge _ _ G v1 v2 == None <-> ∀ e : E, ¬ (v1 == src e /\ v2 == tgt e). +Proof using . + intros. rewrite <- (Decidable.not_not_iff _ (option_decidable _)). + setoid_rewrite not_None_iff. split. + - intros H1 e H2. apply H1. exists e. apply find_edge_Some, H2. + - intros H1 [e H2]. apply (H1 e), find_edge_Some, H2. +Qed. -(* We explictely define the setoid here to avoid using proj1_Setoid instead. *) -Instance finite_node_Setoid n : Setoid (finite_node n) := eq_setoid _. -Instance finite_node_EqDec n : EqDec (finite_node_Setoid n) := @Identifiers.subset_dec n. +Lemma find_edge_pick : ∀ v1 v2 : V, + pick_spec (λ e : E, v1 == src e /\ v2 == tgt e) (find_edge v1 v2). +Proof using . + intros. apply pick_Some_None. intros e. + apply find_edge_Some. apply find_edge_None. +Qed. -Definition FiniteGraph (n : nat) E := Graph (finite_node n) E. +(* The specifications of find_edge make the graph simple *) +Lemma simple_graph : ∀ e1 e2 : E, + @src _ _ G e1 == src e2 /\ @tgt _ _ G e1 == tgt e2 -> e1 == e2. +Proof using . + intros * [Hs Ht]. apply Some_inj. erewrite <-2 (proj2 (find_edge_Some _ _ _)). + apply find_edge_compat. 3,4: split. 1,2,5,6: symmetry. + 1,3,5: apply Hs. all: apply Ht. +Qed. + +Lemma simple_graph_iff : ∀ e1 e2 : E, + @src _ _ G e1 == src e2 /\ @tgt _ _ G e1 == tgt e2 <-> e1 == e2. +Proof using . + intros. split. apply simple_graph. intros <-. split. all: reflexivity. +Qed. + +End some_lemmas. + +(** A finite graph ia a graph where the set [V] of vertices is a prefix of N. *) +(* FIXME: nothing prevents E from being infinite! *) +(* Definition FiniteGraph (n : nat) E := Graph (fin n) E. Existing Class FiniteGraph. -Global Identity Coercion proj_graph : FiniteGraph >-> Graph. +Global Identity Coercion proj_graph : FiniteGraph >-> Graph. *) + +Definition NodesLoc {V E : Type} {G : Graph V E} : Location + := make_Location V. diff --git a/Spaces/Grid.v b/Spaces/Grid.v index d973803b6fb15624dcf9f3f6692f23dd83126b0d..4b6c7aa433efc52500211ec6360f9dadb4055d10 100644 --- a/Spaces/Grid.v +++ b/Spaces/Grid.v @@ -11,7 +11,7 @@ Require Import Psatz SetoidDec ZArith Rbase. Require Import Pactole.Util.Coqlib. -(*Require Import Pactole.Core.Robots.*) +Require Import Pactole.Util.Bijection. Require Export Pactole.Spaces.Graph. @@ -24,10 +24,10 @@ Open Scope Z_scope. Inductive direction := North | South | East | West | Self. -Instance direction_Setoid : Setoid direction := eq_setoid _. +#[export] Instance direction_Setoid : Setoid direction := eq_setoid _. -Instance direction_EqDec : EqDec direction_Setoid. -Proof. +#[export] Instance direction_EqDec : EqDec direction_Setoid. +Proof using . intros x y. simpl. change (x = y -> False) with (x <> y). decide equality. Defined. @@ -35,9 +35,9 @@ Defined. Notation node := (Z*Z)%type. Notation edge := (Z*Z*direction)%type. -Instance node_Setoid : Setoid node := eq_setoid _. -Instance node_EqDec : EqDec node_Setoid. -Proof. +#[export] Instance node_Setoid : Setoid node := eq_setoid _. +#[export] Instance node_EqDec : EqDec node_Setoid. +Proof using . intros x y. destruct (fst x =?= fst y). + destruct (snd x =?= snd y). @@ -46,9 +46,9 @@ destruct (fst x =?= fst y). + right. abstract (destruct x, y; injection; auto). Defined. -Instance edge_Setoid : Setoid edge := eq_setoid _. -Instance edge_EqDec : EqDec edge_Setoid. -Proof. +#[export] Instance edge_Setoid : Setoid edge := eq_setoid _. +#[export] Instance edge_EqDec : EqDec edge_Setoid. +Proof using . intros x y. destruct (fst x =?= fst y). + destruct (snd x =?= snd y). @@ -68,39 +68,24 @@ Definition edge_tgt (e : edge) : node := Arguments edge_tgt !e. (** The Z² grid is a graph. *) -Instance Z2 : Graph node edge. +#[export] Instance Z2 : Graph node edge. simple refine {| V_EqDec := node_EqDec; E_EqDec := edge_EqDec; src := fst; - tgt := edge_tgt; - threshold := fun _ => (1 / 2)%R |}. -Proof. -* (* threshold_pos *) - intros. lra. + tgt := edge_tgt; |}. +Proof using . * (* find_edge *) exact (fun x y : node => if equiv_dec (EqDec := node_EqDec) y x then Some (x, Self) else if y =?= (fst x + 1, snd x) then Some (x, East) else if y =?= (fst x, snd x + 1) then Some (x, North) else if y =?= (fst x - 1, snd x) then Some (x, West) else if y =?= (fst x, snd x - 1) then Some (x, South) else None). -* (* find_edge_None *) - intros x y. cbn -[equiv]. repeat destruct_match. - + abstract (split; try tauto; []; intro He; apply (He (x, Self)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, East)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, North)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, West)); auto). - + abstract (split; try tauto; []; intro He; apply (He (x, South)); auto). - + split; intros _; auto; []. - abstract (intros [x' d] [Hx He]; simpl in Hx; subst x'; - rewrite <- He in *; destruct d; unfold edge_tgt in *; simpl in *; auto). * (* find_edge_Some *) - intros x y e. cbn -[equiv]. repeat destruct_match; simpl; - try abstract (solve [ split; intro Heq; subst; unfold edge_tgt; simpl in *; try tauto; []; - destruct e as [p d], Heq as [? Heq]; simpl in *; f_equal; trivial; []; - subst; unfold edge_tgt in *; - destruct p, d; simpl in *; reflexivity || (injection Heq; lia) ]); []. - split; try tauto; []. intros []. - abstract (subst; destruct e as [? []]; unfold edge_tgt in *; simpl in *; auto). + abstract (intros [p d] x y; cbn-[equiv]; + assert (forall x : Z, x <> x + 1) by lia; + assert (forall x : Z, x <> x - 1) by lia; + assert (forall x : Z, x + 1 <> x - 1) by lia; + repeat destruct_match; destruct x, y, p, d; cbn in *; intuition congruence). Defined. (** ** Change of frame of reference in Z² **) @@ -109,14 +94,14 @@ Require Pactole.Util.Bijection. Require Import Pactole.Core.State. Require Import Pactole.Core.Formalism. -Instance Loc : Location := {| location := node |}. +#[export] Instance Loc : Location := {| location := node |}. (** angle: anglei represents the possible angles for a rotation of reflection: - for a rotation: angle i/2 * pi; - for a reflection: angle i/4 * pi *) Inductive angle := angle0 | angle1 | angle2 | angle3. -Instance angle_Setoid : Setoid angle := eq_setoid _. -Instance angle_EqDec : EqDec angle_Setoid. -Proof. +#[export] Instance angle_Setoid : Setoid angle := eq_setoid _. +#[export] Instance angle_EqDec : EqDec angle_Setoid. +Proof using . intros x y. simpl. change (x = y -> False) with (x <> y). decide equality. @@ -136,13 +121,13 @@ Definition opp_angle (r : angle) := Definition translation (v : Z*Z) : Bijection.bijection (Z*Z). refine {| Bijection.section := fun x => (fst x + fst v, snd x + snd v); Bijection.retraction := fun x => (fst x - fst v, snd x - snd v) |}. -Proof. +Proof using . intros x y. simpl. abstract (split; intro; subst; destruct x || destruct y; f_equal; simpl; lia). Defined. -Instance translation_compat : Proper (equiv ==> equiv) translation. -Proof. intros ? ? Heq x. now rewrite Heq. Qed. +#[export] Instance translation_compat : Proper (equiv ==> equiv) translation. +Proof. intros ? ? Heq x. Fail timeout 2 now rewrite Heq. cbn in Heq. now subst. Qed. (** Rotation *) Definition mk_rotation r : Z*Z -> Z*Z := @@ -153,15 +138,19 @@ Definition mk_rotation r : Z*Z -> Z*Z := | angle3 => fun x => (snd x, - fst x) end. +Lemma mk_rotation_compat : forall r, Proper (equiv ==> equiv) (mk_rotation r). +Proof. intros [] [] [] Heq; cbn; congruence. Qed. + Definition rotation (r : angle) : Bijection.bijection (Z*Z). - refine {| Bijection.section := mk_rotation r; - Bijection.retraction := mk_rotation (opp_angle r) |}. -Proof. -intros x y. simpl. -abstract (split; intro; subst; destruct r; simpl; destruct x || destruct y; simpl; f_equal; lia). + refine {| Bijection.section := mk_rotation r; + Bijection.retraction := mk_rotation (opp_angle r); + Bijection.section_compat := mk_rotation_compat r |}. +Proof using . +intros x y. cbn. +abstract (split; intro; subst; destruct r; cbn; destruct x || destruct y; cbn; f_equal; lia). Defined. -Instance rotation_compat : Proper (equiv ==> equiv) rotation := reflexive_proper _. +(* #[export] Instance rotation_compat : Proper (equiv ==> equiv) rotation := reflexive_proper _. *) (** Reflection *) Definition mk_reflection r : Z*Z -> Z*Z := @@ -172,36 +161,40 @@ Definition mk_reflection r : Z*Z -> Z*Z := | angle3 => fun x => (- snd x, - fst x) end. +Lemma mk_reflection_compat : forall r, Proper (equiv ==> equiv) (mk_reflection r). +Proof. intros [] [] [] Heq; cbn; congruence. Qed. + Definition reflection (r : angle) : Bijection.bijection (Z*Z). refine {| Bijection.section := mk_reflection r; - Bijection.retraction := mk_reflection r |}. -Proof. -intros x y. simpl. -abstract (split; intro; subst; destruct r; simpl; destruct x || destruct y; simpl; f_equal; lia). + Bijection.retraction := mk_reflection r; + Bijection.section_compat := mk_reflection_compat r |}. +Proof using . +intros x y. cbn. +abstract (split; intro; subst; destruct r; cbn; destruct x || destruct y; cbn; f_equal; lia). Defined. -Instance reflection_compat : Proper (equiv ==> equiv) reflection := reflexive_proper _. +(* #[export] Instance reflection_compat : Proper (equiv ==> equiv) reflection := reflexive_proper _. *) (** *** Change of frame of reference **) (** Translation **) -Instance FCTranslation : frame_choice (Z*Z) := {| +#[export] Instance FCTranslation : frame_choice (Z*Z) := {| frame_choice_bijection := translation; frame_choice_Setoid := _; frame_choice_bijection_compat := _ |}. (** Rigid Motion **) -Instance rigid_motion_compat : +#[export] Instance rigid_motion_compat : Proper (equiv ==> equiv) (fun rm => rotation (snd rm) ∘ translation (fst rm)). -Proof. intros ? ? [Hv Ha]; do 2 f_equiv; assumption. Qed. +Proof using . intros ? ? [Hv Ha]; do 2 f_equiv; assumption. Qed. -Instance FCRigidMotion : frame_choice (Z*Z*angle) := {| +#[export] Instance FCRigidMotion : frame_choice (Z*Z*angle) := {| frame_choice_bijection := fun rm => rotation (snd rm) ∘ translation (fst rm); frame_choice_Setoid := prod_Setoid node_Setoid angle_Setoid; frame_choice_bijection_compat := rigid_motion_compat |}. (** Similarities *) -Instance FCSimilarity : frame_choice (bool*(Z*Z)*angle)%type. +#[export] Instance FCSimilarity : frame_choice (bool*(Z*Z)*angle)%type. simple refine {| frame_choice_bijection := fun '(b, v, a) => rotation a ∘ translation v ∘ (if b : bool then reflection angle0 else @Bijection.id node _); diff --git a/Spaces/Isometry.v b/Spaces/Isometry.v new file mode 100644 index 0000000000000000000000000000000000000000..50f02376ee51c00a64bcd5fdf2432bf174c5c12c --- /dev/null +++ b/Spaces/Isometry.v @@ -0,0 +1,204 @@ +(**************************************************************************) +(* Mechanised Framework for Local Interactions & Distributed Algorithms *) +(* T. Balabonski, P. Courtieu, L. Rieg, X. Urbain *) +(* PACTOLE project *) +(* *) +(* This file is distributed under the terms of the CeCILL-C licence. *) +(* *) +(**************************************************************************) + +(**************************************************************************) +(** Mechanised Framework for Local Interactions & Distributed Algorithms + T. Balabonski, P. Courtieu, L. Rieg, X. Urbain + PACTOLE project + + This file is distributed under the terms of the CeCILL-C licence + *) +(**************************************************************************) + + +Require Import Utf8. +Require Import SetoidDec. +Require Import Rbase Rbasic_fun. +Require Import Lra. +Require Import Pactole.Util.Coqlib. +Require Import Pactole.Util.Bijection. +Require Import Pactole.Spaces.EuclideanSpace. +Require Pactole.Spaces.Similarity. +Set Implicit Arguments. + + +(********************) +(** * Isometries **) +(********************) + +Open Scope R_scope. + +Section IsometryDefinition. + +Context {T : Type}. +Context `{RealMetricSpace T}. + +(** Isometries are functions that preserve distances. + Unlike bijections that only need a setoid, we need here a metric space. *) +Record isometry := { + iso_f :> bijection T; + dist_prop : forall x y, dist (iso_f x) (iso_f y) = dist x y}. + +Global Instance isometry_Setoid : Setoid isometry. +simple refine {| equiv := fun sim1 sim2 => equiv (iso_f sim1) (iso_f sim2) |}. +Proof using . +* apply bij_Setoid. +* split. + + repeat intro. reflexivity. + + repeat intro. now symmetry. + + repeat intro. etransitivity; eauto. +Defined. + +Global Instance f_compat : Proper (equiv ==> equiv) iso_f. +Proof using . intros sim1 sim2 Hsim ?. now apply Hsim. Qed. + +Theorem injective : forall iso : isometry, Preliminary.injective equiv equiv iso. +Proof using . +intros iso z t Heqf. +rewrite <- dist_defined in Heqf |- *. +now rewrite iso.(dist_prop) in Heqf. +Qed. + +(** The identity isometry *) +Definition id : isometry := {| + iso_f := @Bijection.id T _; + dist_prop := ltac:(reflexivity) |}. + +(** Composition of isometries *) +Definition comp (f g : isometry) : isometry. +refine {| iso_f := @compose (Bijection.bijection T) _ _ f g |}. +Proof using . abstract (intros; simpl; now rewrite f.(dist_prop), g.(dist_prop)). Defined. + +Global Instance IsometryComposition : Composition isometry. +refine {| compose := comp |}. +Proof using . intros f1 f2 Hf g1 g2 Hg x. cbn. now rewrite Hf, Hg. Defined. + +Lemma compose_assoc : forall f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. +Proof using . repeat intro. reflexivity. Qed. + +Lemma compose_id_l : forall sim, id ∘ sim == sim. +Proof using . intros sim x. simpl. reflexivity. Qed. + +Lemma compose_id_r : forall sim, sim ∘ id == sim. +Proof using . intros sim x. simpl. reflexivity. Qed. + +(** Inverse of an isometry *) +Definition inv (iso : isometry) : isometry. +refine {| iso_f := inverse iso.(iso_f) |}. +Proof using . +intros x y. +rewrite <- iso.(dist_prop). simpl. now repeat rewrite section_retraction. +Defined. + +Global Instance IsometryInverse : Inverse isometry. +refine {| inverse := inv |}. +Proof using . intros f g Hfg x. simpl. now f_equiv. Defined. + +Lemma compose_inverse_l : forall iso : isometry, (iso â»Â¹ ∘ iso) == id. +Proof using . intros iso x. simpl. now rewrite retraction_section; autoclass. Qed. + +Lemma compose_inverse_r : forall iso : isometry, iso ∘ (iso â»Â¹) == id. +Proof using . intros iso x. simpl. now rewrite section_retraction; autoclass. Qed. + +Lemma inverse_compose : forall f g : isometry, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). +Proof using . intros f g x. simpl. reflexivity. Qed. + +Lemma simpl_inverse_l : forall (iso : isometry) x, iso â»Â¹ (iso x) == x. +Proof using . intros iso x. apply compose_inverse_l. Qed. + +Lemma simpl_inverse_r : forall (iso : isometry) x, iso (iso â»Â¹ x) == x. +Proof using . intros iso x. apply compose_inverse_r. Qed. + +Lemma dist_swap : forall (iso : isometry) x y, dist (isoâ»Â¹ x) y == dist (iso y) x. +Proof using. +intros iso x y. rewrite <- (compose_inverse_r iso x) at 2. +cbn. now rewrite dist_prop, dist_sym. +Qed. + +(** Center of a isometry, that is, the point that gets mapped to the origin. *) +Definition center (iso : isometry) : T := isoâ»Â¹ origin. + +Lemma center_prop : forall iso : isometry, iso (center iso) == origin. +Proof using . intro. unfold center. apply compose_inverse_r. Qed. + +Global Instance center_compat : Proper (equiv ==> equiv) center. +Proof using . intros iso ? Hiso. apply (injective iso). now rewrite center_prop, Hiso, center_prop. Qed. + +Definition isometry_similarity (iso : isometry) : Similarity.similarity T := {| + Similarity.sim_f := iso_f iso; + Similarity.zoom := 1%R; + Similarity.dist_prop := ltac:(intros; rewrite Rmult_1_l; apply dist_prop) |}. + +Definition similarity_isometry (sim : Similarity.similarity T) + (Hzoom : Similarity.zoom sim = 1) : isometry := {| + iso_f := Similarity.sim_f sim; + dist_prop := ltac:(now intros; rewrite Similarity.dist_prop, Hzoom, Rmult_1_l) |}. + +Instance isometry_similarity_compat : Proper (equiv ==> equiv) isometry_similarity. +Proof using . intros ? ? Heq. apply Heq. Qed. + +Lemma similarity_isometry_compat : forall sim1 sim2 (Hsim : sim1 == sim2) Hsim1 Hsim2, + similarity_isometry sim1 Hsim1 == similarity_isometry sim2 Hsim2. +Proof using . intros ? ? Heq ? ?. apply Heq. Qed. + +Lemma isometry_similarity_zoom : forall iso, Similarity.zoom (isometry_similarity iso) = 1. +Proof using . reflexivity. Qed. + +Lemma isometry_similarity_inv : forall iso Hsim, + similarity_isometry (isometry_similarity iso) Hsim == iso. +Proof using . intros iso Hiso x. reflexivity. Qed. + +Lemma similarity_isometry_inv : forall sim Hsim, + isometry_similarity (similarity_isometry sim Hsim) == sim. +Proof using . intros sim Hsim x. reflexivity. Qed. + +(* TODO: prove that isometries preserve barycenters *) +End IsometryDefinition. +Arguments isometry T {_} {_} {_} {_}. + + +Section Translation. +Context {T : Type}. +Context `{rnsT : RealNormedSpace T}. + +(** The translation isometry *) +Lemma bij_translation_Inversion : forall v x y : T, add x v == y ↔ add y (opp v) == x. +Proof using . +intros. split; intro Heq; rewrite Heq || rewrite <- Heq; rewrite <- add_assoc. +- now rewrite add_opp, add_origin. +- setoid_rewrite add_comm at 2. now rewrite add_opp, add_origin. +Qed. + +Definition bij_translation (v : T) : @bijection T _. +refine {| + section := fun x => add x v; + retraction := fun x => add x (opp v) |}. +Proof using VS. ++ now repeat intro; apply add_compat. ++ apply bij_translation_Inversion. +Defined. + +Lemma translation_zoom : forall v x y : T, dist (add x v) (add y v) = 1 * dist x y. +Proof using . intros. ring_simplify. apply dist_translation. Qed. + +Definition translation (v : T) : isometry T. +refine {| iso_f := bij_translation v |}. +Proof using . cbn -[dist]. abstract (now intros; rewrite dist_translation). Defined. + +Global Instance translation_compat : Proper (equiv ==> equiv) translation. +Proof using . intros u v Huv x. simpl. now rewrite Huv. Qed. + +Lemma translation_origin : translation origin == id. +Proof using . intro. simpl. now rewrite add_origin. Qed. + +Lemma translation_inverse : forall t, inverse (translation t) == translation (opp t). +Proof using . intros t x. simpl. reflexivity. Qed. + +End Translation. + diff --git a/Spaces/Isomorphism.v b/Spaces/Isomorphism.v index c76a707bc0ed940dfea5f984a9fe6c7a531347de..b4133a6e68f32cc043d6625f3ac74d0e7953dde9 100644 --- a/Spaces/Isomorphism.v +++ b/Spaces/Isomorphism.v @@ -32,22 +32,18 @@ Context {G : Graph V E}. Record isomorphism := { iso_V :> bijection V; iso_E : bijection E; - iso_T : bijection R; - iso_morphism : forall e, iso_V (src e) == src (iso_E e) - /\ iso_V (tgt e) == tgt (iso_E e); - iso_threshold : forall e, iso_T (threshold e) = threshold (iso_E e); - iso_incr : forall a b, (a < b)%R -> (iso_T a < iso_T b)%R; - iso_bound_T : forall r, (0 < iso_T r < 1)%R <-> (0 < r < 1)%R }. + iso_morphism : ∀ e, iso_V (src e) == src (iso_E e) + /\ iso_V (tgt e) == tgt (iso_E e) }. Global Instance isomorphism_Setoid : Setoid isomorphism. -simple refine {| equiv := fun iso1 iso2 => iso1.(iso_V) == iso2.(iso_V) - /\ iso1.(iso_E) == iso2.(iso_E) - /\ iso1.(iso_T) == iso2.(iso_T) |}; autoclass. -Proof. split. -+ intro f. now repeat split. -+ intros f g Hfg; destruct Hfg as [HV [HE HT]]. repeat split; intro; now symmetry. -+ intros f g h Hfg Hgh. destruct Hfg as [? [? ?]], Hgh as [? [? ?]]. - repeat split; intro; etransitivity; eauto. +Proof using . + simple refine {| + equiv := λ iso1 iso2, iso1.(iso_V) == iso2.(iso_V) + /\ iso1.(iso_E) == iso2.(iso_E) |}; autoclass. split. + + intro f. now repeat split. + + intros f g Hfg; destruct Hfg as [HV HE]. split; now symmetry. + + intros f g h Hfg Hgh. destruct Hfg as [? ?], Hgh as [? ?]. + split; etransitivity; eauto. Defined. Instance iso_V_compat : Proper (equiv ==> equiv) iso_V. @@ -56,33 +52,26 @@ Proof using . intros ? ? Heq ?. now apply Heq. Qed. Instance iso_E_compat : Proper (equiv ==> equiv) iso_E. Proof using . intros ? ? Heq ?. now apply Heq. Qed. -Instance iso_T_compat : Proper (equiv ==> equiv) iso_T. -Proof using . intros ? ? Heq ?. now apply Heq. Qed. - +Lemma equiv_iso_V_to_iso_E iso1 iso2: + iso_V iso1 == iso_V iso2 -> iso_E iso1 == iso_E iso2. +Proof using . + intros Heqiso_V e. apply simple_graph. rewrite <-! (proj1 (iso_morphism _ _)), + <-! (proj2 (iso_morphism _ _)), <- Heqiso_V. split. all: reflexivity. +Qed. Definition id : isomorphism. -refine {| iso_V := @id V _; - iso_E := @id E _; - iso_T := @id R _ |}. -Proof. -+ now intros. -+ now intros. -+ now intros. -+ now intros. +Proof using . + refine {| iso_V := id; + iso_E := id |}. + now intros. Defined. - Definition comp (f g : isomorphism) : isomorphism. -refine {| - iso_V := compose f.(iso_V) g.(iso_V); - iso_E := compose f.(iso_E) g.(iso_E); - iso_T := compose f.(iso_T) g.(iso_T) |}. -Proof. -+ intro. simpl. - split; now rewrite <- 2 (proj1 (iso_morphism _ _)) || rewrite <- 2 (proj2 (iso_morphism _ _)). -+ intro. simpl. now rewrite 2 iso_threshold. -+ intros. simpl. now do 2 apply iso_incr. -+ intro. simpl. now rewrite 2 iso_bound_T. +Proof using . + refine {| + iso_V := compose f.(iso_V) g.(iso_V); + iso_E := compose f.(iso_E) g.(iso_E) |}. + intro. simpl. split; now rewrite <- 2 (proj1 (iso_morphism _ _)) || rewrite <- 2 (proj2 (iso_morphism _ _)). Defined. Global Instance IsoComposition : Composition isomorphism. @@ -92,72 +81,63 @@ Proof. intros f1 f2 Hf g1 g2 Hg. repeat split; intro; simpl; now rewrite Hf, Hg. (* Global Instance compose_compat : Proper (equiv ==> equiv ==> equiv) compose. Proof. intros f1 f2 Hf g1 g2 Hg. repeat split; intro; simpl; now rewrite Hf, Hg. Qed. *) -Lemma compose_assoc : forall f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. -Proof using . intros f g h; repeat split; simpl; reflexivity. Qed. +Lemma compvE : ∀ iso1 iso2 : isomorphism, + iso_V (iso1 ∘ iso2) = iso_V iso1 ∘ iso_V iso2 :> bijection V. +Proof using . reflexivity. Qed. + +Lemma compeE : ∀ iso1 iso2 : isomorphism, + iso_E (iso1 ∘ iso2) = iso_E iso1 ∘ iso_E iso2 :> bijection E. +Proof using . reflexivity. Qed. -Set Printing Implicit. +Lemma compose_assoc : ∀ f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. +Proof using . intros f g h; repeat split; simpl; reflexivity. Qed. Definition inv (iso : isomorphism) : isomorphism. +Proof using . refine {| iso_V := inverse iso.(iso_V); - iso_E := inverse iso.(iso_E); - iso_T := inverse iso.(iso_T) - |}. -Proof. -+ intro. simpl. rewrite <- 2 Inversion, (proj1 (iso_morphism _ _)), (proj2 (iso_morphism _ _)). + iso_E := inverse iso.(iso_E) |}. + intro. simpl. rewrite <- 2 Inversion, (proj1 (iso_morphism _ _)), (proj2 (iso_morphism _ _)). split; apply src_compat || apply tgt_compat; now rewrite Inversion. -+ intro. simpl. change eq with (@equiv R _). rewrite <- Inversion, iso_threshold. - apply threshold_compat. now rewrite Inversion. -+ intros a b Hab. - simpl. - assert (Hincr := iso_incr iso). - assert (forall x y, @retraction R _ (iso_T iso) x < @retraction R _ (iso_T iso) y -> x < y)%R. - { intros. - specialize (Hincr (@retraction R _ (iso_T iso) x) (@retraction R _ (iso_T iso) y) H). - now do 2 rewrite section_retraction in Hincr. } - assert (Hnondecr: - (forall x y, x <= y -> @retraction R _ (iso_T iso) x <= @retraction R _ (iso_T iso) y)%R). - { intros x y Hle. apply Rnot_lt_le. apply Rle_not_lt in Hle. intuition. } - destruct (Hnondecr a b) as [| Heq]; auto; intuition; []. - apply (f_equal (iso_T iso)) in Heq. rewrite 2 section_retraction in Heq. lra. -+ intro. simpl. - assert (Hbound := iso_bound_T iso). specialize (Hbound (@retraction R _ (iso_T iso) r)). - now rewrite section_retraction in Hbound. Defined. Global Instance IsoInverse : Inverse isomorphism. -refine {| inverse := inv |}. -Proof. -intros f g [? [? ?]]. -repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. +Proof using . + refine {| inverse := inv |}. + intros f g [? ?]. repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. Defined. +Lemma id_inv : idâ»Â¹ == id. +Proof using . split. all: setoid_rewrite id_inv. all: reflexivity. Qed. + +Lemma id_comp_l : ∀ iso : isomorphism, id ∘ iso == iso. +Proof using. intros. split. all: setoid_rewrite id_comp_l. all: reflexivity. Qed. + +Lemma id_comp_r : ∀ iso : isomorphism, iso ∘ id == iso. +Proof using. intros. split. all: setoid_rewrite id_comp_r. all: reflexivity. Qed. + +Lemma inv_inv : ∀ iso : isomorphism, isoâ»Â¹â»Â¹ == iso. +Proof using . intros. split. all: setoid_rewrite inv_inv. all: reflexivity. Qed. + (* Global Instance inverse_compat : Proper (equiv ==> equiv) inverse. Proof. intros f g [? [? ?]]. repeat split; intro; simpl; change eq with (@equiv R _); f_equiv; auto. Qed. *) -Lemma compose_inverse_l : forall iso : isomorphism, iso â»Â¹ ∘ iso == id. +Lemma compose_inverse_l : ∀ iso : isomorphism, iso â»Â¹ ∘ iso == id. Proof using . intro. repeat split; intro; simpl; try now rewrite retraction_section; autoclass. Qed. -Lemma compose_inverse_r : forall iso : isomorphism, iso ∘ (iso â»Â¹) == id. +Lemma compose_inverse_r : ∀ iso : isomorphism, iso ∘ (iso â»Â¹) == id. Proof using . intro. repeat split; intro; simpl; try now rewrite section_retraction; autoclass. Qed. -Lemma inverse_compose : forall f g : isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). +Lemma inverse_compose : ∀ f g : isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). Proof using . intros f g; repeat split; intro; simpl; reflexivity. Qed. -Lemma injective : forall iso : isomorphism, Preliminary.injective equiv equiv iso. -Proof using . -intros f x y Heq. transitivity (id x); try reflexivity; []. -rewrite <- (compose_inverse_l f). simpl. rewrite Heq. -now apply compose_inverse_l. -Qed. - End Isomorphism. Arguments isomorphism {V} {E} G. -Lemma find_edge_iso_Some `{G : Graph} : forall (iso : isomorphism G) src tgt e, +Lemma find_edge_iso_Some `{G : Graph} : ∀ (iso : isomorphism G) src tgt e, @find_edge _ _ G (iso src) (iso tgt) == Some (iso.(iso_E) e) <-> @find_edge _ _ G src tgt == Some e. Proof using . @@ -174,7 +154,7 @@ revert iso src tgt e. apply strong_and. simpl in Hstep. now rewrite 3 Bijection.retraction_section in Hstep. Qed. -Lemma find_edge_iso_None `{G : Graph} : forall (iso : isomorphism G) src tgt, +Lemma find_edge_iso_None `{G : Graph} : ∀ (iso : isomorphism G) src tgt, @find_edge _ _ G (iso src) (iso tgt) == None <-> @find_edge _ _ G src tgt == None. Proof using . intros iso src tgt. destruct (find_edge src tgt) eqn:Hcase. diff --git a/Spaces/R.v b/Spaces/R.v index f1e51f79c1e98b217b457c75b08a720f4190a40a..6a9803b480a6d37e6644f62c5ec9c533c2139de1 100644 --- a/Spaces/R.v +++ b/Spaces/R.v @@ -20,12 +20,11 @@ Require Import Bool. -Require Import Arith.Div2. Require Import Lia. Require Export Rbase Rbasic_fun. Require Import Sorting. Require Import List. -Require Import Relations. +Require Import Relations RelationClasses. Require Import RelationPairs. Require Import SetoidDec. Require Import Pactole.Util.Coqlib. @@ -37,6 +36,8 @@ Import Permutation. Set Implicit Arguments. Open Scope R_scope. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Typeclasses eauto := (bfs). @@ -62,7 +63,7 @@ Proof. * intro. nra. * intro u. split; intro Heq. + apply Rmult_integral in Heq. now destruct Heq. - + now rewrite Heq, Rmult_0_l. + + cbn in Heq. subst. apply Rmult_0_l. Defined. Lemma norm_R : forall x, norm x = Rabs x. @@ -77,29 +78,29 @@ Ltac Rdec := repeat match goal with | |- context[@equiv_dec _ _ R_EqDec ?x ?x] => let Heq := fresh "Heq" in destruct (@equiv_dec _ _ R_EqDec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] + [clear Heq | exfalso; contradiction Heq; reflexivity] | |- context[@equiv_dec _ _ R_EqDec 1 0] => let Heq := fresh "Heq" in destruct (@equiv_dec _ _ R_EqDec 1 0) as [Heq | Heq]; - [now elim R1_neq_R0 | clear Heq] + [now contradiction R1_neq_R0 | clear Heq] | |- context[@equiv_dec _ _ R_EqDec 0 1] => let Heq := fresh "Heq" in destruct (@equiv_dec _ _ R_EqDec 0 1) as [Heq | Heq]; - [now symmetry in Heq; elim R1_neq_R0 | clear Heq] + [now symmetry in Heq; contradiction R1_neq_R0 | clear Heq] | H : context[@equiv_dec _ _ R_EqDec ?x ?x] |- _ => let Heq := fresh "Heq" in destruct (@equiv_dec _ _ R_EqDec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] + [clear Heq | exfalso; contradiction Heq; reflexivity] | |- context[Rdec ?x ?x] => let Heq := fresh "Heq" in destruct (Rdec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] + [clear Heq | exfalso; contradiction Heq; reflexivity] | |- context[Rdec 1 0] => let Heq := fresh "Heq" in destruct (Rdec 1 0) as [Heq | Heq]; - [now elim R1_neq_R0 | clear Heq] + [now contradiction R1_neq_R0 | clear Heq] | |- context[Rdec 0 1] => let Heq := fresh "Heq" in destruct (Rdec 0 1) as [Heq | Heq]; - [now symmetry in Heq; elim R1_neq_R0 | clear Heq] + [now symmetry in Heq; contradiction R1_neq_R0 | clear Heq] | H : context[Rdec ?x ?x] |- _ => let Heq := fresh "Heq" in destruct (Rdec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] - | H : ?x <> ?x |- _ => elim H; reflexivity + [clear Heq | exfalso; contradiction Heq; reflexivity] + | H : ?x <> ?x |- _ => contradiction H; reflexivity end. Ltac Rdec_full := @@ -115,7 +116,7 @@ Ltac Rdec_full := Ltac Rabs := match goal with - | Hx : ?x <> ?x |- _ => now elim Hx + | Hx : ?x <> ?x |- _ => now contradiction Hx | Heq : ?x == ?y, Hneq : ?y =/= ?x |- _ => symmetry in Heq; contradiction | Heq : ?x == ?y, Hneq : ?y <> ?x |- _ => symmetry in Heq; contradiction | Heq : ?x = ?y, Hneq : ?y =/= ?x |- _ => symmetry in Heq; contradiction @@ -132,7 +133,7 @@ Ltac Rle_dec := end. Global Instance Leibniz_fun_compat : forall f : R -> R, Proper (equiv ==> equiv) f. -Proof. intros f ? ? Heq. now rewrite Heq. Qed. +Proof. intros f ? ? Heq. cbn in *. now rewrite Heq. Qed. (** A location is determined by distances to 2 points. *) Lemma dist_case : forall x y, dist x y = x - y \/ dist x y = y - x. @@ -165,14 +166,14 @@ Definition Rleb (x y : R) := if Rle_lt_dec x y then true else false. Lemma Rleb_spec : forall x y, Rleb x y = true <-> Rle x y. Proof. intros x y; unfold Rleb; destruct (Rle_lt_dec x y); split; intro H; trivial. -inversion H. elim (Rlt_not_le _ _ r H). +inversion H. contradiction (Rlt_not_le _ _ r H). Qed. Corollary Rleb_total : forall x y, Rleb x y = true \/ Rleb y x = true. Proof. intros x y. unfold Rleb. destruct (Rle_lt_dec x y). now left. - right. destruct (Rle_lt_dec y x). reflexivity. elim (Rlt_irrefl x). now apply Rlt_trans with y. + right. destruct (Rle_lt_dec y x). reflexivity. contradiction (Rlt_irrefl x). now apply Rlt_trans with y. Qed. Corollary Rleb_trans : Transitive Rleb. @@ -209,7 +210,7 @@ apply (StronglySorted_ind (fun l => forall l', StronglySorted Rleb l' -> Permuta now apply Permutation_cons_inv with b. Qed. -Instance sort_uniq : Proper (@Permutation R ==> eq) sort. +Global Instance sort_uniq : Proper (@Permutation R ==> eq) sort. Proof. intros l l' Hl. apply StronglySorted_uniq. - apply StronglySorted_sort. exact Rleb_trans. @@ -217,7 +218,7 @@ intros l l' Hl. apply StronglySorted_uniq. - transitivity l. symmetry. apply Permuted_sort. transitivity l'. assumption. apply Permuted_sort. Qed. -Instance sort_uniqA : Proper (PermutationA eq ==> eq) sort. +Global Instance sort_uniqA : Proper (PermutationA eq ==> eq) sort. Proof. intros ? ?. rewrite PermutationA_Leibniz. apply sort_uniq. Qed. Corollary StronglySorted_sort_identity : forall l, StronglySorted Rleb l -> sort l = l. @@ -406,10 +407,13 @@ Hint Immediate injective sim_Minjective : core. Corollary similarity_monotonic : forall sim : similarity R, monotonic Rleb Rleb sim. Proof. intro sim. destruct (similarity_in_R_case sim) as [Hinc | Hdec]. -+ left. intros x y Hxy. do 2 rewrite Hinc. apply similarity_increasing; trivial. - pose (Hratio := zoom_pos sim). lra. -+ right. intros x y Hxy. do 2 rewrite Hdec. apply similarity_decreasing; trivial. - assert (Hratio := zoom_pos sim). lra. ++ left. intros x y Hxy. cbn in Hinc. do 2 rewrite Hinc. + apply similarity_increasing; trivial; []. apply Rlt_le, zoom_pos. ++ right. intros x y Hxy. cbn in Hdec. do 2 rewrite Hdec. + apply similarity_decreasing; trivial; []. + (* TODO Coq 8.17 : does [lra] work again? *) + rewrite <- (Rplus_opp_r (zoom sim)), <- Rminus_0_l. + unfold Rminus. apply Rplus_le_compat; try lra; []. apply Rlt_le, zoom_pos. Qed. (** To conclude that two similarities are equal, it is enough to show that they are equal on two points. *) @@ -424,8 +428,12 @@ assert (Hzoom : zoom sim1 = zoom sim2). apply Rmult_eq_reg_r with (dist pt1 pt2); trivial; []. now rewrite <- 2 dist_prop, H1, H2. } assert (Hk : k1 = k2 \/ k1 = - k2). -{ destruct Hk1, Hk2; subst; rewrite Hzoom, ?Ropp_involutive; tauto. } -assert (k2 <> 0). { generalize (zoom_pos sim2). lra. } +{ destruct Hk1, Hk2; subst; rewrite ?Ropp_involutive; intuition. } +(* assert (k2 <> 0). { generalize (zoom_pos sim2). lra. } *) +assert (k2 <> 0). +{ destruct Hk2; subst k2. + - symmetry. apply Rlt_not_eq, zoom_pos. + - apply Rlt_not_eq, Ropp_lt_gt_0_contravar, Rlt_gt, zoom_pos. } rewrite Hsim1, Hsim2 in *. destruct Hk; subst k1. + (* Having same factor, they also have same center *) @@ -448,30 +456,28 @@ simple refine {| sim_f := {| section := fun x => pt3 + (pt4 - pt3) / (pt2 - pt1) Proof. * abstract (intros; simpl in *; split; intro; subst; field; lra). * intros x y. cbn -[dist]. repeat rewrite dist_R. unfold Rdiv. - rewrite <- Rabs_Rinv. - + repeat rewrite <- Rabs_mult. f_equal. ring. - + simpl in *. lra. + rewrite <- Rabs_inv. repeat rewrite <- Rabs_mult. f_equal. ring. Defined. Lemma build_similarity_compat : forall pt1 pt1' pt2 pt2' pt3 pt3' pt4 pt4' (H12 : pt1 =/= pt2) (H34 : pt3 =/= pt4) (H12' : pt1' =/= pt2') (H34' : pt3' =/= pt4'), pt1 == pt1' -> pt2 == pt2' -> pt3 == pt3' -> pt4 == pt4' -> build_similarity H12 H34 == build_similarity H12' H34'. -Proof. intros * Heq1 Heq2 Heq3 Heq4 x. simpl. now rewrite Heq1, Heq2, Heq3, Heq4 in *. Qed. +Proof. intros * Heq1 Heq2 Heq3 Heq4 x. cbn in *. now rewrite Heq1, Heq2, Heq3, Heq4 in *. Qed. Lemma build_similarity_swap : forall pt1 pt2 pt3 pt4 (Hdiff12 : pt1 =/= pt2) (Hdiff34 : pt3 =/= pt4), build_similarity (symmetry Hdiff12) (symmetry Hdiff34) == build_similarity Hdiff12 Hdiff34. -Proof. repeat intro. simpl in *. field. lra. Qed. +Proof. repeat intro. cbn in *. field. lra. Qed. Lemma build_similarity_eq1 : forall pt1 pt2 pt3 pt4 (Hdiff12 : pt1 =/= pt2) (Hdiff34 : pt3 =/= pt4), build_similarity Hdiff12 Hdiff34 pt1 == pt3. -Proof. intros. simpl in *. field. lra. Qed. +Proof. intros. cbn in *. field. lra. Qed. (* This is wrong without the proper orientation *) Lemma build_similarity_eq2 : forall pt1 pt2 pt3 pt4 (Hdiff12 : pt1 =/= pt2) (Hdiff34 : pt3 =/= pt4), build_similarity Hdiff12 Hdiff34 pt2 == pt4. -Proof. intros. simpl in *. field. lra. Qed. +Proof. intros. cbn in *. field. lra. Qed. Lemma build_similarity_inverse : forall pt1 pt2 pt3 pt4 (Hdiff12 : pt1 =/= pt2) (Hdiff34 : pt3 =/= pt4), (build_similarity Hdiff12 Hdiff34)â»Â¹ == build_similarity Hdiff34 Hdiff12. -Proof. repeat intro. simpl in *. field. lra. Qed. +Proof. repeat intro. cbn in *. field. lra. Qed. diff --git a/Spaces/R2.v b/Spaces/R2.v index d423c58c5eca626861574eee2e6e8c9137f95a77..ca086f9ee3a30da67d024c6ff5dc0a6a4d80302a 100644 --- a/Spaces/R2.v +++ b/Spaces/R2.v @@ -26,18 +26,18 @@ Import List Permutation SetoidList. Require Import SetoidDec. Require Import FunInd. Require Import Pactole.Util.Coqlib. +Require Import Pactole.Util.Bijection. Require Export Pactole.Spaces.EuclideanSpace. Require Import Pactole.Spaces.Similarity. Set Implicit Arguments. Open Scope R_scope. -Coercion Bijection.section : Bijection.bijection >-> Funclass. (** ** R² as a Euclidean space over R **) Definition R2 := (R * R)%type. -Instance R2_Setoid : Setoid R2 := {| equiv := @Logic.eq R2 |}. -Instance R2_EqDec : @EqDec R2 _. +Global Instance R2_Setoid : Setoid R2 := {| equiv := @Logic.eq R2 |}. +Global Instance R2_EqDec : @EqDec R2 _. Proof. intros x y. destruct (Rdec (fst x) (fst y)). @@ -125,12 +125,12 @@ Ltac R2dec := repeat match goal with | |- context[@equiv_dec _ _ R2_EqDec ?x ?x] => let Heq := fresh "Heq" in destruct (@equiv_dec _ _ R2_EqDec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] + [clear Heq | exfalso; contradiction Heq; reflexivity] | H : context[Rdec ?x ?x] |- _ => let Heq := fresh "Heq" in destruct (Rdec x x) as [Heq | Heq]; - [clear Heq | exfalso; elim Heq; reflexivity] - | H : ?x <> ?x |- _ => elim H; reflexivity - | H : ?x =/= ?x |- _ => elim H; reflexivity + [clear Heq | exfalso; contradiction Heq; reflexivity] + | H : ?x <> ?x |- _ => contradiction H; reflexivity + | H : ?x =/= ?x |- _ => contradiction H; reflexivity | Heq : ?x == ?y, Hneq : ~?y == ?x |- _ => symmetry in Heq; contradiction | Heq : ?x == ?y, Hneq : ~?x == ?y |- _ => contradiction | Heq : ?x == ?y |- context[@equiv_dec _ _ R2_EqDec ?x ?y] => @@ -182,7 +182,7 @@ Lemma R2dec_bool_false_iff (x y : R2) : R2dec_bool x y = false <-> x =/= y. Proof using . unfold R2dec_bool. destruct (equiv_dec x y); split; discriminate || auto. -intros abs. rewrite e in abs. now elim abs. +intros abs. rewrite e in abs. now contradiction abs. Qed. @@ -193,10 +193,10 @@ Definition orthogonal (u : R2) : R2 := (/(norm u) * (snd u, (- fst u)%R))%VS. Definition colinear (u v : R2) := perpendicular u (orthogonal v). (* Compatibilities *) -Instance orthogonal_compat : Proper (equiv ==> equiv) orthogonal. +Global Instance orthogonal_compat : Proper (equiv ==> equiv) orthogonal. Proof using . intros u v Heq. now rewrite Heq. Qed. -Instance colinear_compat : Proper (equiv ==> equiv ==> iff) colinear. +Global Instance colinear_compat : Proper (equiv ==> equiv ==> iff) colinear. Proof using . intros u1 u2 Hu v1 v2 Hv. now rewrite Hu, Hv. Qed. (** *** Results about [norm] **) @@ -354,10 +354,10 @@ Qed. Lemma colinear_dec : forall u v, {colinear u v} + {~colinear u v}. Proof. intros u v. unfold colinear. apply perpendicular_dec. Defined. -Instance colinear_Reflexive : Reflexive colinear. +Global Instance colinear_Reflexive : Reflexive colinear. Proof using . intro. apply orthogonal_perpendicular. Qed. -Instance colinear_Symmetric : Symmetric colinear. +Global Instance colinear_Symmetric : Symmetric colinear. Proof using . intros u v H. unfold colinear. now rewrite perpendicular_sym, perpendicular_orthogonal_shift. Qed. Lemma colinear_trans : forall u v w, ~equiv v origin -> colinear u v -> colinear v w -> colinear u w. @@ -430,8 +430,7 @@ intros [x1 y1] Hnull [x2 y2]. unfold unitary, orthogonal, norm, inner_product. s replace (Rpow_def.pow y1 2) with y1² by (unfold Rsqr; ring). replace (x2 * (/ sqrt (x1 * x1 + y1 * y1))² * x1² + x2 * (/ sqrt (x1 * x1 + y1 * y1))² * y1²) with (x2 * (/ sqrt (x1 * x1 + y1 * y1))² * (x1² + y1²)) by ring. - rewrite R_sqr.Rsqr_inv; trivial; []. - rewrite Rsqr_sqrt. + rewrite R_sqr.Rsqr_inv'; trivial; []. rewrite Rsqr_sqrt. - unfold Rsqr. field. intro Habs. apply Hnull. rewrite Habs. apply sqrt_0. - replace 0 with (0 + 0) by ring. apply Rplus_le_compat; apply Rle_0_sqr. + ring_simplify. rewrite <- norm_defined in Hnull. unfold norm, inner_product in Hnull. simpl in Hnull. @@ -441,7 +440,7 @@ intros [x1 y1] Hnull [x2 y2]. unfold unitary, orthogonal, norm, inner_product. s replace (Rpow_def.pow y1 2) with y1² by (unfold Rsqr; ring). replace (x2 * (/ sqrt (x1 * x1 + y1 * y1))² * x1² + x2 * (/ sqrt (x1 * x1 + y1 * y1))² * y1²) with (x2 * (/ sqrt (x1 * x1 + y1 * y1))² * (x1² + y1²)) by ring. - rewrite R_sqr.Rsqr_inv; trivial; []. + rewrite R_sqr.Rsqr_inv'; trivial; []. rewrite Rsqr_sqrt. - unfold Rsqr. field. intro Habs. apply Hnull. rewrite Habs. apply sqrt_0. - replace 0 with (0 + 0) by ring. apply Rplus_le_compat; apply Rle_0_sqr. @@ -524,7 +523,7 @@ intros t u v w Heq. null (u - t)%VS; [| null (v - t)%VS]. + split; apply colinear_origin_l || apply colinear_origin_r. + split; try apply colinear_origin_r. rewrite R2sub_origin in Hnull0. rewrite <- Hnull0 in *. - elim Hnull. + contradiction Hnull. rewrite R2sub_origin, <- dist_defined. apply Rmult_eq_reg_l with 2; try lra; []. ring_simplify. apply Rplus_eq_reg_r with (dist v w). rewrite Rplus_0_l. rewrite Heq at 2. setoid_rewrite dist_sym at 3. ring. @@ -966,7 +965,7 @@ destruct (Rlt_le_dec k kh) as [Hlt | Hle]. right. ring. left. apply Rle_lt_trans with (r2 := k). tauto. assumption. - left. apply Rlt_Rminus. assumption. + left. apply Rlt_0_minus. assumption. destruct A, B; compute; f_equal; ring. destruct A, B; compute; f_equal; ring. @@ -1298,7 +1297,7 @@ change ((translation t) (/ k * fold_left add l origin)%VS) rewrite <- (mul_1 t) at 2. rewrite <- (Rinv_l k); trivial; []. rewrite <- mul_morph, <- mul_distr_add. f_equiv. change eq with equiv. subst k. clear Hk. induction l as [| e l]. -* now elim Hl. +* now contradiction Hl. * destruct l as [| e' l']. + destruct e, t. simpl. f_equal; ring. + specialize (IHl ltac:(discriminate)). @@ -1366,7 +1365,7 @@ Proof using . (List.map (fun xn => (Bijection.section (Similarity.sim_f sim) (fst xn), snd xn)) m) init = fold_left (fun acc pt' => acc + snd pt' * (sim.(Similarity.zoom))² * (dist pt (fst pt'))²) m init). { intro pt. induction m as [| p1 m]; intro init. - + now elim Hm. + + now contradiction Hm. + clear Hm. destruct m as [| p2 m]. * cbn -[dist]. now rewrite sim.(Similarity.dist_prop), R_sqr.Rsqr_mult, Rmult_assoc. * remember (p2 :: m) as mm. @@ -1464,12 +1463,12 @@ Proof using . repeat rewrite ?Rle_bool_true_iff, ?Rle_bool_false_iff in * ; repeat progress normalize_R2dist pt1' pt2' pt3' ;try contradiction; repeat match goal with - | H1: ?A < ?A |- _ => elim (Rlt_irrefl _ h_ltxx) + | H1: ?A < ?A |- _ => contradiction (Rlt_irrefl _ h_ltxx) | H1: ?A < ?B, H2: ?B < ?A |- _ => - assert (h_ltxx:A<A) by (eapply Rlt_trans;eauto);elim (Rlt_irrefl _ h_ltxx) + assert (h_ltxx:A<A) by (eapply Rlt_trans;eauto);contradiction (Rlt_irrefl _ h_ltxx) | H1: ?A < ?B, H2: ?B < ?C, H3: ?C < ?A |- _ => assert (h_ltxx:A<C) by (eapply Rlt_trans;eauto); - assert (h_ltxx':A<A) by (eapply Rlt_trans;eauto);elim (Rlt_irrefl _ h_ltxx') + assert (h_ltxx':A<A) by (eapply Rlt_trans;eauto);contradiction (Rlt_irrefl _ h_ltxx') | H1:?A <> ?B, H2: ?A <= ?B |- _ => assert (A<B) by (apply Rle_neq_lt;auto);clear H2 | H1:?A <> ?B, H2: ?B <= ?A |- _ => assert (B<A) by (apply Rle_neq_lt;auto;apply not_eq_sym;auto);clear H2 end; reflexivity. @@ -1507,7 +1506,7 @@ subst; trivial; try contradiction. + do 2 right. subst. repeat split; trivial. intro Heq. rewrite Heq in *. intuition. + repeat match goal with | H : dist _ _ = _ |- _ => rewrite H in *; clear H - | H : ?x <> ?x |- _ => now elim H + | H : ?x <> ?x |- _ => now contradiction H end. + left. now repeat split. Qed. @@ -1789,7 +1788,7 @@ destruct (equiv_dec pt1 pt2) as [Heq12 | Heq12]. rewrite dist_sym. rewrite dist_defined. auto. } rewrite norm_orthogonal; trivial; []. rewrite R_sqr.Rsqr_1, Rmult_1_r. repeat rewrite Rmult_0_r. rewrite Rplus_0_r. - rewrite dist_sym, norm_dist. setoid_rewrite R_sqr.Rsqr_div; try lra. + rewrite dist_sym, norm_dist. setoid_rewrite R_sqr.Rsqr_div'; try lra; []. unfold Rsqr. intro. assert (Hk : k*k = 0) by lra. now apply Rsqr_0_uniq. - unfold middle. destruct pt1, pt2; cbn -[norm]. f_equal; field; @@ -2042,14 +2041,14 @@ Record circle := {center : R2; radius : R}. Definition enclosing_circle (c : circle) l := forall x, In x l -> dist x (center c) <= (radius c). Definition on_circle (c : circle) x := Rdec_bool (dist x (center c)) (radius c). -Instance enclosing_circle_compat : forall c, Proper (@Permutation _ ==> iff) (enclosing_circle c). +Global Instance enclosing_circle_compat : forall c, Proper (@Permutation _ ==> iff) (enclosing_circle c). Proof using . repeat intro. unfold enclosing_circle. do 2 rewrite <- Forall_forall. apply Forall_Permutation_compat; trivial. intros ? ? ?. now subst. Qed. -Instance on_circle_compat : Proper (eq ==> equiv ==> eq) on_circle. +Global Instance on_circle_compat : Proper (eq ==> equiv ==> eq) on_circle. Proof using . repeat intro. cbn in *. now subst. Qed. Lemma on_circle_true_iff : forall c pt, on_circle c pt = true <-> dist pt (center c) = radius c. @@ -2189,20 +2188,20 @@ Axiom SEC_spec2 : forall l c, enclosing_circle c l -> radius (SEC l) <= radius c (** Extra specification in the degenerate case. *) Axiom SEC_nil : radius (SEC nil) = 0. (** Its definition does not depend on the order of points. *) -Declare Instance SEC_compat : Proper (@Permutation _ ==> Logic.eq) SEC. +Global Declare Instance SEC_compat : Proper (@Permutation _ ==> Logic.eq) SEC. Global Instance SEC_compat_bis : Proper (PermutationA Logic.eq ==> Logic.eq) SEC. Proof using . intros ? ? Heq. rewrite PermutationA_Leibniz in Heq. now rewrite Heq. Qed. (* The last axiom is useful because of the following degeneracy fact. *) Lemma enclosing_circle_nil : forall pt r, enclosing_circle {| center := pt; radius := r |} nil. -Proof using . intros. unfold enclosing_circle. intros x Hin. elim Hin. Qed. +Proof using . intros. unfold enclosing_circle. intros x Hin. contradiction Hin. Qed. Definition center_eq c1 c2 := c1.(center) = c2.(center). Definition radius_eq c1 c2 := c1.(radius) = c2.(radius). (** Unicity proof of the radius of the SEC *) -Instance SEC_radius_compat : +Global Instance SEC_radius_compat : Proper (@Permutation _ ==> center_eq) SEC -> Proper (@Permutation _ ==> radius_eq) SEC. Proof using . intros Hspec l1 l2 Hperm. @@ -2225,7 +2224,7 @@ Qed. (** Points on the SEC. *) Definition on_SEC l := List.filter (on_circle (SEC l)) l. -Instance on_SEC_compat : Proper (PermutationA Logic.eq ==> PermutationA Logic.eq) on_SEC. +Global Instance on_SEC_compat : Proper (PermutationA Logic.eq ==> PermutationA Logic.eq) on_SEC. Proof using . intros l1 l2 Hl. unfold on_SEC. rewrite Hl at 2. rewrite filter_extensionality_compat; try reflexivity. @@ -2264,7 +2263,7 @@ Lemma max_dist_le : forall pt x l, In x l -> dist x pt <= max_dist pt l. Proof using . intros pt x l Hin. unfold max_dist. generalize 0. induction l as [| e l]; intro acc; simpl. -* elim Hin. +* contradiction Hin. * destruct Hin as [? | Hin]; subst. + apply Rle_trans with (Rmax acc (dist x pt)). - apply Rmax_r. @@ -2275,7 +2274,7 @@ Qed. Lemma max_dist_exists : forall pt l, l <> nil -> exists x, In x l /\ dist x pt = max_dist pt l. Proof using . intros pt l Hl. induction l as [| e1 l]. -* now elim Hl. +* now contradiction Hl. * destruct l as [| e2 l]. + exists e1. split; try now left. unfold max_dist. simpl. symmetry. apply Rmax_right. change (0 <= dist e1 pt). apply dist_nonneg. @@ -2429,12 +2428,12 @@ Lemma farthest_from_in_except_In : forall exc c l, (exists pt, pt <> exc /\ In p In (farthest_from_in_except exc c c l) l. Proof using . intros exc c l Hl. induction l as [| e l]. -* now elim Hl. +* destruct Hl as [? H]. apply H. * cbn. destruct (equiv_dec e exc) as [Heq | Heq]. + rewrite Heq in *. destruct l. - - destruct Hl as [pt' [Habs Hin]]. elim Habs. now inversion Hin. + - destruct Hl as [pt' [Habs Hin]]. contradiction Habs. now inversion Hin. - right. apply IHl. destruct Hl as [pt' [Hneq Hin]]. exists pt'. split; trivial. - inversion Hin; subst; trivial; now elim Hneq. + inversion Hin; subst; trivial; now contradiction Hneq. + destruct (Rle_dec (dist e c) (dist c c)) as [H | H]. - assert (He : equiv e c). { rewrite <- dist_defined. transitivity (dist c c). @@ -2627,7 +2626,7 @@ Qed. Lemma on_SEC_singleton : forall pt, on_SEC (pt :: nil) = pt :: nil. Proof using . intro. cbn. rewrite SEC_singleton. unfold on_circle. cbn. rewrite R2_dist_defined_2. -destruct (Rdec_bool 0 0) eqn:Htest; trivial. rewrite Rdec_bool_false_iff in Htest. now elim Htest. +destruct (Rdec_bool 0 0) eqn:Htest; trivial. rewrite Rdec_bool_false_iff in Htest. now contradiction Htest. Qed. Lemma on_SEC_singleton_is_singleton : forall pt l, NoDup l -> on_SEC l = pt :: nil -> l = pt :: nil. @@ -2688,7 +2687,7 @@ destruct (equiv_dec pt1 pt2) as [Heq | Heq]. + assert (Hperm : exists l, Permutation l' (pt2 :: l)). { rewrite <- InA_Leibniz in Hpt2. setoid_rewrite <- PermutationA_Leibniz. apply PermutationA_split; autoclass. - inversion_clear Hpt2; trivial. subst. now elim Heq. } + inversion_clear Hpt2; trivial. subst. now contradiction Heq. } destruct Hperm as [l Hperm]. rewrite Hperm in *. clear Hpt2 Hperm l'. change (/2 * dist pt1 pt2) with (radius {| center := middle pt1 pt2; radius := /2 * dist pt1 pt2 |}). rewrite <- SEC_dueton. apply SEC_incl_compat. intro. cbn. intuition. @@ -2755,7 +2754,7 @@ destruct Hl as [Hl | [[pt1 [Hl Hnil]] | [pt1 [pt2 [Hneq [Hpt1 Hpt2]]]]]]. { assert (Hlen : (length l <> 0)%nat) by auto using length_0. rewrite Hl. clear Hl Hnil. induction (length l) as [| [| n]]. - + now elim Hlen. + + now contradiction Hlen. + simpl. apply SEC_dueton. + assert (Hperm : Permutation (pt :: alls pt1 (S (S n))) (pt1 :: pt :: alls pt1 (S n))) by (simpl; constructor). @@ -2846,7 +2845,7 @@ destruct Hl as [Hl | [[pt1 [Hl Hnil]] | [pt1 [pt2 [Hneq [Hpt1 Hpt2]]]]]]. - apply Fourier_util.Rle_mult_inv_pos; lra. - unfold Rdiv. intro Habs. symmetry in Habs. apply Rmult_integral in Habs. assert (Hlt_inv_d := Rinv_0_lt_compat _ Hlt_d). - destruct Habs as [? | Habs]; lra || rewrite Rinv_mult_distr in Habs; lra. } + destruct Habs as [? | Habs]; lra || rewrite Rinv_mult in Habs; lra. } assert (Hdist : dist câ‚‚ c₃ = ε / 2). { unfold c₃, ratio. rewrite <- add_origin at 1. setoid_rewrite add_comm. rewrite dist_translation, dist_sym. rewrite norm_dist, opp_origin, add_origin. @@ -3108,8 +3107,8 @@ destruct (equiv_dec pt1 pt2) as [Heq12 | Heq12]; assert (Hpt2' := Hincl pt2' ltac:(intuition)). simpl in Hpt1', Hpt2'. decompose [or] Hpt1'; decompose [or] Hpt2'; clear Hpt1' Hpt2'; repeat subst; try match goal with - | H : False |- _ => elim H - | H : ~equiv ?x ?x |- _ => elim H; reflexivity + | H : False |- _ => contradiction H + | H : ~equiv ?x ?x |- _ => contradiction H; reflexivity end. - exists pt3. do 4 constructor. - exists pt2. do 4 constructor. @@ -3253,7 +3252,7 @@ repeat match goal with | H : equiv _ _ |- _ => rewrite H in * | H : InA _ _ _ |- _ => inversion_clear H end. -- now elim Hdiff. +- now contradiction Hdiff. - rewrite R2_dist_defined_2 in *. symmetry in Heq2. rewrite dist_defined in Heq2. now symmetry in Heq2. - rewrite R2_dist_defined_2 in *. now rewrite dist_defined in Heq1. Qed. diff --git a/Spaces/RealMetricSpace.v b/Spaces/RealMetricSpace.v index 3d8336015bc0bf97098781a2b0c32a7ef97ecc56..7e5c1a03d24f7e78aa33db6412ec61e432699908 100644 --- a/Spaces/RealMetricSpace.v +++ b/Spaces/RealMetricSpace.v @@ -22,9 +22,9 @@ Class RealMetricSpace (T : Type) {S : Setoid T} `{@EqDec T S} {VS : RealVectorSp dist_sym : forall u v, dist v u = dist u v; triang_ineq : forall u v w, (dist u w <= dist u v + dist v w)%R}. -Arguments dist T%type _ _ _ _ u%VS v%VS. +Arguments dist T%_type _ _ _ _ u%_VS v%_VS. -Instance dist_compat `{RealMetricSpace} : Proper (equiv ==> equiv ==> Logic.eq) dist. +Global Instance dist_compat `{RealMetricSpace} : Proper (equiv ==> equiv ==> Logic.eq) dist. Proof. intros x x' Hx y y' Hy. apply Rle_antisym. + replace (dist x' y') with (0 + dist x' y' + 0)%R by ring. symmetry in Hy. @@ -41,7 +41,7 @@ Lemma dist_nonneg `{RealMetricSpace} : forall u v, (0 <= dist u v)%R. Proof. intros x y. apply Rmult_le_reg_l with 2%R. + apply Rlt_R0_R2. -+ do 2 rewrite double. rewrite Rplus_0_r. ++ do 2 rewrite <- Rplus_diag. rewrite Rplus_0_r. assert (Hx : equiv x x) by reflexivity. rewrite <- dist_defined in Hx. rewrite <- Hx. setoid_rewrite dist_sym at 3. apply triang_ineq. Qed. diff --git a/Spaces/RealNormedSpace.v b/Spaces/RealNormedSpace.v index c49a937ff02f5e72604248ce3e3384c86afb70cc..3547ef0f6c3843fb765d5bebda2b73411c8d737e 100644 --- a/Spaces/RealNormedSpace.v +++ b/Spaces/RealNormedSpace.v @@ -30,7 +30,7 @@ Class RealNormedSpace (T : Type) {S : Setoid T} {EQ : @EqDec T S} {VS : RealVect triang_ineq : forall u v, (norm (add u v) <= norm u + norm v)%R}. Global Existing Instance norm_compat. -Arguments norm T%type _ _ _ _ u%VS. +Arguments norm T%_type _ _ _ _ u%_VS. Notation "∥ u ∥" := (norm u) : VectorSpace_scope. (** *** Proofs of derivable properties about RealNormedSpace **) @@ -379,16 +379,16 @@ Section BarycenterResults. Proof using . intros E dm Hnotempty Hdm p Hp. assert (Hlength_pos: 0 < INR (List.length E)). - { apply lt_0_INR. destruct E; try (now elim Hnotempty); []. simpl. lia. } + { apply lt_0_INR. destruct E; try (now contradiction Hnotempty); []. simpl. lia. } rewrite norm_dist. subst. unfold isobarycenter. setoid_replace p%VS with (- / INR (length E) * (- INR (length E) * p))%VS - by (rewrite mul_morph, Ropp_inv_permute, <- Rinv_l_sym, mul_1; lra || reflexivity). + by (rewrite mul_morph, <- Rinv_opp, <- Rinv_l_sym, mul_1; lra || reflexivity). rewrite <- minus_morph, <- mul_distr_add, norm_mul, Rabs_Ropp, Rabs_right; try (now apply Rle_ge, Rlt_le, Rinv_0_lt_compat); []. apply Rmult_le_reg_l with (r := INR (length E)); trivial; []. rewrite <- Rmult_assoc, Rinv_r, Rmult_1_l; try lra; []. induction E as [| a [| b E]]. - + now elim Hnotempty. + + now contradiction Hnotempty. + specialize (Hp a ltac:(now left)). cbn -[mul add norm]. setoid_replace ((-(1) * p) + (0 + a))%VS with (a - p)%VS by now rewrite add_origin_l, add_comm, minus_morph, mul_1. @@ -421,4 +421,4 @@ Section BarycenterResults. End BarycenterResults. -Arguments unitary {T%type} {_} {_} {_} {_} u%VS. +Arguments unitary {T%_type} {_} {_} {_} {_} u%_VS. diff --git a/Spaces/RealVectorSpace.v b/Spaces/RealVectorSpace.v index 565bf57b4d2abf11b29b470f49413cd36eb0303f..53c8c624460554a082c3a9bfa6666ff3aaf26c94 100644 --- a/Spaces/RealVectorSpace.v +++ b/Spaces/RealVectorSpace.v @@ -51,17 +51,17 @@ Global Existing Instance opp_compat. Declare Scope VectorSpace_scope. Delimit Scope VectorSpace_scope with VS. -Arguments add T%type _ _ _ u%VS v%VS. -Arguments mul T%type _ _ _ k%R u%VS. -Arguments opp T%type _ _ _ u%VS. -Arguments add_assoc {T} {_} {_} {_} u%VS v%VS w%VS. -Arguments add_comm {T} {_} {_} {_} u%VS v%VS. -Arguments add_origin {T} {_} {_} {_} u%VS. -Arguments add_opp {T} {_} {_} {_} u%VS. -Arguments mul_distr_add {T} {_} {_} {_} a%R u%VS v%VS. -Arguments mul_morph {T} {_} {_} {_} a%R b%R u%VS. -Arguments add_morph {T} {_} {_} {_} a%R b%R u%VS. -Arguments mul_1 {T} {_} {_} {_} u%VS. +Arguments add T%_type _ _ _ u%_VS v%_VS. +Arguments mul T%_type _ _ _ k%_R u%_VS. +Arguments opp T%_type _ _ _ u%_VS. +Arguments add_assoc {T} {_} {_} {_} u%_VS v%_VS w%_VS. +Arguments add_comm {T} {_} {_} {_} u%_VS v%_VS. +Arguments add_origin {T} {_} {_} {_} u%_VS. +Arguments add_opp {T} {_} {_} {_} u%_VS. +Arguments mul_distr_add {T} {_} {_} {_} a%_R u%_VS v%_VS. +Arguments mul_morph {T} {_} {_} {_} a%_R b%_R u%_VS. +Arguments add_morph {T} {_} {_} {_} a%_R b%_R u%_VS. +Arguments mul_1 {T} {_} {_} {_} u%_VS. Notation "0" := (origin) : VectorSpace_scope. Notation "u + v" := (add u v) : VectorSpace_scope. @@ -146,14 +146,14 @@ Proof using . intros k k' u Hu Heq. destruct (Rdec k k') as [| Hneq]; trivial. assert (Heq0 : (k - k') * u == 0). { unfold Rminus. rewrite <- add_morph, minus_morph, Heq. apply add_opp. } -elim Hu. rewrite <- (mul_1 u). rewrite <- (Rinv_l (k - k')). +contradiction Hu. rewrite <- (mul_1 u). rewrite <- (Rinv_l (k - k')). - rewrite <- mul_morph. rewrite Heq0. apply mul_origin. - intro Habs. apply Hneq. now apply Rminus_diag_uniq. Qed. Definition middle `{RealVectorSpace} u v := (1/2) * (u + v). -Instance middle_compat `{RealVectorSpace} : Proper (equiv ==> equiv ==> equiv) middle. +Global Instance middle_compat `{RealVectorSpace} : Proper (equiv ==> equiv ==> equiv) middle. Proof using . intros u1 u2 Hu v1 v2 Hv. unfold middle. now rewrite Hu, Hv. Qed. Lemma mul_integral `{RealVectorSpace} : forall k u, k * u == 0 -> k = 0%R \/ u == 0. @@ -174,7 +174,7 @@ abstract (intros x y Hxy; apply add_compat; try reflexivity; []; apply mul_compat; try apply Hxy; []; apply add_compat, opp_compat; reflexivity). Defined. -Instance straight_path_compat {T} `{RealVectorSpace T} : +Global Instance straight_path_compat {T} `{RealVectorSpace T} : Proper (equiv ==> equiv ==> equiv) straight_path. Proof using . intros pt1 pt2 Heq pt1' pt2' Heq' x. simpl. @@ -191,7 +191,7 @@ refine (Build_path _ _ (fun x => @mul _ _ _ RVS x pt) _). abstract (intros x y Hxy; apply mul_compat; reflexivity || apply Hxy). Defined. -Instance local_straight_path_compat {T} `{RealVectorSpace T} : +Global Instance local_straight_path_compat {T} `{RealVectorSpace T} : Proper (equiv ==> equiv) local_straight_path. Proof using . intros pt1 pt2 Heq x. simpl. now apply mul_compat. Qed. @@ -265,7 +265,7 @@ Section Barycenter. 0 < snd (barycenter_aux E init). Proof using . intros E init Hinit Hnil HE. - destruct E as [| e E]; try (now elim Hnil); []. + destruct E as [| e E]; try (now contradiction Hnil); []. simpl. apply Rlt_le_trans with (snd e + snd init)%R. + inv HE. lra. + change (snd e + snd init)%R with (snd ((snd e * fst e + fst init)%VS, snd e + snd init))%R. diff --git a/Spaces/Ring.v b/Spaces/Ring.v index 839f8843878b19ab2439a589a2296f4b5070f75c..b9368aa0dac8d4f1b963c4cc691c3767a990cbd2 100644 --- a/Spaces/Ring.v +++ b/Spaces/Ring.v @@ -8,428 +8,577 @@ *) (**************************************************************************) +Require Import Utf8 Lia Psatz SetoidDec Rbase. +From Pactole Require Import Util.Coqlib Util.Bijection Spaces.Graph + Spaces.Isomorphism Models.NoByzantine Util.Fin. -Require Import Lia Psatz SetoidDec Rbase. -Require Import Pactole.Util.Coqlib. -Require Import Pactole.Core.Identifiers. -Require Export Pactole.Spaces.Graph. -Require Import Pactole.Spaces.Isomorphism. +(** ** A ring **) +Inductive direction := Forward | Backward | SelfLoop. -Typeclasses eauto := (bfs). -Remove Hints eq_setoid : typeclass_instances. +Global Instance direction_Setoid : Setoid direction := eq_setoid direction. -(** ** A ring **) +Lemma direction_eq_dec_subproof : ∀ d1 d2 : direction, {d1 = d2} + {d1 <> d2}. +Proof using . decide equality. Defined. + +Global Instance direction_EqDec : EqDec direction_Setoid + := direction_eq_dec_subproof. + +(* Returns the nat to give as parameter to + addm to translate in the direction dir *) +Definition dir2nat {n : nat} {ltc_2_n : 2 <c n} (d : direction) : nat := + match d with + | SelfLoop => 0 + | Forward => 1 + | Backward => Nat.pred n + end. -(** What we need to define a ring. *) -Class RingSpec := { - ring_size : nat; - ring_size_spec : 1 < ring_size }. -Coercion ring_size_spec : RingSpec >-> lt. +(* ''inverse'' of dir2nat *) +Definition nat2Odir {n : nat} {ltc_2_n : 2 <c n} (m : nat) : option direction := + match m with + | 0 => Some SelfLoop + | 1 => Some Forward + | _ => if m =? Nat.pred n then Some Backward else None + end. + +Global Instance ltc_1_n {n : nat} {ltc_2_n : 2 <c n} : 1 <c n + := (lt_s_u 1 ltac:(auto)). + +Definition ring_edge {n : nat} {ltc_2_n : 2 <c n} := (fin n * direction)%type. Section Ring. -Context {RR : RingSpec}. -Notation ring_node := (finite_node ring_size). +Context {n : nat} {ltc_2_n : 2 <c n}. -Inductive direction := Forward | Backward | SelfLoop. -Definition ring_edge := (ring_node * direction)%type. - -Global Instance direction_Setoid : Setoid direction. -simple refine {| - equiv := fun d1 d2 => if (Nat.eq_dec ring_size 2) - then match d1, d2 with - | SelfLoop, SelfLoop => True - | SelfLoop, _ | _, SelfLoop => False - | _, _ => True - end - else d1 = d2 |}; trivial; []. -Proof. split. -+ intro. repeat destruct_match; reflexivity. -+ intros ? ? ?. repeat destruct_match; auto; now symmetry. -+ intros ? ? ?. repeat destruct_match; auto; congruence. -Defined. +Global Instance ring_edge_Setoid : Setoid ring_edge := eq_setoid ring_edge. + +Lemma ring_edge_dec : ∀ e1 e2 : ring_edge, {e1 = e2} + {e1 ≠e2}. +Proof using . apply pair_dec. apply fin_EqDec. apply direction_EqDec. Qed. + +Global Instance ring_edge_EqDec : EqDec ring_edge_Setoid := ring_edge_dec. + +Definition dir2nat_compat : Proper (equiv ==> equiv) dir2nat := _. + +Lemma dir20 : dir2nat SelfLoop = 0. +Proof using . reflexivity. Qed. + +Lemma dir21 : dir2nat Forward = 1. +Proof using . reflexivity. Qed. + +Lemma dir2pred_n : dir2nat Backward = Nat.pred n. +Proof using . reflexivity. Qed. + +Lemma dir2nat_lt : ∀ d : direction, dir2nat d < n. +Proof using . + intros. destruct d. all: cbn. 2: apply lt_pred_u. + all: apply lt_s_u. all: auto. +Qed. + +Lemma dir2natI : Util.Preliminary.injective equiv equiv dir2nat. +Proof using . + intros d1 d2 H'. destruct d1, d2. all: inversion H' as [H]. + all: try reflexivity. all: exfalso. 1,4: symmetry in H. + all: eapply neq_pred_u_s. 2,4,6,8: exact H. all: auto. +Qed. -Global Instance ring_edge_Setoid : Setoid ring_edge := prod_Setoid _ _. - -Lemma direction_eq_dec : forall d d': direction, {d = d'} + {d <> d'}. -Proof. decide equality. Defined. - -Global Instance direction_EqDec : EqDec direction_Setoid. - refine (fun d1 d2 => if (Nat.eq_dec ring_size 2) - then match d1, d2 with - | SelfLoop, SelfLoop => left _ - | SelfLoop, _ | _, SelfLoop => right _ - | _, _ => left _ - end - else _); simpl; try rewrite e; simpl; try tauto. -Proof. -abstract (repeat destruct_match; tauto || apply direction_eq_dec). +Definition nat2Odir_compat : Proper (equiv ==> equiv) nat2Odir := _. + +Lemma nat2Odir_Some : ∀ (d : direction) (m : nat), + nat2Odir m = Some d <-> dir2nat d = m. +Proof using . + unfold dir2nat, nat2Odir. intros d m. split; intros H. + all: destruct d, m as [|p]. all: try inversion H. all: try reflexivity. + 4:{ exfalso. eapply neq_pred_u_s. 2: exact H. auto. } + 4: rewrite Nat.eqb_refl. all: destruct p. all: try inversion H. + all: try reflexivity. 4:{ exfalso. eapply neq_pred_u_s. 2: exact H. + auto. } all: destruct (S (S p) =? Nat.pred n) eqn:Hd. + all: try inversion H. symmetry. apply Nat.eqb_eq, Hd. +Qed. + +Lemma nat2Odir_None : ∀ m : nat, + nat2Odir m = None <-> (∀ d : direction, dir2nat d ≠m). +Proof using . + intros. split; intros H. + - intros d Habs. apply nat2Odir_Some in Habs. rewrite H in Habs. + inversion Habs. + - destruct (nat2Odir m) eqn:Hd. 2: reflexivity. exfalso. + apply (H d), nat2Odir_Some, Hd. +Qed. + +Lemma nat2Odir_pick : ∀ m : nat, + pick_spec (λ d : direction, dir2nat d = m) (nat2Odir m). +Proof using . + intros. apply pick_Some_None. intros d. + all: rewrite opt_Setoid_eq by reflexivity. + apply nat2Odir_Some. apply nat2Odir_None. +Qed. + +Definition nat2dir (m : nat) (H : ∃ d, dir2nat d = m) : direction. +Proof using . + destruct (nat2Odir_pick m) as [d Hd | Hd]. exact d. + abstract (exfalso; destruct H as [d H]; apply (Hd d), H). Defined. -Global Instance ring_edge_EqDec : EqDec ring_edge_Setoid := prod_EqDec _ _. +Lemma dir2natK : ∀ (d1 : direction) (H : ∃ d2, dir2nat d2 = dir2nat d1), + nat2dir (dir2nat d1) H = d1. +Proof using . + intros. unfold nat2dir. destruct (nat2Odir_pick (dir2nat d1)) as [m Hd | Hd]. + eapply dir2natI, Hd. contradiction (Hd d1). reflexivity. +Qed. + +Lemma nat2dirK : + ∀ (m : nat) (H : ∃ d, dir2nat d = m), dir2nat (nat2dir m H) = m. +Proof using . + intros. unfold nat2dir. destruct (nat2Odir_pick m) as [g Hd | Hd]. + exact Hd. exfalso. destruct H as [d H]. apply (Hd d). exact H. +Qed. + +(* Returns the fin n to give as parameter to + addf to translate in the direction dir *) +Definition dir2nod (d : direction) : fin n := Fin (dir2nat_lt d). -(* the following lemmas are used to easily prove that - (Z.to_nat (l mod Z.of_nat n)) = (l mod Z.of_nat n) *) -Lemma to_Z_sup_0 : forall l : Z, (0 <= l mod Z.of_nat ring_size)%Z. -Proof using . intros. apply Zdiv.Z_mod_lt. destruct RR. simpl. lia. Qed. +(* ''inverse'' of dir2nat *) +Definition nod2Odir (v : fin n) : option direction := nat2Odir v. -Lemma to_Z_inf_n (x : Z): Z.to_nat (x mod Z.of_nat ring_size)%Z < ring_size. +Definition dir2nod_compat : Proper (equiv ==> equiv) dir2nod := _. + +Definition nod2Odir_compat : Proper (equiv ==> equiv) nod2Odir := _. + +Lemma dir2nod2nat : ∀ d : direction, dir2nod d = dir2nat d :> nat. +Proof using . reflexivity. Qed. + +Lemma dir2nodE : ∀ d : direction, dir2nod d = mod2fin (dir2nat d). Proof using . -intros. rewrite <- Nat2Z.id, <- Z2Nat.inj_lt; -try apply Zdiv.Z_mod_lt; destruct RR; simpl; lia. + intros. apply fin2natI. rewrite dir2nod2nat. + symmetry. apply mod2fin_small, dir2nat_lt. Qed. -Definition to_Z (v : ring_node) : Z := Z.of_nat (proj1_sig v). -Definition of_Z (x : Z) : ring_node := - exist _ (Z.to_nat (x mod Z.of_nat ring_size)) (to_Z_inf_n x). +Lemma nod2OdirE : ∀ v : fin n, nod2Odir v = nat2Odir v. +Proof using . reflexivity. Qed. -Global Instance to_Z_compat : Proper (equiv ==> Z.eq) to_Z. -Proof using . repeat intro. hnf in *. now f_equal. Qed. +Lemma dir2fin0 : dir2nod SelfLoop = fin0. +Proof using . apply fin2natI. rewrite dir2nod2nat. apply dir20. Qed. -Global Instance of_Z_compat : Proper (Z.eq ==> equiv) of_Z. -Proof using . intros ? ? Heq. now rewrite Heq. Qed. +Lemma dir2fin1 : dir2nod Forward = fin1. +Proof using . + apply fin2natI. symmetry. rewrite dir2nod2nat, dir21. apply fin12nat. +Qed. -Lemma to_Z_injective : Preliminary.injective equiv Logic.eq to_Z. +Lemma dir2max : dir2nod Backward = fin_max. +Proof using . apply fin2natI. rewrite dir2nod2nat. apply dir2pred_n. Qed. + +Lemma dir2nodI : Util.Preliminary.injective equiv equiv dir2nod. Proof using . -intros [x Hx] [y Hy] Heq. -unfold to_Z in Heq. hnf in Heq |- *. simpl in Heq. -apply Nat2Z.inj in Heq. subst. f_equal. apply le_unique. + intros d1 d2 H. apply dir2natI. rewrite <-2 dir2nod2nat, H. reflexivity. Qed. -Lemma to_Z_small : forall v, (0 <= to_Z v < Z.of_nat ring_size)%Z. -Proof using . intro. unfold to_Z. split; try lia; []. apply Nat2Z.inj_lt. apply proj2_sig. Qed. +Lemma nod2Odir_Some : ∀ (d : direction) (v : fin n), + nod2Odir v = Some d <-> dir2nod d = v. +Proof using . + intros. rewrite nod2OdirE. split. all: intros H. apply fin2natI, + nat2Odir_Some, H. apply nat2Odir_Some. rewrite <- dir2nod2nat, H. reflexivity. +Qed. -Lemma Z2Z : forall l, (to_Z (of_Z l) = l mod Z.of_nat ring_size)%Z. +Lemma nod2Odir_None : ∀ v : fin n, + nod2Odir v = None <-> (∀ d : direction, dir2nod d ≠v). Proof using . -intros. unfold to_Z, of_Z. simpl. -rewrite Z2Nat.id; trivial; []. -apply Z.mod_pos_bound. destruct RR. simpl. lia. + intros. rewrite nod2OdirE. split. + - intros H d Habs. eapply (proj1 (nat2Odir_None _)). apply H. + rewrite <- Habs. symmetry. apply dir2nod2nat. + - intros H. apply nat2Odir_None. intros d Habs. apply (H d), fin2natI, Habs. Qed. -Lemma V2V : forall v, of_Z (to_Z v) == v. +Lemma nod2Odir_pick : ∀ v : fin n, + pick_spec (λ d : direction, dir2nod d = v) (nod2Odir v). Proof using . -intros [k Hk]. hnf. unfold to_Z, of_Z. apply eq_proj1. simpl. -rewrite <- Zdiv.mod_Zmod, Nat2Z.id, Nat.mod_small; lia. + intros. apply pick_Some_None. intros d. + all: rewrite opt_Setoid_eq by reflexivity. + apply nod2Odir_Some. apply nod2Odir_None. +Qed. + +Definition nod2dir (v : fin n) (H : ∃ d, dir2nod d = v) : direction. +Proof using . + apply (nat2dir v). abstract (destruct H as [d H]; exists d; + rewrite <- H; symmetry; apply dir2nod2nat). +Defined. + +Lemma dir2nodK : ∀ (d1 : direction) (H : ∃ d2, dir2nod d2 = dir2nod d1), + nod2dir (dir2nod d1) H = d1. +Proof using . intros. erewrite <- (dir2natK d1). reflexivity. Qed. + +Lemma nod2dirK : + ∀ (v : fin n) (H : ∃ d, dir2nod d = v), dir2nod (nod2dir v H) = v. +Proof using . + intros. apply fin2natI. erewrite <- (nat2dirK v). reflexivity. Qed. (** From a node, if we move in one direction, get get to another node. *) -Definition move_along (v : ring_node) (dir : direction) := - match dir with - | SelfLoop => v - | Forward => of_Z (to_Z v + 1) - | Backward => of_Z (to_Z v - 1) +Definition move_along (v : fin n) (d : direction) + := addf v (dir2nod d). + +Definition move_along_compat : + Proper (equiv ==> equiv ==> equiv) move_along := _. + +Lemma move_alongE : ∀ (v : fin n) (d : direction), + move_along v d = addf v (dir2nod d). +Proof using . intros. reflexivity. Qed. + +Definition move_along2nat : ∀ (v : fin n) (d : direction), + fin2nat (move_along v d) = (v + dir2nat d) mod n. +Proof using . intros. apply addf2nat. Qed. + +Lemma move_alongI_ : ∀ d : direction, + Util.Preliminary.injective equiv equiv (λ v, move_along v d). +Proof using . intros. apply addIf. Qed. + +Lemma move_along_I : ∀ v : fin n, + Util.Preliminary.injective equiv equiv (move_along v). +Proof using . intros v d1 d2 H. eapply dir2nodI, addfI, H. Qed. + +Lemma move_along_0 : ∀ v : fin n, move_along v SelfLoop = v. +Proof using . apply addf0. Qed. + +Lemma move_along0_ : ∀ d : direction, move_along fin0 d = dir2nod d. +Proof using . intros. rewrite move_alongE. apply add0f. Qed. + +Lemma move_alongAC : ∀ (v : fin n) (d1 d2 : direction), + move_along (move_along v d2) d1 = move_along (move_along v d1) d2. +Proof using . intros. rewrite 4 move_alongE. apply addfAC. Qed. + +Lemma addm_move_along : ∀ (v : fin n) (m : nat) (d : direction), + addm (move_along v d) m = move_along (addm v m) d. +Proof using . intros. rewrite 2 move_alongE, 2 addm_addf. apply addfAC. Qed. + +Lemma subm_move_along : ∀ (v : fin n) (m : nat) (d : direction), + subm (move_along v d) m = move_along (subm v m) d. +Proof using . intros. rewrite 2 move_alongE, 2 submE. apply addfAC. Qed. + +Lemma addf_move_along : ∀ (v1 v2 : fin n) (d : direction), + addf (move_along v1 d) v2 = move_along (addf v1 v2) d. +Proof using . intros. rewrite 2 move_alongE. apply addfAC. Qed. + +Lemma subf_move_along : ∀ (v1 v2 : fin n) (d : direction), + subf v2 (move_along v1 d) = subf (subf v2 v1) (dir2nod d). +Proof using . intros. rewrite move_alongE. apply subf_addf. Qed. + +Lemma subf_move_along' : ∀ (v1 v2 : fin n) (d : direction), + subf (move_along v1 d) v2 = move_along (subf v1 v2) d. +Proof using . intros. rewrite 2 move_alongE, 2 subfE. apply addfAC. Qed. + +(* returns the ''opposite'' direction *) +Definition swap_direction (d : direction) := + match d with + | Forward => Backward + | Backward => Forward + | SelfLoop => SelfLoop end. -Global Instance move_along_compat : Proper (equiv ==> equiv ==> equiv) move_along. +Definition swap_direction_compat : + Proper (equiv ==> equiv) swap_direction := _. + +Lemma swap_directionI : + Util.Preliminary.injective equiv equiv swap_direction. +Proof using . + unfold swap_direction. intros d1 d2 H. destruct d1, d2. + all: try inversion H. all: reflexivity. +Qed. + +Lemma swap_directionK : ∀ d : direction, + swap_direction (swap_direction d) == d. +Proof using . + intros d. unfold swap_direction. destruct d. all: reflexivity. +Qed. + +Definition swbd : bijection direction + := cancel_bijection swap_direction _ swap_directionI + swap_directionK swap_directionK. + +Lemma swbdE : swbd = swap_direction :> (direction -> direction). +Proof using . reflexivity. Qed. + +Lemma swbdVE : swbdâ»Â¹ = swap_direction :> (direction -> direction). +Proof using . reflexivity. Qed. + +Lemma swbdV : swbdâ»Â¹ == swbd. +Proof using . intros dir. rewrite swbdE, swbdVE. reflexivity. Qed. + +Lemma swbdK : swbd ∘ swbd == Bijection.id. +Proof using . intros dir. rewrite compE, swbdE. apply swap_directionK. Qed. + +Lemma dir2nod_swap_direction : ∀ d : direction, + dir2nod (swap_direction d) = oppf (dir2nod d). +Proof using . + intros. destruct d. all: cbn. all: symmetry. 1,2: rewrite dir2max, dir2fin1. + apply oppf1. apply oppf_max. rewrite dir2fin0. apply oppf0. +Qed. + +Lemma move_along_swap_direction : ∀ (v : fin n) (d : direction), + move_along v (swap_direction d) = subf v (dir2nod d). +Proof using . + intros. rewrite move_alongE, dir2nod_swap_direction. symmetry. apply subfE. +Qed. + +Lemma move_alongK' : ∀ (v : fin n) (d : direction), + subf (move_along v d) (dir2nod d) = v. +Proof using . intros. rewrite move_alongE. apply addfVKV. Qed. + +Lemma submVKV' : ∀ (v : fin n) (d : direction), + move_along (subf v (dir2nod d)) d = v. +Proof using . intros. rewrite move_alongE. apply subfVKV. Qed. + +Lemma move_alongK : ∀ (v : fin n) (d : direction), + move_along (move_along v d) (swap_direction d) = v. +Proof using . + intros. rewrite move_along_swap_direction. apply move_alongK'. +Qed. + +Lemma move_along_dir2nod_subf : ∀ (v1 v2 : fin n) (d : direction), + dir2nod d = subf v2 v1 <-> move_along v1 d = v2. +Proof using . + intros. rewrite move_alongE. split; intros H. rewrite H. apply subfVK. + rewrite <- H. symmetry. apply addfKV. +Qed. + +Lemma symf_move_along : ∀ (c v : fin n) (d : direction), + symf c (move_along v d) = move_along (symf c v) (swap_direction d). +Proof using . + intros. rewrite move_along_swap_direction, move_alongE. apply symf_addf. +Qed. + +Lemma move_along_symf : ∀ (v1 v2 : fin n) (d : direction), + move_along (symf v1 v2) d = symf v1 (move_along v2 (swap_direction d)). Proof using . -intros v1 v2 Hv e1 e2. simpl in *. unfold move_along. subst. -destruct (Nat.eq_dec ring_size 2) as [Hsize | Hsize]; -repeat destruct_match; try tauto || discriminate; [|]; -intros _; unfold of_Z; apply eq_proj1; simpl; rewrite Hsize; simpl; -assert (Hv2 := proj2_sig v2); -destruct v2 as [[| [| [| v]]] ?]; simpl in *; lia. + intros. rewrite symf_move_along, swap_directionK. reflexivity. Qed. -Definition Ring (thd : ring_edge -> R) - (thd_pos : forall e, (0 < thd e < 1)%R) - (thd_compat : Proper (equiv ==> Logic.eq) thd) - : FiniteGraph ring_size (ring_edge). -Proof. -refine ({| +Lemma nod2Odir_subf_Some : + ∀ (v1 v2 : fin n) (d : direction), nod2Odir (subf v1 v2) = Some d + <-> nod2Odir (subf v2 v1) = Some (swap_direction d). +Proof using . + intros. rewrite 2 nod2Odir_Some, dir2nod_swap_direction, <- (oppf_subf v2). + symmetry. apply injective_eq_iff, oppfI. +Qed. + +Lemma nod2Odir_subf_None : ∀ v1 v2 : fin n, + nod2Odir (subf v2 v1) = None <-> ∀ d : direction, move_along v1 d ≠v2. +Proof using . + intros. rewrite nod2Odir_None. split. all: intros H d Heq. all: apply (H d). + all: apply move_along_dir2nod_subf, Heq. +Qed. + +Definition find_edge_subterm (v1 v2 : fin n) : option ring_edge + := match (nod2Odir (subf v2 v1)) with + | Some d1 => Some (v1, d1) + | None => None + end. + +Lemma find_edge_subproof : + ∀ (e : ring_edge) (v1 v2 : fin n), find_edge_subterm v1 v2 == Some e + <-> v1 == fst e /\ v2 == move_along (fst e) (snd e). +Proof using . + intros. rewrite opt_Setoid_eq by reflexivity. unfold find_edge_subterm. + cbn. rewrite (eq_sym_iff v2), <- move_along_dir2nod_subf. + destruct (nod2Odir_pick (subf v2 v1)) as [d Hd|Hd]. + - rewrite (injective_eq_iff Some_eq_inj), <- (pair_eqE (v1, d)). cbn[fst snd]. + split. all: intros [H1 H2]. all: subst. all: split; [reflexivity |]. + apply Hd. apply dir2nodI. rewrite H2. apply Hd. + - split. intros H. inversion H. intros [? H]. subst. exfalso. eapply Hd, H. +Qed. + +Global Instance Ring : Graph (fin n) ring_edge := {| + V_EqDec := @fin_EqDec n; + E_EqDec := ring_edge_EqDec; src := fst; - tgt := fun e => move_along (fst e) (snd e); - threshold := thd; - find_edge := fun v1 v2 => if v1 =?= of_Z (to_Z v2 + 1) then Some (v1, Backward) else - if of_Z (to_Z v1 + 1) =?= v2 then Some (v1, Forward) else - if v1 =?= v2 then Some (v1, SelfLoop) else None; - V_EqDec := @finite_node_EqDec ring_size; - E_EqDec := ring_edge_EqDec |}). -* exact thd_pos. -* (* src_compat *) - intros e1 e2 He. apply He. -* (* tgt_compat *) - intros e1 e2 [Ht Hd]. apply eq_proj1. - repeat rewrite Ht. clear Ht. revert Hd. - simpl. repeat destruct_match; simpl; intro; - try tauto || discriminate; [| |]; - try (destruct (fst e2) as [k ?]; unfold to_Z; simpl; - match goal with H : ring_size = 2 |- _ => try rewrite H in *; clear H end; - destruct k as [| [| k]]; simpl; lia); []. - rewrite move_along_compat; trivial; try rewrite Hd; reflexivity. -(* * (* find_edge_compat *) - intros v1 v2 Hv12 v3 v4 Hv34. hnf in *. subst. - repeat destruct_match; hnf in *; simpl in *; try easy || congruence; [| |]; - (split; trivial; []); now destruct_match. *) -Open Scope Z_scope. -* (* find_edge_None *) - assert (Hsize := ring_size_spec). - unfold move_along. - intros a b; split; unfold find_edge; - destruct (a =?= of_Z (to_Z b + 1)) as [Heq_a | Hneq_a]. - + tauto. - + destruct (of_Z (to_Z a + 1) =?= b) as [Heq_b | Hneq_b], (a =?= b); try tauto; []. - intros _ e [Hsrc Htgt]. - destruct (snd e); rewrite Hsrc in Htgt. - - hnf in *. subst b. intuition. - - elim Hneq_a. rewrite <- Htgt. rewrite Z2Z. apply eq_proj1. - unfold of_Z. simpl. rewrite Zdiv.Zplus_mod_idemp_l. - ring_simplify (to_Z a - 1 + 1). - unfold to_Z. rewrite <- Zdiv.mod_Zmod, Nat2Z.id, Nat.mod_small; try lia; []. - apply proj2_sig. - - contradiction. - + intros Hedge. elim (Hedge (a, Backward)). - split; simpl; try reflexivity; []. - rewrite Heq_a, Z2Z. apply eq_proj1. - unfold Z.sub, of_Z. simpl. rewrite Zdiv.Zplus_mod_idemp_l. - ring_simplify (to_Z b + 1 + -1). - unfold to_Z. rewrite <- Zdiv.mod_Zmod, Nat2Z.id, Nat.mod_small; lia || apply proj2_sig. - + intro Hedge. destruct (of_Z (to_Z a + 1) =?= b) as [Heq_b | Hneq_b]. - - elim (Hedge (a, Forward)). - split; simpl; try reflexivity; []. now rewrite <- Heq_b. - - destruct (a =?= b) as [Heq |]; try reflexivity; []. - elim (Hedge (a, SelfLoop)). - split; simpl; try reflexivity; []. apply Heq. -* (* find_edge_Some *) - assert (Hsize_pos := ring_size_spec). - unfold move_along. - clear dependent thd. - intros v1 v2 e. - simpl in *. repeat (destruct_match; simpl in *); subst; - match goal with - | H : ring_size = 2%nat |- _ => rename H into Hsize - | H : ring_size <> 2%nat |- _ => rename H into Hsize - end; - destruct e as [v dir]; simpl in *; split; intros []; subst; - try split; try tauto || discriminate. - + apply eq_proj1. unfold of_Z, to_Z. simpl. - destruct v2 as [[| [| ?]] ?]; simpl; rewrite Hsize; simpl; lia. - + apply eq_proj1. unfold of_Z, to_Z. simpl. - destruct v2 as [[| [| ?]] ?]; simpl; rewrite Hsize; simpl; lia. - + unfold of_Z, to_Z in *. destruct v2 as [v2 ?]; simpl in *; rewrite Hsize in *; simpl in *. - match goal with H : exist _ _ _ = _ |- _ => rename H into Heq end. - apply (f_equal (@proj1_sig _ _)) in Heq. simpl in Heq. rewrite Hsize in *. - destruct v2 as [| [| ?]]; simpl in *; lia. - + apply eq_proj1. destruct v as [v Hv]. unfold of_Z, to_Z. simpl. rewrite Hsize in *. - destruct v as [| [| ?]]; simpl; lia. - + match goal with H : _ = v |- _ => rename H into Heq end. - apply (f_equal (@proj1_sig _ _)) in Heq. - unfold to_Z, of_Z in Heq. destruct v as [v Hv]. simpl in Heq. - rewrite Hsize in *. destruct v as [| [| ?]]; simpl in *; lia || lia. - + match goal with H : v = _ |- _ => rename H into Heq end. - apply (f_equal (@proj1_sig _ _)) in Heq. - unfold to_Z, of_Z in Heq. destruct v as [v Hv]. simpl in Heq. - rewrite Hsize in *. destruct v as [| [| ?]]; simpl in *; lia || lia. - + match goal with H : of_Z (to_Z v + 1) = of_Z (to_Z v - 1) -> False |- _ => apply H end. - apply eq_proj1. unfold of_Z, to_Z. simpl. - rewrite <- (Zdiv.Z_mod_plus_full (_ - 1) 1 (Z.of_nat ring_size)). - do 2 f_equal. destruct v as [v ?]; simpl; rewrite Hsize; simpl. ring. - + exfalso. match goal with H : _ = _ |- _ => rename H into Heq end. - apply (f_equal to_Z) in Heq. symmetry in Heq. revert Heq. - rewrite 2 Z2Z. destruct v2 as [v2 ?]. unfold of_Z, to_Z. simpl. - rewrite Zdiv.Zplus_mod_idemp_l. - rewrite <- (Z.mod_small (Z.of_nat v2) (Z.of_nat ring_size)) at 2; try lia; []. - replace (Z.of_nat v2 + 1 + 1) with (Z.of_nat v2 + 2) by ring. - apply Zadd_small_mod_non_conf. lia. - + rewrite Z2Z. apply eq_proj1. destruct v2 as [v2 Hv2]. unfold to_Z, of_Z; simpl. - rewrite Zdiv.Zminus_mod_idemp_l. - ring_simplify (Z.of_nat v2 + 1 - 1). rewrite Z.mod_small, Nat2Z.id; lia. - + exfalso. match goal with H : _ = _ |- _ => rename H into Heq end. - apply (f_equal to_Z)in Heq. rewrite Z2Z in Heq. destruct v2 as [v2 ?]. - unfold to_Z, of_Z in Heq. simpl in Heq. symmetry in Heq. revert Heq. - rewrite <- (Z.mod_small (Z.of_nat v2) (Z.of_nat ring_size)) at 2; try lia; []. - apply Zadd_small_mod_non_conf. lia. - + exfalso. match goal with H : _ = _ |- _ => rename H into Heq end. - apply (f_equal to_Z) in Heq. revert Heq. - rewrite 2 Z2Z. replace (to_Z v + 1) with (to_Z v - 1 + 2) by ring. - apply Zadd_small_mod_non_conf. lia. - + exfalso. match goal with H : _ = _ |- _ => rename H into Heq end. - apply (f_equal to_Z) in Heq. revert Heq. - rewrite Z2Z. rewrite <- (Z.mod_small _ _ (to_Z_small v)) at 2. - apply Zadd_small_mod_non_conf. lia. - + exfalso. match goal with H : _ = _ |- _ => rename H into Heq end. - apply (f_equal to_Z) in Heq. revert Heq. - rewrite Z2Z. rewrite <- (Z.mod_small _ _ (to_Z_small v)) at 1. - replace (to_Z v) with (to_Z v - 1 + 1) at 1 by ring. - apply Zadd_small_mod_non_conf. lia. - + match goal with H : _ = of_Z (_ + 1) -> False |- _ => apply H end. - apply to_Z_injective. rewrite 2 Z2Z, Zdiv.Zplus_mod_idemp_l. - ring_simplify (to_Z v - 1 + 1). symmetry. apply Z.mod_small, to_Z_small. -Defined. + tgt := λ e, move_along (fst e) (snd e); + find_edge := find_edge_subterm; + find_edge_Some := find_edge_subproof |}. -(** If we do not care about threshold values, we just take 1/2 everywhere. *) -Definition nothresholdRing : FiniteGraph ring_size (ring_edge) := - Ring (fun _ => 1/2)%R - ltac:(abstract (intro; lra)) - (fun _ _ _ => eq_refl). -End Ring. +Global Instance Ring_isomorphism_Setoid : + Setoid (isomorphism Ring) := isomorphism_Setoid. + +Global Instance Ring_isomorphism_Inverse : Inverse (isomorphism Ring) + := @IsoInverse _ _ Ring. +Global Instance Ring_isomorphism_Composition + : Composition (isomorphism Ring) := @IsoComposition _ _ Ring. (** ** Ring operations **) -Section RingTranslation. -Context {RR : RingSpec}. -Local Instance localRing : FiniteGraph ring_size ring_edge := nothresholdRing. -Notation ring_node := (finite_node ring_size). - -(** *** Translation along a ring **) - -(* TODO: generalize the definition of translation to thresholds. *) -Lemma bij_trans_V_proof : forall c x y, - of_Z (to_Z x - c) == y <-> of_Z (to_Z y + c) == x. -Proof using . -intros c [x ?] [y ?]. unfold of_Z, to_Z. -split; intro Heq; apply (f_equal (@proj1_sig _ _)) in Heq; -simpl in *; subst; apply eq_proj1; simpl; -rewrite Z2Nat.id; auto using to_Z_sup_0. -- rewrite Z.add_mod_idemp_l; try lia; []. - ring_simplify (Z.of_nat x - c + c)%Z. - rewrite Z.mod_small, Nat2Z.id; lia. -- rewrite Zdiv.Zminus_mod_idemp_l. - ring_simplify (Z.of_nat y + c - c)%Z. - rewrite Z.mod_small, Nat2Z.id; lia. -Qed. - -Definition bij_trans_V (c : Z) : @Bijection.bijection ring_node (@V_Setoid _ _ localRing) := {| - Bijection.section := fun x => of_Z (to_Z x - c); - Bijection.retraction := fun x => of_Z (to_Z x + c); - Bijection.Inversion := bij_trans_V_proof c |}. - -Definition bij_trans_E (c : Z) : @Bijection.bijection ring_edge (@E_Setoid _ _ localRing). -refine {| - Bijection.section := fun x => (of_Z (to_Z (fst x) - c), snd x); - Bijection.retraction := fun x => (of_Z (to_Z (fst x) + c), snd x) |}. -Proof. -+ abstract (intros ? ? Heq; hnf in *; simpl; destruct Heq as [Heq ?]; now rewrite Heq). -+ abstract (intros [x dx] [y dy]; split; intros [Hxy Hd]; - split; try (now apply bij_trans_V_proof); [|]; - destruct (Nat.eq_dec ring_size 2); - solve [simpl in *; destruct dx, dy; tauto | symmetry; apply Hd]). +Lemma srcE : ∀ e : ring_edge, src e = fst e. +Proof using . reflexivity. Qed. + +Lemma tgtE : ∀ e : ring_edge, tgt e = move_along (fst e) (snd e). +Proof using . reflexivity. Qed. + +Lemma find_edgeE : ∀ v1 v2 : fin n, + find_edge v1 v2 = match (nod2Odir (subf v2 v1)) with + | Some d1 => Some (v1, d1) + | None => None + end. +Proof using . reflexivity. Qed. + +Lemma find_edge_move_along : ∀ (v : fin n) (d : direction), + find_edge v (move_along v d) = Some (v, d). +Proof using . + intros. rewrite <- opt_Setoid_eq. setoid_rewrite find_edge_Some. + split. all: reflexivity. +Qed. + +Definition trans (v : fin n) : isomorphism Ring. +Proof using . + refine {| + iso_V := asbf vâ»Â¹; + iso_E := prod_eq_bij (asbf vâ»Â¹) Bijection.id |}. + abstract (intros; split; [reflexivity | apply subf_move_along']). Defined. +Global Opaque trans. -Definition trans (c : Z) : isomorphism localRing. -refine {| - iso_V := bij_trans_V c; - iso_E := bij_trans_E c; - iso_T := @Bijection.id _ R_Setoid |}. -Proof. -* (* iso_morphism *) - intro e. split. - + simpl. reflexivity. - + apply to_Z_injective. - unfold tgt. simpl. unfold move_along. - destruct (snd e) eqn:Hsnd; repeat rewrite Z2Z. - - repeat rewrite ?Zdiv.Zplus_mod_idemp_l, ?Zdiv.Zminus_mod_idemp_l. f_equal. ring. - - rewrite 2 Zdiv.Zminus_mod_idemp_l. f_equal. ring. - - reflexivity. -* (* iso_threshold *) - intro. simpl. reflexivity. -* (* iso_incr *) - intro. simpl. tauto. -* (* iso_bound_T *) - intro. simpl. tauto. -Defined. (* TODO: abstract the proofs *) - -Instance trans_compat : Proper (equiv ==> @equiv _ isomorphism_Setoid) trans. -Proof using . -intros c1 c2 Hc. simpl in *. subst. -repeat split; try reflexivity; []. -repeat destruct_match; tauto. -Qed. - -Lemma trans_origin : @equiv _ isomorphism_Setoid (trans 0) Isomorphism.id. -Proof using . -split; [| split; [| reflexivity]]. -+ intro x. apply eq_proj1. cbn -[of_Z]. now rewrite Z.sub_0_r, V2V. -+ intros [x d]. cbn -[equiv]. now rewrite Z.sub_0_r, V2V. -Qed. - -Lemma trans_same : forall k, Bijection.section (trans (to_Z k)) k == of_Z 0. -Proof using . intro k. simpl. f_equal. ring. Qed. - -(** *** Symmetry of a ring w.r.t. a point [c] **) - -Definition swap_direction dir := - match dir with - | Forward => Backward - | Backward => Forward - | SelfLoop => SelfLoop - end. +Definition trans_compat : Proper (equiv ==> equiv) trans := _. + +Lemma transvE : ∀ v : fin n, trans v = asbf vâ»Â¹ :> bijection (fin n). +Proof using . reflexivity. Qed. + +Lemma transvVE : ∀ v : fin n, trans vâ»Â¹ == asbf v :> bijection (fin n). +Proof using . intros. rewrite <- (Bijection.inv_inv (asbf v)). reflexivity. Qed. + +Lemma transeE : + ∀ v : fin n, iso_E (trans v) = prod_eq_bij (asbf vâ»Â¹) Bijection.id. +Proof using . reflexivity. Qed. + +Lemma transeVE : + ∀ v : fin n, iso_E (trans vâ»Â¹) == prod_eq_bij (asbf v) Bijection.id. +Proof using . + intros. rewrite <- (Bijection.inv_inv (asbf v)), + <- Bijection.id_inv, <- prod_eq_bij_inv. reflexivity. +Qed. + +Lemma trans0 : trans fin0 == id. +Proof using . + split. rewrite transvE. 2: rewrite transeE. all: rewrite asbf0, + Bijection.id_inv. reflexivity. apply prod_eq_bij_id. +Qed. -Lemma bij_sym_V_proof : forall c x y, of_Z (c - to_Z x) == y <-> of_Z (c - to_Z y) == x. -Proof using . -intros c x y. simpl. split; intro; subst; rewrite Z2Z; apply eq_proj1; -unfold of_Z; simpl proj1_sig; rewrite Zdiv.Zminus_mod_idemp_r; -match goal with x : ring_node |- _ => - ring_simplify (c - (c - to_Z x))%Z; destruct x as [m Hm] end; -unfold to_Z; simpl; rewrite Z.mod_small, Nat2Z.id; lia. -Qed. - -Definition bij_sym_V (c : Z) : @Bijection.bijection ring_node (@V_Setoid _ _ localRing) := {| - Bijection.section := fun x => of_Z (c - to_Z x); - Bijection.retraction := fun x => of_Z (c - to_Z x); - Bijection.Inversion := bij_sym_V_proof c |}. - -Definition bij_sym_E (c : Z) : @Bijection.bijection ring_edge (@E_Setoid _ _ localRing). -refine {| - Bijection.section := fun x => (of_Z (c - to_Z (fst x)), swap_direction (snd x)); - Bijection.retraction := fun x => (of_Z (c - to_Z (fst x)), swap_direction (snd x)) |}. -Proof. -+ abstract (intros x y [Heq ?]; hnf in *; simpl; rewrite Heq; - destruct (Nat.eq_dec ring_size 2), (snd x), (snd y); simpl in *; tauto || discriminate). -+ abstract (intros [x dx] [y dy]; split; intros [Hxy Hd]; - split; try (now apply bij_sym_V_proof); [|]; simpl in *; - destruct (Nat.eq_dec ring_size 2), dx, dy; simpl in *; tauto || discriminate). +Lemma transI : ∀ v1 v2 : fin n, trans v1 == trans v2 → v1 = v2. +Proof using . intros * H. apply subfI with (f1:=v1), H. Qed. + +Lemma transVI : ∀ v1 v2 : fin n, (trans v1)â»Â¹ == (trans v2)â»Â¹ → v1 = v2. +Proof using . intros * H. apply addfI with (f1:=v1), H. Qed. + +Lemma transAC : + ∀ v1 v2 : fin n, trans v2 ∘ (trans v1) == trans v1 ∘ (trans v2). +Proof using . + intros. split. rewrite 2 (compvE (trans _)), 2 transvE. apply asbfVAC. + rewrite 2 (compeE (trans _)), 2 transeE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfVAC v2). reflexivity. +Qed. + +Lemma transVAC : ∀ v1 v2 : fin n, + (trans v2)â»Â¹ ∘ (trans v1)â»Â¹ == (trans v1)â»Â¹ ∘ (trans v2)â»Â¹. +Proof using . + intros. split. rewrite 2 (compvE (trans _â»Â¹)), 2 transvVE. apply asbfAC. + rewrite 2 (compeE (trans _â»Â¹)), 2 transeVE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfAC v2). reflexivity. +Qed. + +Lemma transCV : ∀ v1 v2 : fin n, + trans v1 ∘ (trans v2)â»Â¹ == (trans v2)â»Â¹ ∘ (trans v1). +Proof using . + intros. split. rewrite (compvE (trans _)), (compvE (trans _â»Â¹)), transvE, + transvVE. symmetry. apply asbfCV. rewrite (compeE (trans _)), + (compeE (trans _â»Â¹)), transeE, transeVE. setoid_rewrite prod_eq_bij_comp. + rewrite (asbfCV v2). reflexivity. +Qed. + +Lemma transA : + ∀ v1 v2 : fin n, trans (trans v1 v2) == (trans v1â»Â¹) ∘ (trans v2). +Proof using . + intros. split. rewrite (compvE (trans _â»Â¹)), 3 transvE, transvVE. + apply asbfVAV. rewrite (compeE (trans _â»Â¹)), transvE, 2 transeE, transeVE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfVAV, Bijection.id_comp_l. + reflexivity. +Qed. + +Lemma transAV : + ∀ v1 v2 : fin n, trans ((trans v1)â»Â¹ v2) == (trans v1) ∘ (trans v2). +Proof using . + intros. split. rewrite (compvE (trans _)), 3 transvE, transvVE. + apply asbfVA. rewrite (compeE (trans _)), transvVE, 3 transeE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfVA, Bijection.id_comp_l. + reflexivity. +Qed. + +Lemma transVA : + ∀ v1 v2 : fin n, trans (trans v1 v2)â»Â¹ == (trans v1) ∘ (trans v2â»Â¹). +Proof using . + intros. split. rewrite (compvE (trans _)), transvE, 2 transvVE. + apply asbfAV. rewrite (compeE (trans _)), transvE, transeE, 2 transeVE. + setoid_rewrite prod_eq_bij_comp. rewrite asbfAV, Bijection.id_comp_l. + reflexivity. +Qed. + +Lemma transVAV : + ∀ v1 v2 : fin n, trans ((trans v1)â»Â¹ v2)â»Â¹ == (trans v1â»Â¹) ∘ (trans v2â»Â¹). +Proof using . + intros. split. rewrite (compvE (trans _â»Â¹)), 3 transvVE. + apply asbfA. rewrite (compeE (trans _â»Â¹)), transvVE, 3 transeVE. + setoid_rewrite prod_eq_bij_comp. + rewrite asbfA, Bijection.id_comp_l. reflexivity. +Qed. + +Lemma move_along_trans : ∀ (v1 v2 : fin n) (d : direction), + move_along (trans v1 v2) d == trans v1 (move_along v2 d). +Proof using . + intros. rewrite transvE, (asbfVE v1). symmetry. apply subf_move_along'. +Qed. + +Lemma trans_inj : Preliminary.injective equiv equiv trans. +Proof using . + intros v1 v2 [H _]. eapply addfI. rewrite (addfC _ v2). + apply inverse_compat in H. rewrite (transvVE v1), (transvVE v2) in H. + specialize (H v2). rewrite (asbfE v1), (asbfE v2) in H. apply H. +Qed. + +Lemma transs : ∀ v : fin n, trans v v = fin0. +Proof using . intros. rewrite transvE. apply asbfVf. Qed. + +Definition sym (v : fin n) : isomorphism Ring. +Proof using . + refine {| + iso_V := sybf v; + iso_E := prod_eq_bij (sybf v) swbd |}. + (* iso_morphism *) + abstract (intros; split; [reflexivity | apply symf_move_along]). Defined. +Global Opaque sym. + +Definition sym_compat : Proper (equiv ==> equiv) sym := _. + +Lemma symvE : ∀ v : fin n, sym v = sybf v :> bijection (fin n). +Proof using . reflexivity. Qed. -Definition sym (c : Z) : isomorphism localRing. -refine {| - iso_V := bij_sym_V c; - iso_E := bij_sym_E c; - iso_T := @Bijection.id _ R_Setoid |}. -Proof. -* (* iso_morphism *) - intro e. split. - + simpl. reflexivity. - + apply eq_proj1. - unfold tgt. simpl. unfold move_along. - destruct (snd e) eqn:Hsnd; simpl; repeat rewrite Z2Z. - - rewrite Zdiv.Zminus_mod_idemp_r, Zdiv.Zminus_mod_idemp_l. do 2 f_equal. ring. - - rewrite Zdiv.Zplus_mod_idemp_l, Zdiv.Zminus_mod_idemp_r. do 2 f_equal. ring. - - reflexivity. -* (* iso_threshold *) - intro. simpl. reflexivity. -* (* iso_incr *) - intro. simpl. tauto. -* (* iso_bound_T *) - intro. simpl. tauto. -Defined. (* TODO: abstract the proofs *) - -Instance sym_compat : Proper (equiv ==> @equiv _ isomorphism_Setoid) sym. -Proof using . -intros c1 c2 Hc. simpl in *. subst. -repeat split; try reflexivity; []. -repeat destruct_match; tauto. -Qed. - -Lemma sym_involutive : forall c, - @equiv _ isomorphism_Setoid (@compose _ _ IsoComposition (sym c) (sym c)) Isomorphism.id. -Proof using . -intro c. split; [| split; [| simpl; reflexivity]]. -+ intro x. apply eq_proj1. cbn -[of_Z]. rewrite Z2Z. unfold of_Z. simpl. - rewrite Zdiv.Zminus_mod_idemp_r. ring_simplify (c - (c - to_Z x))%Z. - unfold to_Z. rewrite Z.mod_small, Nat2Z.id; trivial; []. assert (Hx := proj2_sig x). - simpl in Hx. - lia. -+ intro x. split. - - (* TODO: same proof as above, factor it? *) - apply eq_proj1. cbn -[of_Z]. rewrite Z2Z. unfold of_Z. simpl. - rewrite Zdiv.Zminus_mod_idemp_r. ring_simplify (c - (c - to_Z (fst x)))%Z. - unfold to_Z. rewrite Z.mod_small, Nat2Z.id; trivial; []. - assert (Hx := proj2_sig (fst x)). - simpl in Hx. - lia. - - simpl. destruct (snd x); simpl; now destruct_match. -Qed. - -End RingTranslation. +Lemma symeE : ∀ v : fin n, iso_E (sym v) = prod_eq_bij (sybf v) swbd. +Proof using . reflexivity. Qed. + +Lemma symvVE : ∀ v : fin n, sym vâ»Â¹ == sybf v :> bijection (fin n). +Proof using . intros v1 v2. rewrite <- (sybfV v1). reflexivity. Qed. + +Lemma symeVE : ∀ v : fin n, iso_E (sym vâ»Â¹) == prod_eq_bij (sybf v) swbd. +Proof using . + intros. rewrite <- sybfV, <- swbdV, <- prod_eq_bij_inv. reflexivity. +Qed. + +Lemma symK : ∀ v : fin n, (sym v) ∘ (sym v) == id. +Proof using . + intros. split. rewrite (compvE (sym v)), symvE. apply sybfK. + rewrite (compeE (sym v)), symeE. setoid_rewrite prod_eq_bij_comp. + rewrite (sybfK v), swbdK, prod_eq_bij_id. reflexivity. +Qed. + +Lemma move_along_sym : ∀ (v1 v2 : fin n) (d : direction), + move_along (sym v1 v2) d = sym v1 (move_along v2 (swap_direction d)). +Proof using . intros. rewrite symvE, (sybfE v1). apply move_along_symf. Qed. + +Lemma symm : ∀ v : fin n, sym v v = v. +Proof using . intros. rewrite symvE. apply sybff. Qed. + +End Ring. diff --git a/Spaces/ThresholdIsomorphism.v b/Spaces/ThresholdIsomorphism.v new file mode 100644 index 0000000000000000000000000000000000000000..8011d799cc64af5a841c05f4700b8ebf5dc92a6f --- /dev/null +++ b/Spaces/ThresholdIsomorphism.v @@ -0,0 +1,100 @@ +Require Import Utf8 SetoidDec Rbase Rbasic_fun Psatz. +From Pactole Require Import Util.Coqlib Util.Bijection Util.Ratio. +From Pactole Require Import Spaces.Graph Spaces.Isomorphism. +Set Implicit Arguments. + +Section ThresholdIsomorphism. + +Context {V E : Type}. +Context {G : ThresholdGraph V E}. + +Record threshold_isomorphism := { + iso_VE :> isomorphism G; + iso_T : bijection strict_ratio; + iso_threshold : ∀ e : E, iso_T (threshold e) == threshold (iso_E iso_VE e) :> strict_ratio; + iso_incr : ∀ a b : strict_ratio, (a < b)%R -> (iso_T a < iso_T b)%R }. + +Global Instance threshold_isomorphism_Setoid : Setoid threshold_isomorphism. +Proof using . + simple refine {| + equiv := λ iso1 iso2, iso1.(iso_VE) == iso2.(iso_VE) + /\ iso1.(iso_T) == iso2.(iso_T) |}; autoclass. split. + + intro f. now repeat split. + + intros f g Hfg; destruct Hfg as [HVE HT]. split; now symmetry. + + intros f g h Hfg Hgh. destruct Hfg as [? ?], Hgh as [? ?]. + split; etransitivity; eauto. +Defined. + +Instance iso_T_compat : Proper (equiv ==> equiv) iso_T. +Proof using . intros ? ? Heq ?. now apply Heq. Qed. + +Definition id : threshold_isomorphism. +Proof using . + refine {| iso_VE := id; + iso_T := Bijection.id |}. + + now intros. + + now intros. +Defined. + +Definition comp (f g : threshold_isomorphism) : threshold_isomorphism. +Proof using . + refine {| + iso_VE := compose f.(iso_VE) g.(iso_VE); + iso_T := compose f.(iso_T) g.(iso_T) |}. + + intro. simpl. now rewrite 2 iso_threshold. + + intros. simpl. now do 2 apply iso_incr. +Defined. + +Global Instance TIsoComposition : Composition threshold_isomorphism. +Proof using . + refine {| compose := comp |}. + intros f1 f2 Hf g1 g2 Hg. split. all: apply compose_compat. + 3,4: f_equiv. 1,3: apply Hf. all: apply Hg. +Defined. + +Lemma compose_assoc : ∀ f g h, f ∘ (g ∘ h) == (f ∘ g) ∘ h. +Proof using . intros f g h; repeat split; simpl; reflexivity. Qed. + +Definition inv (tiso : threshold_isomorphism) : threshold_isomorphism. +Proof using . + refine {| iso_VE := inverse tiso.(iso_VE); + iso_T := inverse tiso.(iso_T) |}. + + intros e. cbn-[equiv]. rewrite <- Inversion, iso_threshold, + section_retraction. reflexivity. + + intros a b Hab. apply Rnot_le_lt. intros [Hd | Hd]. + all: apply (Rlt_not_le b a Hab). left. erewrite <- (section_retraction _ b), + <- (section_retraction _ a). apply iso_incr, Hd. right. + apply proj_ratio_compat, proj_strict_ratio_compat, + (injective (iso_T tiso â»Â¹)), proj_strict_ratio_inj, proj_ratio_inj, Hd. +Defined. + +Global Instance TIsoInverse : Inverse threshold_isomorphism. +Proof using . + refine {| inverse := inv |}. + intros f g [? ?]. split. all: apply inverse_compat. all: assumption. +Defined. + +Lemma id_inv : idâ»Â¹ == id. +Proof using . split. apply id_inv. setoid_rewrite Bijection.id_inv. reflexivity. Qed. + +Lemma id_comp_l : ∀ tiso : threshold_isomorphism, id ∘ tiso == tiso. +Proof using. intros. split. apply id_comp_l. setoid_rewrite Bijection.id_comp_l. reflexivity. Qed. + +Lemma id_comp_r : ∀ tiso : threshold_isomorphism, tiso ∘ id == tiso. +Proof using. intros. split. apply id_comp_r. setoid_rewrite Bijection.id_comp_r. reflexivity. Qed. + +Lemma inv_inv : ∀ tiso : threshold_isomorphism, tisoâ»Â¹â»Â¹ == tiso. +Proof using . intros. split. apply inv_inv. setoid_rewrite Bijection.inv_inv. reflexivity. Qed. + +Lemma compose_inverse_l : ∀ tiso : threshold_isomorphism, tiso â»Â¹ ∘ tiso == id. +Proof using . intro. split. apply compose_inverse_l. setoid_rewrite Bijection.compose_inverse_l. reflexivity. Qed. + +Lemma compose_inverse_r : ∀ tiso : threshold_isomorphism, tiso ∘ (tiso â»Â¹) == id. +Proof using . intro. split. apply compose_inverse_r. setoid_rewrite Bijection.compose_inverse_r. reflexivity. Qed. + +Lemma inverse_compose : ∀ f g : threshold_isomorphism, (f ∘ g) â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). +Proof using . intros f g. split. apply inverse_compose. setoid_rewrite Bijection.inverse_compose. reflexivity. Qed. + +End ThresholdIsomorphism. + +Arguments threshold_isomorphism {V} {E} G. diff --git a/Util/Bijection.v b/Util/Bijection.v index 00818bc0c72950512ccf3d3113f7d9698c1282d9..a7ed108b9a9e59b441049fe638bcd1a1ea30d343 100644 --- a/Util/Bijection.v +++ b/Util/Bijection.v @@ -22,7 +22,7 @@ Set Implicit Arguments. (** Bijections on a type [T] with an equivalence relation [eqT] *) Section Bijections. -Context (T : Type). +Context {T : Type}. Context {HeqT : Setoid T}. Record bijection := { @@ -30,20 +30,25 @@ Record bijection := { retraction : T → T; section_compat : Proper (equiv ==> equiv) section; Inversion : ∀ x y, section x == y ↔ retraction y == x}. -Global Existing Instance section_compat. +#[export]Existing Instance section_compat. -Global Instance bij_Setoid : Setoid bijection. -simple refine {| equiv := fun f g => forall x, f.(section) x == g x |}; auto; []. +#[export]Instance bij_Setoid : Setoid bijection. +simple refine {| equiv := λ f g, ∀ x, f.(section) x == g x |}; auto; []. Proof. split. + repeat intro. reflexivity. + repeat intro. now symmetry. + repeat intro. etransitivity; eauto. Defined. -Global Instance section_full_compat : Proper (equiv ==> (equiv ==> equiv)) section. +(* Use fun_Setoid directly in the definition of bij_Setoid? *) +#[export]Instance bij_Setoid_eq_compat : Proper ((@equiv bijection bij_Setoid) + ==> (@equiv (T -> T) (fun_Setoid T HeqT))) section. +Proof using . intros ?? H. apply H. Qed. + +#[export]Instance section_full_compat : Proper (equiv ==> (equiv ==> equiv)) section. Proof using . intros f g Hfg x y Hxy. rewrite Hxy. now apply Hfg. Qed. -Global Instance retraction_compat : Proper (equiv ==> (equiv ==> equiv)) retraction. +#[export]Instance retraction_compat : Proper (equiv ==> (equiv ==> equiv)) retraction. Proof using . intros f g Hfg x y Hxy. now rewrite <- f.(Inversion), Hxy, Hfg, g.(Inversion). Qed. (** The identity bijection *) @@ -62,14 +67,18 @@ Proof. + abstract (intros x y; now rewrite f.(Inversion), <- g.(Inversion)). Defined. -Global Instance BijectionComposition : Composition bijection. +#[export]Instance BijectionComposition : Composition bijection. refine {| compose := comp |}. Proof. intros f1 f2 Hf g1 g2 Hg x. cbn -[equiv]. rewrite (Hf (g1 x)). f_equiv. apply Hg. Defined. + +Lemma compE : ∀ f g : bijection, f ∘ g = (λ t, f (g t)) :> (T -> T). +Proof using . reflexivity. Qed. + (* -Global Instance compose_compat : Proper (equiv ==> equiv ==> equiv) compose. +#[export]Instance compose_compat : Proper (equiv ==> equiv ==> equiv) compose. Proof. intros f1 f2 Hf g1 g2 Hg x. cbn -[equiv]. rewrite (Hf (g1 x)). f_equiv. apply Hg. @@ -84,11 +93,24 @@ refine {| section := bij.(retraction); retraction := bij.(section) |}. Proof. abstract (intros; rewrite bij.(Inversion); reflexivity). Defined. -Global Instance BijectionInverse : Inverse bijection. +#[export]Instance BijectionInverse : Inverse bijection. refine {| inverse := inv |}. Proof. repeat intro. simpl. now f_equiv. Defined. + +Lemma id_inv : idâ»Â¹ == id. +Proof using . intros t. reflexivity. Qed. + +Lemma id_comp_l : ∀ b : bijection, id ∘ b == b. +Proof using. intros. cbn. reflexivity. Qed. + +Lemma id_comp_r : ∀ b : bijection, b ∘ id == b. +Proof using. intros. cbn. reflexivity. Qed. + +Lemma inv_inv : ∀ b : bijection, bâ»Â¹â»Â¹ == b. +Proof using . intros. cbn. reflexivity. Qed. + (* -Global Instance inverse_compat : Proper (equiv ==> equiv) inverse. +#[export]Instance inverse_compat : Proper (equiv ==> equiv) inverse. Proof. repeat intro. simpl. now f_equiv. Qed. *) Lemma retraction_section : forall (bij : bijection) x, bij.(retraction) (bij.(section) x) == x. @@ -106,6 +128,18 @@ Proof using . repeat intro. simpl. now rewrite section_retraction. Qed. Lemma inverse_compose : forall f g : bijection, (f ∘ g)â»Â¹ == (g â»Â¹) ∘ (f â»Â¹). Proof using . repeat intro. reflexivity. Qed. +Definition cancel_bijection (f g : T -> T) + (fP : Proper (equiv ==> equiv) f) (f_inj : injective equiv equiv f) + (fK : ∀ t, g (f t) == t) (gK : ∀ t, f (g t) == t) : bijection. +Proof using . + refine {| + section := f; + retraction := g; + section_compat := fP |}. + abstract (intros t1 t2; split; intros H; [rewrite <- (gK t2) in H; + apply f_inj in H; rewrite H; reflexivity | rewrite <- H; apply gK]). +Defined. + (** Bijections are in particular injective. *) Lemma injective : forall bij : bijection, injective equiv equiv bij. Proof using . intros bij x y Heq. now rewrite <- (retraction_section bij x), Heq, retraction_section. Qed. @@ -115,3 +149,109 @@ End Bijections. Arguments bijection T {_}. Arguments section {_} {_} !_ x. Arguments retraction {_} {_} !_ x. + +Section prod_bij. + +Context {A B : Type} (SA : Setoid A) (SB : Setoid B). + +Definition prod_bij (BA : bijection A) (BB : bijection B) : bijection (A * B). +Proof using . + refine {| + section := λ p, (BA (fst p), BB (snd p)); + retraction := λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p)) |}. + abstract (intros ?? H; rewrite H; reflexivity). + abstract (intros; split; intros H; rewrite <- H; split; + apply Inversion; reflexivity). +Defined. + +Lemma prod_bijE : ∀ (BA : bijection A) (BB : bijection B), + prod_bij BA BB = (λ p : A * B, (BA (fst p), BB (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Lemma prod_bijVE : ∀ (BA : bijection A) (BB : bijection B), + prod_bij BA BBâ»Â¹ = (λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +#[export]Instance prod_bij_compat : Proper (equiv ==> equiv ==> equiv) prod_bij. +Proof using . + intros * BA1 BA2 HA BB1 BB2 HB p. rewrite 2 prod_bijE. split. + rewrite HA. reflexivity. rewrite HB. reflexivity. +Qed. + +Lemma prod_bij_id : prod_bij (@id A SA) (@id B SB) == id. +Proof using . intros * p. rewrite prod_bijE. split. all: reflexivity. Qed. + +Lemma prod_bij_comp : ∀ (BA1 BA2 : bijection A) (BB1 BB2 : bijection B), + prod_bij BA1 BB1 ∘ (prod_bij BA2 BB2) == prod_bij (BA1 ∘ BA2) (BB1 ∘ BB2). +Proof using . intros * p. rewrite compE, 3 prod_bijE, 2 compE. reflexivity. Qed. + +Lemma prod_bij_inv : ∀ (BA : bijection A) (BB : bijection B), + (prod_bij BA BB)â»Â¹ == prod_bij (BAâ»Â¹) (BBâ»Â¹). +Proof using . intros * p. rewrite prod_bijVE, prod_bijE. reflexivity. Qed. + +Definition prod_eq_bij (BA : @bijection A (eq_setoid A)) + (BB : @bijection B (eq_setoid B)) : @bijection (A * B) (eq_setoid (A * B)). +Proof using . + refine {| + section := λ p, (BA (fst p), BB (snd p)); + retraction := λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p)) |}. + abstract (intros; split; cbn; intros H; rewrite <- H, <- pair_eqE; + cbn[fst snd]; [rewrite 2 retraction_section | rewrite 2 section_retraction]; + split; reflexivity). +Defined. + +Lemma prod_eq_bijE : ∀ (BA : bijection A) (BB : bijection B), prod_eq_bij BA BB + = (λ p : A * B, (BA (fst p), BB (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +Lemma prod_eq_bijVE : ∀ (BA : bijection A) (BB : bijection B), + prod_eq_bij BA BBâ»Â¹ = (λ p, (BAâ»Â¹ (fst p), BBâ»Â¹ (snd p))) :> (A * B -> A * B). +Proof using . reflexivity. Qed. + +#[export]Instance prod_eq_bij_compat : + Proper (equiv ==> equiv ==> equiv) prod_eq_bij. +Proof using . + cbn. intros * BA1 BA2 HA BB1 BB2 HB p. rewrite 2 prod_eq_bijE, <- pair_eqE. + split. rewrite HA. reflexivity. rewrite HB. reflexivity. +Qed. + +Lemma prod_eq_bij_id : prod_eq_bij id id == id. +Proof using . + intros * p. rewrite prod_eq_bijE. apply pair_eqE. split. all: reflexivity. +Qed. + +Lemma prod_eq_bij_comp : ∀ (BA1 BA2 : bijection A) (BB1 BB2 : bijection B), + prod_eq_bij BA1 BB1 ∘ (prod_eq_bij BA2 BB2) + == prod_eq_bij (BA1 ∘ BA2) (BB1 ∘ BB2). +Proof using . + intros * p. rewrite compE, 3 prod_eq_bijE, 2 compE. reflexivity. +Qed. + +Lemma prod_eq_bij_inv : ∀ (BA : bijection A) (BB : bijection B), + (prod_eq_bij BA BB)â»Â¹ == prod_eq_bij (BAâ»Â¹) (BBâ»Â¹). +Proof using . intros * p. rewrite prod_eq_bijVE, prod_eq_bijE. reflexivity. Qed. + +End prod_bij. + +Section equiv_bij. + +Context {T : Type} [S1 S2 : Setoid T]. +Context (H : ∀ t1 t2 : T, @equiv T S1 t1 t2 <-> @equiv T S2 t1 t2). + +Definition equiv_bij : @bijection T S1 -> @bijection T S2. +Proof using H. + intros [s r c p]. refine {| section := s; retraction := r |}. + abstract (intros ?? Heq; apply H, c, H, Heq). + abstract (intros; rewrite <-2 H; apply p). +Defined. + +Lemma equiv_bijE : ∀ b : @bijection T S1, equiv_bij b = b :> (T -> T). +Proof using . intros []. reflexivity. Qed. + +Lemma equiv_bijVE : ∀ b : @bijection T S1, equiv_bij bâ»Â¹ = bâ»Â¹ :> (T -> T). +Proof using . intros []. reflexivity. Qed. + +Lemma equiv_bij_id : equiv_bij (@id T S1) == (@id T S2). +Proof using . intros t. rewrite equiv_bijE. reflexivity. Qed. + +End equiv_bij. diff --git a/Util/Enum.v b/Util/Enum.v new file mode 100644 index 0000000000000000000000000000000000000000..e7e63556348417e60a9c1baa318c320ef8573396 --- /dev/null +++ b/Util/Enum.v @@ -0,0 +1,203 @@ +Require Import Utf8 SetoidDec SetoidList Arith_base Lia. +From Pactole Require Import Util.Coqlib Util.Fin. + +Set Implicit Arguments. + +(* TODO: should we add a fold operator? *) +(* FIXME: change the equalities to use equiv and the Setoid class *) + +Program Fixpoint build_enum N k (Hle : k <= N) acc : list (fin N) := + match k with + | 0 => acc + | S m => @build_enum N m _ (@Fin _ m _ :: acc) + end. +Next Obligation. + lia. +Qed. + +(** A list containing all elements of [fin N]. *) +Definition enum N : list (fin N) := build_enum (Nat.le_refl N) nil. + +(** Specification of [enum]. *) +Lemma In_build_enum : ∀ N k (Hle : k <= N) l x, + In x (build_enum Hle l) <-> In x l \/ fin2nat x < k. +Proof using . + intros N k. induction k; intros Hle l x; simpl. + + intuition auto with *. + + rewrite IHk. simpl. split; intro Hin. + - destruct Hin as [[Hin | Hin] | Hin]; intuition; []. + subst. simpl. right. lia. + - destruct Hin as [Hin | Hin]; intuition; []. + assert (Hcase : fin2nat x < k \/ fin2nat x = k) by lia. + destruct Hcase as [Hcase | Hcase]; intuition; []. + subst. do 2 left. destruct x; f_equal; simpl in *. apply le_unique. +Qed. + +Lemma In_enum : ∀ N x, In x (enum N) <-> fin2nat x < N. +Proof using . intros. unfold enum. rewrite In_build_enum. simpl. intuition. Qed. + +(** Length of [enum]. *) +Lemma build_enum_length : ∀ N k (Hle : k <= N) l, + length (build_enum Hle l) = k + length l. +Proof using . + intros N k. induction k; intros Hle l; simpl. + + reflexivity. + + rewrite IHk. simpl. lia. +Qed. + +Lemma enum_length : ∀ N, length (enum N) = N. +Proof using . intro. unfold enum. now rewrite build_enum_length. Qed. + +(** [enum] does not contain duplicates. *) +Lemma build_enum_NoDup : ∀ N k (Hle : k <= N) l, + (∀ x, In x l -> k <= fin2nat x) -> NoDup l -> NoDup (build_enum Hle l). +Proof using . + intros N k. induction k; intros Hle l Hin Hl; simpl; auto; []. + apply IHk. + + intros x [Hx | Hx]. + - now subst. + - apply Hin in Hx. lia. + + constructor; trivial; []. + intro Habs. apply Hin in Habs. simpl in Habs. lia. +Qed. + +Lemma enum_NoDup : ∀ N, NoDup (enum N). +Proof using . + intro. unfold enum. + apply build_enum_NoDup; simpl; intuition; constructor. +Qed. + +(** [enum] is sorted in increasing order. *) +Notation Flt := (fun x y => lt (fin2nat x) (fin2nat y)). + +Lemma build_enum_Sorted : + ∀ N k (Hle : k <= N) l, (∀ x, In x l -> k <= fin2nat x) + -> Sorted Flt l -> Sorted Flt (build_enum Hle l). +Proof using . + intros N k. induction k; intros Hle l Hin Hl; simpl; auto; []. + apply IHk. + + intros x [Hx | Hx]. + - now subst. + - apply Hin in Hx. lia. + + constructor; trivial; []. + destruct l; constructor; []. simpl. apply Hin. now left. +Qed. + +Lemma enum_Sorted : ∀ N, Sorted Flt (enum N). +Proof using . + intro. unfold enum. + apply build_enum_Sorted; simpl; intuition. +Qed. + +(** Extensional equality of functions is decidable over finite domains. *) +Lemma build_enum_app_nil : ∀ N k (Hle : k <= N) l, + build_enum Hle l = build_enum Hle nil ++ l. +Proof using . + intros N k. induction k; intros Hle l; simpl. + + reflexivity. + + now rewrite (IHk _ (_ :: nil)), IHk, <- app_assoc. +Qed. + +Theorem build_enum_eq : ∀ {A} eqA N (f g : fin N -> A) k (Hle : k <= N) l, + eqlistA eqA (List.map f (build_enum Hle l)) (List.map g (build_enum Hle l)) -> + ∀ x, fin2nat x < k -> eqA (f x) (g x). +Proof using . + intros A eqA N f g k. induction k; intros Hle l Heq x Hx; simpl. + * destruct x; simpl in *; lia. + * assert (Hlt : k <= N) by lia. + assert (Hcase : fin2nat x < k \/ fin2nat x = k) by lia. + destruct Hcase as [Hcase | Hcase]. + + apply IHk with (x := x) in Heq; auto. + + subst k. cbn in Heq. rewrite build_enum_app_nil, map_app, map_app in Heq. + destruct (eqlistA_app_split _ _ _ _ Heq) as [_ Heq']. + - now do 2 rewrite map_length, build_enum_length. + - simpl in Heq'. inv Heq'. + assert (Heqx : x = Fin Hle). + { clear. destruct x; simpl. f_equal. apply le_unique. } + now rewrite Heqx. +Qed. + +Corollary enum_eq : ∀ {A} eqA N (f g : fin N -> A), + eqlistA eqA (List.map f (enum N)) (List.map g (enum N)) + -> ∀ x, eqA (f x) (g x). +Proof using . + unfold enum. intros A eqA N f g Heq x. + apply build_enum_eq with (x := x) in Heq; auto; []. apply fin_lt. +Qed. + +(** Cutting [enum] after some number of elements. *) +Lemma firstn_build_enum_le : ∀ N k (Hle : k <= N) l k' (Hk : k' <= N), + k' <= k -> firstn k' (build_enum Hle l) = @build_enum N k' Hk nil. +Proof using . + intros N k. induction k; intros Hk l k' Hk' Hle. + * assert (k' = 0) by lia. now subst. + * rewrite build_enum_app_nil, firstn_app, build_enum_length. + replace (k' - (S k + length (@nil (fin N)))) with 0 by lia. + rewrite app_nil_r. + destruct (Nat.eq_dec k' (S k)) as [Heq | Heq]. + + subst k'. rewrite firstn_all2. + - f_equal. apply le_unique. + - rewrite build_enum_length. simpl. lia. + + simpl build_enum. erewrite IHk. + - f_equal. + - lia. +Qed. + +Lemma firstn_build_enum_lt : ∀ N k (Hle : k <= N) l k', k <= k' -> + firstn k' (build_enum Hle l) = build_enum Hle (firstn (k' - k) l). +Proof using . + intros N k. induction k; intros Hle l k' Hk. + + now rewrite Nat.sub_0_r. + + rewrite build_enum_app_nil, firstn_app, build_enum_length, Nat.add_0_r. + rewrite firstn_all2, <- build_enum_app_nil; trivial; []. + rewrite build_enum_length. simpl. lia. +Qed. + +Lemma firstn_enum_le : ∀ N k (Hle : k <= N), + firstn k (enum N) = build_enum Hle nil. +Proof using . intros. unfold enum. now apply firstn_build_enum_le. Qed. + +Lemma firstn_enum_lt : ∀ N k, N <= k -> firstn k (enum N) = enum N. +Proof using . + intros. unfold enum. now rewrite firstn_build_enum_lt, firstn_nil. +Qed. + +Lemma firstn_enum_spec : ∀ N k x, In x (firstn k (enum N)) <-> fin2nat x < k. +Proof using . + intros N k x. destruct (le_lt_dec k N) as [Hle | Hlt]. + + rewrite (firstn_enum_le Hle), In_build_enum. simpl. intuition. + + rewrite (firstn_enum_lt (Nat.lt_le_incl _ _ Hlt)). + split; intro Hin. + - transitivity N; trivial; []. apply fin_lt. + - apply In_enum, fin_lt. +Qed. + +(** Removing some number of elements from the head of [enum]. *) +Lemma skipn_build_enum_lt : ∀ N k (Hle : k <= N) l k', k <= k' -> + skipn k' (build_enum Hle l) = skipn (k' - k) l. +Proof using . + intros N k Hle l k' Hk'. apply app_inv_head with + (firstn k' (build_enum Hle l)). rewrite firstn_skipn, firstn_build_enum_lt; + trivial; []. rewrite (build_enum_app_nil Hle (firstn _ _)). + now rewrite build_enum_app_nil, <- app_assoc, firstn_skipn. +Qed. + +Lemma skipn_enum_lt : ∀ N k, N <= k -> skipn k (enum N) = nil. +Proof using . + intros. unfold enum. now rewrite skipn_build_enum_lt, skipn_nil. +Qed. + +Lemma skipn_enum_spec : + ∀ N k x, In x (skipn k (enum N)) <-> k <= fin2nat x < N. +Proof using . + intros N k x. split; intro Hin. + + assert (Hin' : ~In x (firstn k (enum N))). + { intro Habs. rewrite <- InA_Leibniz in *. revert x Habs Hin. + apply NoDupA_app_iff; autoclass; []. + rewrite firstn_skipn. rewrite NoDupA_Leibniz. apply enum_NoDup. } + rewrite firstn_enum_spec in Hin'. split; auto with zarith; []. + apply fin_lt. + + assert (Hin' : In x (enum N)) by apply In_enum, fin_lt. + rewrite <- (firstn_skipn k), in_app_iff, firstn_enum_spec in Hin'. + intuition lia. +Qed. diff --git a/Util/FMaps/FMapFacts.v b/Util/FMaps/FMapFacts.v index 98db9062583feeb4917802253d3c0e0f9a5d8487..f2241b0e1e9252198e19242ed80c76c804b9b434 100644 --- a/Util/FMaps/FMapFacts.v +++ b/Util/FMaps/FMapFacts.v @@ -2,7 +2,7 @@ Require Import Bool. Require Import Structures.DecidableType. Require Import SetoidDec. Require Import Pactole.Util.FMaps.FMapInterface. - +Require Import Pactole.Util.SetoidDefs. Set Implicit Arguments. Unset Strict Implicit. @@ -16,6 +16,8 @@ Hint Extern 1 (Equivalence _) => constructor; congruence : core. Hint Extern 1 (equiv ?x ?x) => reflexivity : core. Hint Extern 2 (equiv ?y ?x) => now symmetry : core. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Notation Leibniz := (@eq _) (only parsing). @@ -28,7 +30,7 @@ Section WeakFacts. Lemma eq_bool_alt : forall b b', b=b' <-> (b=true <-> b'=true). Proof using . - destruct b; destruct b'; intuition. + destruct b; destruct b'; clear; intuition. Qed. Lemma eq_option_alt : forall (elt:Type)(o o':option elt), @@ -66,7 +68,7 @@ Section WeakFacts. Lemma MapsTo_iff : forall m x y e, x == y -> (MapsTo x e m <-> MapsTo y e m). Proof using HF. - split; apply MapsTo_1; auto using symmetry. + split; apply MapsTo_1 ; auto (*using symmetry stopped working/being useful in 8.20*). Qed. Lemma mem_in_iff : forall m x, In x m <-> mem x m = true. @@ -95,7 +97,7 @@ Section WeakFacts. Proof using HF. split; intros. rewrite eq_option_alt. intro e. rewrite <- find_mapsto_iff. - split; intro H'; try discriminate. elim H; exists e; auto. + split; intro H'; try discriminate. contradiction H; exists e; auto. intros (e,He); rewrite find_mapsto_iff,H in He; discriminate. Qed. @@ -374,7 +376,7 @@ Section WeakFacts. Hint Resolve add_neq_o : map. Lemma add_o : forall m x y e, - find y (add x e m) = if x == y then Some e else find y m. + find y (add x e m) = if x =?= y then Some e else find y m. Proof using HF. intros; destruct (equiv_dec x y); auto with map. Qed. @@ -395,7 +397,7 @@ Section WeakFacts. mem y (add x e m) = (equiv x y) || mem y m. Proof. intros; do 2 rewrite mem_find_b; rewrite add_o; unfold eqb. - destruct (equiv_dec x y); simpl; auto. + destruct (equiv_dec x y); cbn; auto. Qed. *) Lemma remove_eq_o : forall m x y, @@ -419,7 +421,7 @@ Section WeakFacts. Hint Resolve remove_neq_o : map. Lemma remove_o : forall m x y, - find y (remove x m) = if x == y then None else find y m. + find y (remove x m) = if x =?= y then None else find y m. Proof using HF. intros; destruct (equiv_dec x y); auto with map. Qed. @@ -455,7 +457,7 @@ Section WeakFacts. intros. generalize (find_mapsto_iff (map f m) x) (find_mapsto_iff m x) (fun b => map_mapsto_iff m x b f). - destruct (find x (map f m)); destruct (find x m); simpl; auto; intros. + destruct (find x (map f m)); destruct (find x m); cbn; auto; intros. rewrite <- H; rewrite H1; exists e0; rewrite H0; auto. destruct (H e) as [_ H2]. rewrite H1 in H2. @@ -468,7 +470,7 @@ Section WeakFacts. mem x (map f m) = mem x m. Proof using HF. intros; do 2 rewrite mem_find_b; rewrite map_o. - destruct (find x m); simpl; auto. + destruct (find x m); cbn; auto. Qed. Lemma mapi_b : forall m x (f:key->elt->elt'), @@ -476,7 +478,7 @@ Section WeakFacts. Proof using HF. intros. generalize (mem_in_iff (mapi f m) x) (mem_in_iff m x) (mapi_in_iff m x f). - destruct (mem x (mapi f m)); destruct (mem x m); simpl; auto; intros. + destruct (mem x (mapi f m)); destruct (mem x m); cbn; auto; intros. symmetry; rewrite <- H0; rewrite <- H1; rewrite H; auto. rewrite <- H; rewrite H1; rewrite H0; auto. Qed. @@ -488,7 +490,7 @@ Section WeakFacts. intros. generalize (find_mapsto_iff (mapi f m) x) (find_mapsto_iff m x) (fun b => mapi_mapsto_iff m x b H). - destruct (find x (mapi f m)); destruct (find x m); simpl; auto; intros. + destruct (find x (mapi f m)); destruct (find x m); cbn; auto; intros. rewrite <- H0; rewrite H2; exists e0; rewrite H1; auto. destruct (H0 e) as [_ H3]. rewrite H2 in H3. @@ -526,7 +528,7 @@ Section WeakFacts. findA (eqb x) l = findA (fun y => if eqb_dec x y then true else false) l. Proof. - intros; induction l; simpl. + intros; induction l; cbn. reflexivity. unfold eqb; destruct a; destruct (equiv_dec x k); destruct (eqb_dec x k); auto; contradiction. @@ -555,12 +557,12 @@ Section WeakFacts. rewrite InA_alt in He. destruct He as ((y,e'),(Ha1,Ha2)). compute in Ha1; destruct Ha1; subst e'. - exists (y,e); split; simpl; auto. + exists (y,e); split; cbn; auto. unfold eqb; destruct (equiv_dec x y); intuition. rewrite <- H; rewrite H0. destruct H1 as (H1,_). destruct H1 as ((y,e),(Ha1,Ha2)); [intuition|]. - simpl in Ha2. + cbn in Ha2. unfold eqb in *; destruct (equiv_dec x y); auto; try discriminate. exists e; rewrite InA_alt. exists (y,e); intuition. @@ -714,9 +716,10 @@ Section WeakFacts. Global Instance Empty_m elt : Proper (Equal ==> iff) (@Empty key _ _ _ elt). Proof using HF. - unfold Empty; intros m m' Hm; intuition. - rewrite <-Hm in H0; eauto. - rewrite Hm in H0; eauto. + unfold Empty; intros m m' Hm. + split; repeat intro. + - rewrite <-Hm in H0; eapply H;eauto. + - rewrite Hm in H0; eapply H;eauto. Qed. Global Instance is_empty_m elt : @@ -793,9 +796,9 @@ Section MoreWeakFacts. k1 == k2 -> InA eq_key_elt (k1,e1) l -> InA eq_key (k2,e2) l. Proof using . intros k1 k2 e1 e2 l Hk. rewrite 2 InA_alt. - intros ((k',e') & (Hk',He') & H); simpl in *. + intros ((k',e') & (Hk',He') & H); cbn in *. exists (k',e'); split; auto. - red; red; simpl; eauto. + red; red; cbn; eauto. transitivity k1; auto; symmetry; auto. Qed. @@ -843,7 +846,7 @@ Section MoreWeakFacts. rewrite elements_mapsto_iff. rewrite InA_alt; exists a; auto. } destruct (elements m); auto. - elim (H0 p); simpl; auto. + contradiction (H0 p); cbn; auto. + red; intros. rewrite elements_mapsto_iff in H0. rewrite InA_alt in H0; destruct H0. @@ -866,18 +869,18 @@ Section MoreWeakFacts. NoDupA eq_key l -> (MapsTo k e (of_list l) <-> InA eq_key_elt (k,e) l). Proof using HF. - induction l as [|(k',e') l IH]; simpl; intros k e Hnodup. + induction l as [|(k',e') l IH]; cbn-[equiv]; intros k e Hnodup. + rewrite empty_mapsto_iff, InA_nil; intuition. + inversion_clear Hnodup as [| ? ? Hnotin Hnodup']. specialize (IH k e Hnodup'); clear Hnodup'. rewrite add_mapsto_iff, InA_cons, <- IH. unfold eq_key_elt in *. split; destruct 1 as [H|H]; try (intuition; fail). - - destruct (equiv_dec k k'); [left|right]; split; try (now (idtac + symmetry); auto); [|]. - * now destruct H. - * elim c. now destruct H. + - left. destruct H. split. + * cbn[fst] in H. now symmetry. + * cbn in H0. now symmetry. - destruct (equiv_dec k k'). - * left. elim Hnotin. apply InA_eq_key_elt_eq_key with k e; intuition. + * left. contradiction Hnotin. apply InA_eq_key_elt_eq_key with k e; intuition. * right. now split; try symmetry. Qed. @@ -885,7 +888,7 @@ Section MoreWeakFacts. NoDupA eq_key l -> find k (of_list l) = findA (eqb k) l. Proof. - induction l as [|(k',e') l IH]; simpl; intros k Hnodup. + induction l as [|(k',e') l IH]; cbn; intros k Hnodup. apply empty_o. inversion_clear Hnodup as [| ? ? Hnotin Hnodup']. specialize (IH k Hnodup'); clear Hnodup'. @@ -942,11 +945,11 @@ Section MoreWeakFacts. { intros k e; unfold l. now rewrite <- find_mapsto_iff, InA_rev, elements_mapsto_iff. } clearbody l. clearbody ff. clear Hstep f. revert m Hsame. induction l. (* empty *) - intros m Hsame; simpl. + intros m Hsame; cbn. apply Hempty. intros k e. rewrite find_mapsto_iff, Hsame; intro Habs; inversion Habs. (* step *) - intros m Hsame; destruct a as (k,e); simpl. + intros m Hsame; destruct a as (k,e); cbn. apply Hstep' with (of_list l); auto. - inversion_clear Hdup. contradict H. destruct H as (e',He'). apply InA_eq_key_elt_eq_key with k e'; auto. now rewrite <- of_list_1. @@ -958,7 +961,7 @@ Section MoreWeakFacts. * destruct (find k' m) eqn:Habs; trivial. exfalso. rewrite Hsame in Habs. inversion_clear Habs. - -- elim c. now destruct H1. + -- contradiction c. now destruct H1. -- rewrite <- of_list_1, find_mapsto_iff in H1; trivial. rewrite H1 in Hin. discriminate. - apply IHl. @@ -983,7 +986,7 @@ Section MoreWeakFacts. apply fold_rec; intros. apply Pmorphism with (empty _); auto. intro k. rewrite empty_o. case_eq (find k m0); auto; intros e'; rewrite <- find_mapsto_iff. - intro H'; elim (H k e'); auto. + intro H'; contradiction (H k e'); auto. apply Pmorphism with (add k e m'); try intro; auto. Qed. @@ -1062,7 +1065,7 @@ Section MoreWeakFacts. intros m' Heq k'. rewrite empty_o. case_eq (find k' m'); auto; intros e'; rewrite <- find_mapsto_iff. - intro; elim (Heq k' e'); auto. + intro; contradiction (Heq k' e'); auto. intros k e a m' m'' _ _ Hadd Heq k'. rewrite Hadd, 2 add_o, Heq; auto. Qed. @@ -1093,7 +1096,7 @@ Section MoreWeakFacts. Proof using HF st. intros. apply fold_rec_nodep with (P:=fun a => eqA a i). reflexivity. - intros. elim (H k e); auto. + intros. contradiction (H k e); auto. Qed. (** As noticed by P. Casteran, asking for the general [SetoidList.transpose] @@ -1142,10 +1145,10 @@ Section MoreWeakFacts. (apply NoDupA_rev; [typeclasses eauto | apply elements_3]). apply fold_right_equivlistA_restr with (R:=complement eq_key)(eqA:=eq_key_elt); auto with map; try typeclasses eauto. - intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; simpl in *; apply Comp; auto. + intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; cbn in *; apply Comp; auto. unfold complement, eq_key, eq_key_elt; repeat red. intros ? ? Heq ? ? Heq'. rewrite Heq, Heq'. tauto. - intros (k,e) (k',e'); unfold eq_key; simpl; auto with *. + intros (k,e) (k',e'); unfold eq_key; cbn; auto with *. rewrite <- NoDupA_altdef; auto. intros (k,e). rewrite 2 InA_rev; try apply eq_key_elt_Equiv. @@ -1159,27 +1162,23 @@ Section MoreWeakFacts. eqA (fold f m2 i) (f k e (fold f m1 i)). Proof using Comp HF Tra st. assert (eq_key_elt_refl : forall p, eq_key_elt p p). - red; auto. + { red. auto. } assert (eq_key_elt_sym : forall p p', eq_key_elt p p' -> eq_key_elt p' p). - intros (x1,x2) (y1,y2); unfold eq_key_elt; simpl; intuition. + { intros (x1,x2) (y1,y2); unfold eq_key_elt; cbn; intuition. } assert (eq_key_elt_trans : forall p p' p'', eq_key_elt p p' -> eq_key_elt p' p'' -> eq_key_elt p p''). - intros (x1,x2) (y1,y2) (z1,z2); unfold eq_key_elt; simpl. - intuition; subst; auto; transitivity y1; auto. + { intros (x1,x2) (y1,y2) (z1,z2); unfold eq_key_elt. etransitivity; eauto. } intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right. set (f':=fun y x0 => f (fst y) (snd y) x0) in *. change (f k e (fold_right f' i (rev (elements m1)))) with (f' (k,e) (fold_right f' i (rev (elements m1)))). apply fold_right_add_restr with - (R:=complement eq_key)(eqA:=eq_key_elt)(eqB:=eqA); auto. - typeclasses eauto. + (R:=complement eq_key)(eqA:=eq_key_elt)(eqB:=eqA); eauto with typeclass_instances. intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; unfold f'; - simpl in *. apply Comp; auto. - unfold complement, eq_key, eq_key_elt; repeat red. - intros ? ? Habs Heq. apply Habs. now symmetry. + cbn in *. apply Comp; auto. unfold complement, eq_key, eq_key_elt; repeat red. intros ? ? Heq ? ? Heq'. rewrite Heq, Heq'. tauto. - unfold f'; intros (k1,e1) (k2,e2); unfold eq_key; simpl; auto. + unfold f'; intros (k1,e1) (k2,e2); unfold eq_key; cbn; auto. apply NoDupA_rev; auto using eq_key_elt_Equivalence; apply NoDupA_eq_key_eq_key_elt; apply elements_3. apply NoDupA_rev; auto using eq_key_elt_Equivalence; @@ -1194,7 +1193,7 @@ Section MoreWeakFacts. rewrite InA_cons; do 2 (rewrite InA_rev by apply eq_key_elt_Equivalence); destruct a as (a,b); fold eq_key_elt; do 2 rewrite <- elements_mapsto_iff. - do 2 rewrite find_mapsto_iff; unfold eq_key_elt; simpl. + do 2 rewrite find_mapsto_iff; unfold eq_key_elt; cbn. rewrite H0. rewrite add_o. destruct (equiv_dec k a); intuition try (easy || congruence). @@ -1262,7 +1261,7 @@ Section MoreWeakFacts. generalize (elements_mapsto_iff m). destruct (elements m); try discriminate. exists p; auto. - rewrite H0; destruct p; simpl; auto. + rewrite H0; destruct p; cbn; auto. Qed. Lemma cardinal_inv_2b : @@ -1270,7 +1269,7 @@ Section MoreWeakFacts. Proof using HF. intros. generalize (@cardinal_inv_2 m); destruct @cardinal. - elim H;auto. + contradiction H;auto. eauto. Qed. @@ -1325,16 +1324,16 @@ Section MoreWeakFacts. intro m. pattern m, (fold f' m (empty _)). apply fold_rec. intros m' Hm' k e. rewrite empty_mapsto_iff. intuition. - elim (Hm' k e); auto. + contradiction (Hm' k e); auto. intros k e acc m1 m2 Hke Hn Hadd IH k' e'. change (Equal m2 (add k e m1)) in Hadd; rewrite Hadd. - unfold f'; simpl. - case_eq (f k e); intros Hfke; simpl; + unfold f'; cbn. + case_eq (f k e); intros Hfke; cbn; rewrite !add_mapsto_iff, IH; clear IH; intuition. rewrite <- Hfke; apply Hf; auto. destruct (equiv_dec k k') as [Hk|Hk]; [left|right]; auto. - elim Hn; exists e'; rewrite Hk; auto. + contradiction Hn; exists e'; rewrite Hk; auto. assert (f k e = f k' e') by (apply Hf; auto). congruence. Qed. @@ -1345,11 +1344,11 @@ Section MoreWeakFacts. set (f':=fun k e b => if f k e then b else false). intro m. pattern m, (fold f' m true). apply fold_rec. - intros m' Hm'. split; auto. intros _ k e Hke. elim (Hm' k e); auto. + intros m' Hm'. split; auto. intros _ k e Hke. contradiction (Hm' k e); auto. intros k e b m1 m2 _ Hn Hadd IH. clear m. change (Equal m2 (add k e m1)) in Hadd. - unfold f'; simpl. case_eq (f k e); intros Hfke. + unfold f'; cbn. case_eq (f k e); intros Hfke. (* f k e = true *) rewrite IH. clear IH. split; intros Hmapsto k' e' Hke'. rewrite Hadd, add_mapsto_iff in Hke'. @@ -1372,18 +1371,18 @@ Section MoreWeakFacts. intro m. pattern m, (fold f' m false). apply fold_rec. intros m' Hm'. split; try (intros; discriminate). - intros ((k,e),(Hke,_)); simpl in *. elim (Hm' k e); auto. + intros ((k,e),(Hke,_)); cbn in *. contradiction (Hm' k e); auto. intros k e b m1 m2 _ Hn Hadd IH. clear m. change (Equal m2 (add k e m1)) in Hadd. - unfold f'; simpl. case_eq (f k e); intros Hfke. + unfold f'; cbn. case_eq (f k e); intros Hfke. (* f k e = true *) split; [intros _|auto]. - exists (k,e); simpl; split; auto. + exists (k,e); cbn; split; auto. rewrite Hadd, add_mapsto_iff; auto. (* f k e = false *) - rewrite IH. clear IH. split; intros ((k',e'),(Hke1,Hke2)); simpl in *. - exists (k',e'); simpl; split; auto. + rewrite IH. clear IH. split; intros ((k',e'),(Hke1,Hke2)); cbn in *. + exists (k',e'); cbn; split; auto. rewrite Hadd, add_mapsto_iff; right; split; auto. intro abs; contradict Hn; exists e'; rewrite abs; auto. rewrite Hadd, add_mapsto_iff in Hke1. destruct Hke1 as [(?,?)|(?,?)]. @@ -1414,7 +1413,7 @@ Section MoreWeakFacts. m1 = fst (partition f m) -> (MapsTo k e m1 <-> MapsTo k e m /\ f k e = true). Proof using HF Hf. - unfold partition; simpl; intros. subst m1. + unfold partition; cbn; intros. subst m1. apply filter_iff; auto. Qed. @@ -1422,10 +1421,10 @@ Section MoreWeakFacts. m2 = snd (partition f m) -> (MapsTo k e m2 <-> MapsTo k e m /\ f k e = false). Proof using HF Hf. - unfold partition; simpl; intros. subst m2. + unfold partition; cbn; intros. subst m2. rewrite filter_iff. split; intros (H,H'); split; auto. - destruct (f k e); simpl in *; auto. + destruct (f k e); cbn in *; auto. rewrite H'; auto. repeat red; intros. f_equal. apply Hf; auto. Qed. @@ -1453,13 +1452,13 @@ Section MoreWeakFacts. destruct (In_dec m1 k) as [H|H]; [left|right]; auto. destruct Hm as (Hm,Hm'). destruct Hk as (e,He); rewrite Hm' in He; destruct He. - elim H; exists e; auto. + contradiction H; exists e; auto. exists e; auto. Defined. Lemma Disjoint_sym : forall m1 m2, Disjoint m1 m2 -> Disjoint m2 m1. Proof using . - intros m1 m2 H k (H1,H2). elim (H k); auto. + intros m1 m2 H k (H1,H2). contradiction (H k); auto. Qed. Lemma Partition_sym : forall m m1 m2, @@ -1475,10 +1474,10 @@ Section MoreWeakFacts. Proof using . intros m m1 m2 (Hdisj,Heq). split. intro He. - split; intros k e Hke; elim (He k e); rewrite Heq; auto. + split; intros k e Hke; contradiction (He k e); rewrite Heq; auto. intros (He1,He2) k e Hke. rewrite Heq in Hke. destruct Hke. - elim (He1 k e); auto. - elim (He2 k e); auto. + contradiction (He1 k e); auto. + contradiction (He2 k e); auto. Qed. Lemma Partition_Add : @@ -1507,13 +1506,13 @@ Section MoreWeakFacts. destruct (equiv_dec x k) as [He|Hne]; auto. rewrite <- He; apply find_1; auto. (* disjoint *) - intros k (H1,H2). elim (Hdisj k). split; auto. + intros k (H1,H2). contradiction (Hdisj k). split; auto. rewrite remove_in_iff in H1; destruct H1; auto. (* mapsto *) intros k' e'. rewrite Heq, 2 remove_mapsto_iff, Hor. intuition. - intro abs; elim (Hdisj x); split; [exists e|exists e']; auto. + intro abs; contradiction (Hdisj x); split; [exists e|exists e']; auto. apply MapsTo_1 with k'; auto. (* second case : x in m2 *) @@ -1525,13 +1524,13 @@ Section MoreWeakFacts. destruct (equiv_dec x k) as [He|Hne]; auto. rewrite <- He; apply find_1; auto. (* disjoint *) - intros k (H1,H2). elim (Hdisj k). split; auto. + intros k (H1,H2). contradiction (Hdisj k). split; auto. rewrite remove_in_iff in H2; destruct H2; auto. (* mapsto *) intros k' e'. rewrite Heq, 2 remove_mapsto_iff, Hor. intuition. - intro abs; elim (Hdisj x); split; [exists e'|exists e]; auto. + intro abs; contradiction (Hdisj x); split; [exists e'|exists e]; auto. apply MapsTo_1 with k'; auto. Qed. @@ -1591,7 +1590,7 @@ Section MoreWeakFacts. replace (fold f m 0) with (fold f m1 (fold f m2 0)). rewrite <- cardinal_fold. intros. - apply fold_rel with (R:=fun u v => u = v + cardinal m2); simpl; auto. + apply fold_rel with (R:=fun u v => u = v + cardinal m2); cbn; auto. symmetry; apply Partition_fold with (eqA:=@Logic.eq _); try red; auto. compute; auto. Qed. @@ -1613,15 +1612,15 @@ Section MoreWeakFacts. rewrite Hm'. intuition. exists e; auto. - elim (Hm k); split; auto; exists e; auto. + contradiction (Hm k); split; auto; exists e; auto. rewrite (@partition_iff_2 f Hf m m2') by auto. unfold f. rewrite <- not_mem_in_iff. destruct Hm as (Hm,Hm'). rewrite Hm'. intuition. - elim (Hm k); split; auto; exists e; auto. - elim H1; exists e; auto. + contradiction (Hm k); split; auto; exists e; auto. + contradiction H1; exists e; auto. Qed. Lemma update_mapsto_iff : forall m m' k e, @@ -1635,7 +1634,7 @@ Section MoreWeakFacts. intros m0 Hm0 k e. assert (~In k m0) by (intros (e0,He0); apply (Hm0 k e0); auto). intuition. - elim (Hm0 k e); auto. + contradiction (Hm0 k e); auto. intros k e m0 m1 m2 _ Hn Hadd IH k' e'. change (Equal m2 (add k e m1)) in Hadd. @@ -1647,7 +1646,7 @@ Section MoreWeakFacts. Proof. intros m m' k e H. rewrite update_mapsto_iff in H. destruct (In_dec m' k) as [H'|H']; [left|right]; intuition. - elim H'; exists e; auto. + contradiction H'; exists e; auto. Defined. Lemma update_in_iff : forall m m' k, @@ -1659,7 +1658,7 @@ Section MoreWeakFacts. destruct (In_dec m' k) as [H|H]. destruct H as (e,H). intros _; exists e. rewrite update_mapsto_iff; left; auto. - destruct 1 as [H'|H']; [|elim H; auto]. + destruct 1 as [H'|H']; [|contradiction H; auto]. destruct H' as (e,H'). exists e. rewrite update_mapsto_iff; right; auto. Qed. @@ -1751,7 +1750,7 @@ Section MoreWeakFacts. intros k k' e e' i Hneq x. rewrite !add_o. destruct (equiv_dec k x); destruct (equiv_dec k' x); auto. - elim Hneq. etransitivity; eauto. + contradiction Hneq. etransitivity; eauto. apply fold_init with (eqA:=Equal); auto. intros k k' Hk e e' He m m' Hm; rewrite Hk,He,Hm; red; auto. Qed. @@ -1766,14 +1765,14 @@ Section MoreWeakFacts. pattern (mem k m2); rewrite Hm2. (* UGLY, see with Matthieu *) destruct (mem k m2'); rewrite Hii'; auto. apply fold_Equal with (eqA:=Equal); auto. - intros k k' Hk e e' He m m' Hm; simpl in *. + intros k k' Hk e e' He m m' Hm; cbn in *. pattern (mem k m2'); rewrite Hk. (* idem *) destruct (mem k' m2'); rewrite ?Hk,?He,Hm; red; auto. intros k k' e e' i Hneq x. case_eq (mem k m2'); case_eq (mem k' m2'); intros; auto. rewrite !add_o. destruct (equiv_dec k x); destruct (equiv_dec k' x); auto. - elim Hneq. etransitivity; eauto. + contradiction Hneq. etransitivity; eauto. Qed. Global Instance diff_m elt : @@ -1784,16 +1783,16 @@ Section MoreWeakFacts. apply fold_rel with (R:=Equal); try red; auto. intros k e i i' H Hii' x. pattern (mem k m2); rewrite Hm2. (* idem *) - destruct (mem k m2'); simpl; rewrite Hii'; auto. + destruct (mem k m2'); cbn; rewrite Hii'; auto. apply fold_Equal with (eqA:=Equal); auto. - intros k k' Hk e e' He m m' Hm; simpl in *. + intros k k' Hk e e' He m m' Hm; cbn in *. pattern (mem k m2'); rewrite Hk. (* idem *) - destruct (mem k' m2'); simpl; rewrite ?Hk,?He,Hm; red; auto. + destruct (mem k' m2'); cbn; rewrite ?Hk,?He,Hm; red; auto. intros k k' e e' i Hneq x. - case_eq (mem k m2'); case_eq (mem k' m2'); intros; simpl; auto. + case_eq (mem k m2'); case_eq (mem k' m2'); intros; cbn; auto. rewrite !add_o. destruct (equiv_dec k x); destruct (equiv_dec k' x); auto. - elim Hneq. etransitivity; eauto. + contradiction Hneq. etransitivity; eauto. Qed. End MoreWeakFacts. @@ -1836,7 +1835,7 @@ Section OrdProperties. apply SortA_equivlistA_eqlistA; auto. Qed. - Ltac clean_eauto := unfold K.eq_key_elt, K.ltk; simpl; + Ltac clean_eauto := unfold K.eq_key_elt, K.ltk; cbn; intuition; try solve [order]. Definition gtb (p p':key*elt) := @@ -1848,26 +1847,26 @@ Section OrdProperties. Lemma gtb_1 : forall p p', gtb p p' = true <-> ltk p' p. Proof. - intros (x,e) (y,e'); unfold gtb, K.ltk; simpl. + intros (x,e) (y,e'); unfold gtb, K.ltk; cbn. destruct (compare_dec x y); intuition; try discriminate; order. Qed. Lemma leb_1 : forall p p', leb p p' = true <-> ~ltk p' p. Proof. - intros (x,e) (y,e'); unfold leb, gtb, K.ltk; simpl. + intros (x,e) (y,e'); unfold leb, gtb, K.ltk; cbn. destruct (compare_dec x y); intuition; try discriminate; order. Qed. Lemma gtb_compat : forall p, Proper (eq_key_elt==>eq) (gtb p). Proof. - red; intros (x,e) (a,e') (b,e'') H; red in H; simpl in *; destruct H. + red; intros (x,e) (a,e') (b,e'') H; red in H; cbn in *; destruct H. generalize (gtb_1 (x,e) (a,e'))(gtb_1 (x,e) (b,e'')); destruct (gtb (x,e) (a,e')); destruct (gtb (x,e) (b,e'')); auto. - unfold KeyOrderedType.ltk in *; simpl in *; intros. + unfold KeyOrderedType.ltk in *; cbn in *; intros. symmetry; rewrite H2. apply eq_lt with a; auto. rewrite <- H1; auto. - unfold KeyOrderedType.ltk in *; simpl in *; intros. + unfold KeyOrderedType.ltk in *; cbn in *; intros. rewrite H1. apply eq_lt with b; auto. rewrite <- H2; auto. @@ -1887,8 +1886,8 @@ Section OrdProperties. unfold elements_lt, elements_ge, leb; intros. apply filter_split with (eqA:=eq_key) (ltA:=ltk); eauto with map. intros; destruct x; destruct y; destruct p. - rewrite gtb_1 in H; unfold K.ltk in H; simpl in *. - unfold gtb, K.ltk in *; simpl in *. + rewrite gtb_1 in H; unfold K.ltk in H; cbn in *. + unfold gtb, K.ltk in *; cbn in *. destruct (compare_dec k1 k0); intuition; try discriminate; order. Qed. @@ -1910,7 +1909,7 @@ Section OrdProperties. intros. rewrite filter_InA in H1; auto with *; destruct H1. rewrite leb_1 in H2. - destruct y; unfold KeyOrderedType.ltk in *; simpl in *. + destruct y; unfold KeyOrderedType.ltk in *; cbn in *. rr; rewrite <- elements_mapsto_iff in H1. assert (~ x == k). contradict H. @@ -1920,24 +1919,24 @@ Section OrdProperties. intros. rewrite filter_InA in H1; auto with *; destruct H1. rewrite gtb_1 in H3. - destruct y; destruct x0; unfold KeyOrderedType.ltk in *; simpl in *. + destruct y; destruct x0; unfold KeyOrderedType.ltk in *; cbn in *. inversion_clear H2. - red in H4; simpl in *; destruct H4. + red in H4; cbn in *; destruct H4. order. rewrite filter_InA in H4; auto with *; destruct H4. rewrite leb_1 in H4. - unfold KeyOrderedType.ltk in *; simpl in *; order. + unfold KeyOrderedType.ltk in *; cbn in *; order. rr; red; intros a; destruct a. rewrite InA_app_iff, InA_cons, 2 filter_InA, <-2 elements_mapsto_iff, leb_1, gtb_1, find_mapsto_iff, (H0 k), <- find_mapsto_iff, add_mapsto_iff; try apply eq_key_elt_Equiv; auto with *. - unfold eq_key_elt, KeyOrderedType.eq_key_elt, KeyOrderedType.ltk; simpl. + unfold eq_key_elt, KeyOrderedType.eq_key_elt, KeyOrderedType.ltk; cbn. destruct (compare_dec k x); replace (x =/= k) with (~ x == k) in *; intuition. right; split; auto; order. order. - elim H. + contradiction H. exists e0; apply MapsTo_1 with k; auto. right; right; split; auto; order. order. @@ -1956,15 +1955,15 @@ Section OrdProperties. destruct x0; destruct y. rr; rewrite <- elements_mapsto_iff in H1. unfold KeyOrderedType.eq_key_elt, KeyOrderedType.ltk in *; - simpl in *; destruct H3. + cbn in *; destruct H3. apply lt_eq with x; auto. - apply H; simpl in *; subst; exists e0; assumption. + apply H; cbn in *; subst; exists e0; assumption. inversion H3. red; intros a; destruct a. rr; rewrite InA_app_iff, InA_cons, InA_nil, <- 2 elements_mapsto_iff, find_mapsto_iff, (H0 k), <- find_mapsto_iff, add_mapsto_iff by (apply eq_key_elt_Equiv). - unfold eq_key_elt, KeyOrderedType.eq_key_elt, complement; simpl; intuition. + unfold eq_key_elt, KeyOrderedType.eq_key_elt, complement; cbn; intuition. destruct (equiv_dec x k); intuition auto. exfalso. assert (In k m). @@ -1986,7 +1985,7 @@ Section OrdProperties. destruct y; destruct x0. rr; rewrite <- elements_mapsto_iff in H2. unfold KeyOrderedType.eq_key_elt, KeyOrderedType.ltk in *; - simpl in *; destruct H3. + cbn in *; destruct H3. apply eq_lt with x; auto. apply H; exists e0; assumption. inversion H3. @@ -1994,7 +1993,7 @@ Section OrdProperties. rewrite InA_cons, <- 2 elements_mapsto_iff, find_mapsto_iff, (H0 k), <- find_mapsto_iff, add_mapsto_iff by (apply eq_key_elt_Equiv). - unfold eq_key_elt, KeyOrderedType.eq_key_elt; simpl. intuition. + unfold eq_key_elt, KeyOrderedType.eq_key_elt; cbn. intuition. destruct (equiv_dec x k); auto. exfalso. assert (In k m). @@ -2038,30 +2037,30 @@ Section OrdProperties. generalize (elements_3 m). revert x e H y x0 H0 H1. induction (elements m). - simpl; intros; try discriminate. + cbn; intros; try discriminate. intros. - destruct a; destruct l; simpl in *. + destruct a; destruct l; cbn in *. injection H; clear H; intros; subst. inversion_clear H1. - repeat red in H; simpl in *. destruct H; order. - elim H0; eauto. - inversion H; simpl in *. + repeat red in H; cbn in *. destruct H; order. + contradiction H0; eauto. + inversion H; cbn in *. change (max_elt_aux (p::l) = Some (x,e)) in H. generalize (IHl x e H); clear IHl; intros IHl. inversion_clear H1; [ | inversion_clear H2; eauto ]. - red in H3; simpl in H3; destruct H3. + red in H3; cbn in H3; destruct H3. destruct p as (p1,p2). destruct (equiv_dec p1 x). apply lt_eq with p1; auto. inversion_clear H2. inversion_clear H5. - simpl in *; subst; rewrite H1. + cbn in *; subst; rewrite H1. inversion H6; exact H5. - simpl in *; subst. + cbn in *; subst. transitivity p1; auto. inversion_clear H2. inversion_clear H5. - red in H2; simpl in H2; order. + red in H2; cbn in H2; order. inversion_clear H2. eapply IHl; eauto. intro Z; apply H4; order. @@ -2074,8 +2073,8 @@ Section OrdProperties. unfold max_elt in *. rewrite elements_mapsto_iff. induction (elements m). - simpl; try discriminate. - destruct a; destruct l; simpl in *. + cbn; try discriminate. + destruct a; destruct l; cbn in *. injection H; intros; subst; constructor; red; auto. constructor 2; auto. Qed. @@ -2087,7 +2086,7 @@ Section OrdProperties. unfold max_elt in *. rewrite elements_Empty. induction (elements m); auto. - destruct a; destruct l; simpl in *; try discriminate. + destruct a; destruct l; cbn in *; try discriminate. assert (H':=IHl H); discriminate. Qed. @@ -2109,16 +2108,16 @@ Section OrdProperties. try discriminate. destruct p; injection H; intros; subst. inversion_clear H1. - red in H2; destruct H2; simpl in *; order. + red in H2; destruct H2; cbn in *; order. inversion_clear H4. rewrite (@InfA_alt _ eq_key_elt) in H3; eauto with *. apply (H3 (y,x0)); auto. constructor; repeat intro. - destruct x1; repeat red in H4; simpl in H4; order. + destruct x1; repeat red in H4; cbn in H4; order. destruct x1; destruct y0; destruct z. - unfold lt_key, KeyOrderedType.ltk in *; simpl in *; order. + unfold lt_key, KeyOrderedType.ltk in *; cbn in *; order. unfold KeyOrderedType.eq_key_elt, lt_key, KeyOrderedType.ltk; - repeat intro; simpl in *; intuition order. + repeat intro; cbn in *; intuition order. Qed. Lemma min_elt_MapsTo : @@ -2128,8 +2127,8 @@ Section OrdProperties. unfold min_elt in *. rewrite elements_mapsto_iff. destruct (elements m). - simpl; try discriminate. - destruct p; simpl in *. + cbn; try discriminate. + destruct p; cbn in *. injection H; intros; subst; constructor; red; auto. Qed. @@ -2140,7 +2139,7 @@ Section OrdProperties. unfold min_elt in *. rewrite elements_Empty. destruct (elements m); auto. - destruct p; simpl in *; discriminate. + destruct p; cbn in *; discriminate. Qed. End Min_Max_Elt. @@ -2219,7 +2218,7 @@ Section OrdProperties. do 2 rewrite fold_1. do 2 rewrite <- fold_left_rev_right. apply fold_right_eqlistA with (eqA:=eq_key_elt) (eqB:=eqA); auto. - intros (k,e) (k',e') (Hk,He) a a' Ha; simpl in *; apply Hf; auto. + intros (k,e) (k',e') (Hk,He) a a' Ha; cbn in *; apply Hf; auto. apply eqlistA_rev. apply elements_Equal_eqlistA. auto. Qed. @@ -2235,10 +2234,10 @@ Section OrdProperties. transitivity (fold_right f' i (rev (elements m1 ++ (x,e)::nil))). apply fold_right_eqlistA with (eqA:=eq_key_elt) (eqB:=eqA); auto. intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; - unfold f'; simpl in *; apply H; auto. + unfold f'; cbn in *; apply H; auto. apply eqlistA_rev. apply elements_Add_Above; auto. - rewrite distr_rev; simpl. + rewrite distr_rev; cbn. reflexivity. Qed. @@ -2254,10 +2253,10 @@ Section OrdProperties. transitivity (fold_right f' i (rev (((x,e)::nil)++elements m1))). apply fold_right_eqlistA with (eqA:=eq_key_elt) (eqB:=eqA); auto. intros (k1,e1) (k2,e2) (Hk,He) a1 a2 Ha; - unfold f'; simpl in *; apply H; auto. + unfold f'; cbn in *; apply H; auto. apply eqlistA_rev. - simpl; apply elements_Add_Below; auto. - rewrite distr_rev; simpl. + cbn; apply elements_Add_Below; auto. + rewrite distr_rev; cbn. rewrite fold_right_app. reflexivity. Qed. diff --git a/Util/FMaps/FMapInterface.v b/Util/FMaps/FMapInterface.v index f33579924c1c579b570b5f475d1a1593d63ce68e..e8495defc32392b02ca0557398c4c563dd287674 100644 --- a/Util/FMaps/FMapInterface.v +++ b/Util/FMaps/FMapInterface.v @@ -2,6 +2,7 @@ Require Import SetoidList. Require Import SetoidDec. +Require Import Pactole.Util.SetoidDefs. (* Require Import Morphisms. *) Generalizable All Variables. @@ -11,24 +12,6 @@ Generalizable All Variables. Definition Cmp {elt : Type} (cmp : elt -> elt -> bool) e1 e2 := cmp e1 e2 = true. -Instance prod_Setoid A B {SA : Setoid A} {SB : Setoid B} : Setoid (A * B). -simple refine {| equiv := fun xn yp => fst xn == fst yp /\ snd xn == snd yp |}; auto; []. -Proof. split. -+ repeat intro; now split. -+ repeat intro; split; now symmetry. -+ intros ? ? ? [? ?] [? ?]; split; etransitivity; eauto. -Defined. - -Instance prod_EqDec A B `{EqDec A} `{EqDec B} : @EqDec (A * B) _. -refine (fun xn yp => if equiv_dec (fst xn) (fst yp) then - if equiv_dec (snd xn) (snd yp) then left _ else right _ - else right _). -Proof. -- now split. -- abstract (intros [? ?]; contradiction). -- abstract (intros [? ?]; contradiction). -Defined. - (** * [FMaps] : the interface of maps We define the class [FMap] of structures that implement finite @@ -94,7 +77,7 @@ Class FMap `{EqDec key} := { FMap_Setoid : forall `{Setoid elt}, Setoid (dict elt); FMap_EqDec :> forall `{EqDec elt}, @EqDec (dict elt) FMap_Setoid *) }. -Arguments dict key%type_scope {S0} {H} {FMap} elt%type_scope. +Arguments dict key%_type_scope {S0} {H} {FMap} elt%_type_scope. (** Map notations (see below) are interpreted in scope [map_scope], delimited with key [scope]. We bind it to the type [map] and to @@ -102,18 +85,18 @@ Arguments dict key%type_scope {S0} {H} {FMap} elt%type_scope. Declare Scope map_scope. Delimit Scope map_scope with map. Bind Scope map_scope with dict. -Arguments MapsTo {key%type_scope} {_} {_} {_} {elt%type_scope} _ _ _%map_scope. -Arguments is_empty {key%type_scope} {_} {_} {_} {elt%type_scope} _%map_scope. -Arguments mem {key%type_scope} {_} {_} {_} {elt%type_scope} _ _%map_scope. -Arguments add {key%type_scope} {_} {_} {_} {elt%type_scope} _ _ _%map_scope. -Arguments find {key%type_scope} {_} {_} {_} {elt%type_scope} _ _%map_scope. -Arguments remove {key%type_scope} {_} {_} {_} {elt%type_scope} _ _%map_scope. -Arguments equal {key%type_scope} {_} {_} {_} {elt%type_scope} _ _%map_scope _%map_scope. -Arguments map {key%type_scope} {_} {_} {_} {elt%type_scope} {elt'%type_scope} _ _%map_scope. -Arguments mapi {key%type_scope} {_} {_} {_} {elt%type_scope} {elt'%type_scope} _ _%map_scope. -Arguments fold {key%type_scope} {_} {_} {_} {elt%type_scope} {_%type_scope} _ _%map_scope _. -Arguments cardinal {key%type_scope} {_} {_} {_} {elt%type_scope} _%map_scope. -Arguments elements {key%type_scope} {_} {_} {_} {elt%type_scope} _%map_scope. +Arguments MapsTo {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _ _%_map_scope. +Arguments is_empty {key%_type_scope} {_} {_} {_} {elt%_type_scope} _%_map_scope. +Arguments mem {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _%_map_scope. +Arguments add {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _ _%_map_scope. +Arguments find {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _%_map_scope. +Arguments remove {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _%_map_scope. +Arguments equal {key%_type_scope} {_} {_} {_} {elt%_type_scope} _ _%_map_scope _%_map_scope. +Arguments map {key%_type_scope} {_} {_} {_} {elt%_type_scope} {elt'%_type_scope} _ _%_map_scope. +Arguments mapi {key%_type_scope} {_} {_} {_} {elt%_type_scope} {elt'%_type_scope} _ _%_map_scope. +Arguments fold {key%_type_scope} {_} {_} {_} {elt%_type_scope} {_%_type_scope} _ _%_map_scope _. +Arguments cardinal {key%_type_scope} {_} {_} {_} {elt%_type_scope} _%_map_scope. +Arguments elements {key%_type_scope} {_} {_} {_} {elt%_type_scope} _%_map_scope. (** All projections should be made opaque for tactics using [delta]-conversion, otherwise the underlying instances may appear during proofs, which then @@ -271,22 +254,22 @@ Class FMapSpecs_adjust `(FMap key) := { }. *) Class FMapSpecs `(F : FMap key) := { - FFMapSpecs_MapsTo :> FMapSpecs_MapsTo F; - FFMapSpecs_mem :> FMapSpecs_mem F; - FFMapSpecs_empty :> FMapSpecs_empty F; - FFMapSpecs_is_empty :> FMapSpecs_is_empty F; - FFMapSpecs_add :> FMapSpecs_add F; - FFMapSpecs_remove :> FMapSpecs_remove F; - FFMapSpecs_find :> FMapSpecs_find F; - FFMapSpecs_elements :> FMapSpecs_elements F; - FFMapSpecs_cardinal :> FMapSpecs_cardinal F; - FFMapSpecs_fold :> FMapSpecs_fold F; - FFMapSpecs_equal :> FMapSpecs_equal F; - FFMapSpecs_map :> FMapSpecs_map F; - FFMapSpecs_mapi :> FMapSpecs_mapi F; -(* FFMapSpecs_map2 :> FMapSpecs_map2 F; *) -(* FFMapSpecs_insert :> FMapSpecs_insert F; *) -(* FFMapSpecs_adjust :> FMapSpecs_adjust F *) + #[global] FFMapSpecs_MapsTo :: FMapSpecs_MapsTo F; + #[global] FFMapSpecs_mem :: FMapSpecs_mem F; + #[global] FFMapSpecs_empty :: FMapSpecs_empty F; + #[global] FFMapSpecs_is_empty :: FMapSpecs_is_empty F; + #[global] FFMapSpecs_add :: FMapSpecs_add F; + #[global] FFMapSpecs_remove :: FMapSpecs_remove F; + #[global] FFMapSpecs_find :: FMapSpecs_find F; + #[global] FFMapSpecs_elements :: FMapSpecs_elements F; + #[global] FFMapSpecs_cardinal :: FMapSpecs_cardinal F; + #[global] FFMapSpecs_fold :: FMapSpecs_fold F; + #[global] FFMapSpecs_equal :: FMapSpecs_equal F; + #[global] FFMapSpecs_map :: FMapSpecs_map F; + #[global] FFMapSpecs_mapi :: FMapSpecs_mapi F; +(* #[global] FFMapSpecs_map2 :: FMapSpecs_map2 F; *) +(* #[global] FFMapSpecs_insert :: FMapSpecs_insert F; *) +(* #[global] FFMapSpecs_adjust :: FMapSpecs_adjust F *) }. (* About Build_FMapSpecs. *) (* About FMapSpecs. *) diff --git a/Util/FMaps/FMapList.v b/Util/FMaps/FMapList.v index 4786e36fcbe12834a303321ef6f70c2cd70ab7ec..7b7378dca4ba00be3bc25cbd2071fef72b3e5394 100644 --- a/Util/FMaps/FMapList.v +++ b/Util/FMaps/FMapList.v @@ -87,9 +87,9 @@ Lemma list_add_3 : forall (m : t elt) x y e e', x =/= y -> InA (equiv@@1) (y, e) (list_add x e' m) -> InA (equiv@@1) (y, e) m. Proof using . intros m x y e e' Hxy Hy. simpl. induction m as [| [z p] l]. -+ inversion_clear Hy; try inversion H0. now elim Hxy. ++ inversion_clear Hy; try inversion H0. now contradiction Hxy. + simpl in *. destruct (equiv_dec x z). - - right. inversion_clear Hy; trivial. now elim Hxy. + - right. inversion_clear Hy; trivial. now contradiction Hxy. - inversion_clear Hy; now left + (right; apply IHl). Qed. @@ -168,7 +168,7 @@ Definition t_mapi {elt'} (f : key -> elt -> elt') (s : tt elt) : tt elt' := End ListOperations. (** The full set of operations. *) -Instance MapList key `{EqDec key} : FMap := {| +Global Instance MapList key `{EqDec key} : FMap := {| dict := fun elt => sig (@NoDupA (key * elt) (equiv@@1)); MapsTo := fun elt k e m => InA equiv (k, e) (proj1_sig m); empty := fun elt => (exist _ nil (NoDupA_nil (equiv@@1))); @@ -191,7 +191,7 @@ Local Notation eq_pair := (fun xn yp => fst xn == fst yp /\ snd xn = snd yp). (** ** Proofs of the specifications **) -Instance MapListFacts_MapsTo key `{EqDec key} : FMapSpecs_MapsTo MapList. +Global Instance MapListFacts_MapsTo key `{EqDec key} : FMapSpecs_MapsTo MapList. Proof using . split. intros elt [m Hm] x y e Hxy Hx. simpl in *. induction m; inversion_clear Hx. @@ -199,10 +199,10 @@ induction m; inversion_clear Hx. - right. inversion_clear Hm. now apply IHm. Qed. -Instance MapListFacts_mem key `{EqDec key} : FMapSpecs_mem MapList. +Global Instance MapListFacts_mem key `{EqDec key} : FMapSpecs_mem MapList. Proof using . split. * intros elt [m Hm] x [e Hin]. simpl in *. induction m as [| [y n] l]; inversion_clear Hin. - + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. elim Hxy. now destruct H0. + + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. contradiction Hxy. now destruct H0. + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. inversion_clear Hm. auto. * intros elt [m Hm] x Hmem. unfold In. simpl in *. induction m as [| [y n] l]; simpl in Hmem. + discriminate. @@ -212,18 +212,18 @@ Proof using . split. destruct Hmem as [e ?]. exists e. now right. Qed. -Instance MapListFacts_empty key `{EqDec key} : FMapSpecs_empty MapList. +Global Instance MapListFacts_empty key `{EqDec key} : FMapSpecs_empty MapList. Proof using . split. intros elt x e Hin. inversion Hin. Qed. -Instance MapListFacts_is_empty key `{EqDec key} : FMapSpecs_is_empty MapList. +Global Instance MapListFacts_is_empty key `{EqDec key} : FMapSpecs_is_empty MapList. Proof using . split. * intros elt [m Hm] Hm'. destruct m as [| [x n] l]; trivial. - elim Hm' with x n. now left. + contradiction Hm' with x n. now left. * intros elt [m Hm] Hm'. destruct m as [| [x n] l]; try discriminate. intros x n Hin. inversion Hin. Qed. -Instance MapListFacts_add key `{EqDec key} : FMapSpecs_add MapList. +Global Instance MapListFacts_add key `{EqDec key} : FMapSpecs_add MapList. Proof using . split. * intros elt [m Hm] x y e Hxy. simpl in *. induction m as [| [z p] l]; simpl. + now left. @@ -232,28 +232,28 @@ Proof using . split. + inversion Hy. + simpl. destruct (equiv_dec x z). - right. inversion_clear Hy; trivial. - elim Hxy. destruct H0. now transitivity z. + contradiction Hxy. destruct H0. now transitivity z. - inversion_clear Hm. inversion_clear Hy; now left + (right; apply IHl). * intros elt [m Hm] x y e e' Hxy Hy. simpl in *. induction m as [| [z p] l]. - + inversion_clear Hy; inversion H0. now elim Hxy. + + inversion_clear Hy; inversion H0. now contradiction Hxy. + simpl in *. destruct (equiv_dec x z). - - right. inversion_clear Hy; trivial. now elim Hxy. + - right. inversion_clear Hy; trivial. now contradiction Hxy. - inversion_clear Hm. inversion_clear Hy; now left + (right; apply IHl). Qed. -Instance MapListFacts_remove key `{EqDec key} : FMapSpecs_remove MapList. +Global Instance MapListFacts_remove key `{EqDec key} : FMapSpecs_remove MapList. Proof using . split. * intros elt [m Hm] x y Hxy. simpl. unfold t_remove, In. simpl. induction m as [| [z p] l]. + simpl. intros [? Habs]. inversion Habs. + simpl. inversion_clear Hm. destruct (equiv_dec x z) as [Hxz | Hxz]; auto; []. intros [n Habs]. inversion_clear Habs. - - elim Hxz. destruct H2. now transitivity y. + - contradiction Hxz. destruct H2. now transitivity y. - apply IHl; eauto. * intros elt [m Hm] x y e Hxy Hy. simpl in *. induction m as [| [z p] l]. + inversion Hy. + inversion_clear Hm. simpl. destruct (equiv_dec x z). - inversion_clear Hy; simpl in *; auto; []. - elim Hxy. destruct H2. now transitivity z. + contradiction Hxy. destruct H2. now transitivity z. - inversion_clear Hy; now left + (right; apply IHl). * intros elt [m Hm] x y e Hy. simpl in *. induction m as [| [z p] l]. + inversion_clear Hy. @@ -261,14 +261,14 @@ Proof using . split. inversion_clear Hy; now left + (right; apply IHl). Qed. -Instance MapListFacts_find key `{EqDec key} : FMapSpecs_find MapList. +Global Instance MapListFacts_find key `{EqDec key} : FMapSpecs_find MapList. Proof using . split. * intros elt [m Hm] x e Hin. simpl in *. induction m as [| [y p] l]. + inversion Hin. + inversion_clear Hm. simpl. destruct (equiv_dec x y). - inversion_clear Hin; try (now f_equal); []. assert (Heq : equiv@@1 (x, e) (y, p)) by assumption. - elim H0. eapply InA_eqA; eauto with typeclass_instances; []. + contradiction H0. eapply InA_eqA; eauto with typeclass_instances; []. revert H2. apply InA_impl_compat; trivial; []. now repeat intro. - inversion_clear Hin; now auto. @@ -279,17 +279,17 @@ Proof using . split. - right. auto. Qed. -Instance MapListFacts_elements key `{EqDec key} : FMapSpecs_elements MapList. +Global Instance MapListFacts_elements key `{EqDec key} : FMapSpecs_elements MapList. Proof using . split. * tauto. * tauto. * intros elt [m Hm]. simpl. assumption. Qed. -Instance MapListFacts_cardinal key `{EqDec key} : FMapSpecs_cardinal MapList. +Global Instance MapListFacts_cardinal key `{EqDec key} : FMapSpecs_cardinal MapList. Proof using . split. tauto. Qed. -Instance MapListFacts_fold key `{EqDec key} : FMapSpecs_fold MapList. +Global Instance MapListFacts_fold key `{EqDec key} : FMapSpecs_fold MapList. Proof using . split. intros elt [m Hm] A i f. simpl. revert i. induction m as [| [y p] l]; simpl. - reflexivity. @@ -361,7 +361,7 @@ intros [m Hm]. induction m as [| [k e] m]. Qed. -Instance MapListFacts_equal key `{EqDec key} : FMapSpecs_equal MapList. +Global Instance MapListFacts_equal key `{EqDec key} : FMapSpecs_equal MapList. Proof using . split. * unfold Equivb, equal. intuition. @@ -374,7 +374,7 @@ Proof using . split. firstorder. Qed. -Instance MapListFacts_map key `{EqDec key} : FMapSpecs_map MapList. +Global Instance MapListFacts_map key `{EqDec key} : FMapSpecs_map MapList. Proof using . split. * intros elt elt' [m Hm] x e f Hin. simpl in *. induction m as [| [y p] m]. + inversion Hin. @@ -382,7 +382,7 @@ Proof using . split. - left. simpl. split; trivial. inversion_clear Hin. -- simpl in *. destruct H2. now subst. -- assert (Heq : equiv@@1 (x, e) (y, p)) by assumption. - elim H0. eapply InA_eqA; eauto with typeclass_instances; []. + contradiction H0. eapply InA_eqA; eauto with typeclass_instances; []. revert H2. apply InA_impl_compat; trivial; []. now repeat intro. - inversion_clear Hin; try easy; []. @@ -396,7 +396,7 @@ Proof using . split. exists e'. now right. Qed. -Instance MapListFacts_mapi key `{EqDec key} : FMapSpecs_mapi MapList. +Global Instance MapListFacts_mapi key `{EqDec key} : FMapSpecs_mapi MapList. Proof using . split. * intros elt elt' [m Hm] x e f Hin. simpl in *. induction m as [| [y p] m]. + inversion Hin. @@ -405,7 +405,7 @@ Proof using . split. simpl. split; trivial. inversion_clear Hin. -- simpl in *. destruct H2. now subst. -- assert (Heq : equiv@@1 (x, e) (y, p)) by assumption. - elim H0. eapply InA_eqA; eauto with typeclass_instances; []. + contradiction H0. eapply InA_eqA; eauto with typeclass_instances; []. revert H2. apply InA_impl_compat; trivial; []. now repeat intro. - inversion_clear Hin; try easy; []. @@ -421,5 +421,5 @@ Proof using . split. Qed. (** The full set of specifications. *) -Instance MapListFacts key `{EqDec key} : FMapSpecs MapList. +Global Instance MapListFacts key `{EqDec key} : FMapSpecs MapList. Proof using . split; auto with typeclass_instances. Qed. diff --git a/Util/FSets/FSetFacts.v b/Util/FSets/FSetFacts.v index 43827f950af32cde32577a718eaab4320e851c81..1fc67182a3326171d6b4cfa01a046b0863cc6f20 100644 --- a/Util/FSets/FSetFacts.v +++ b/Util/FSets/FSetFacts.v @@ -22,6 +22,9 @@ Local Open Scope equiv_scope. #[export] Hint Extern 2 (equiv ?y ?x) => now symmetry : core. #[export] Hint Extern 2 (Equivalence.equiv ?y ?x) => now symmetry : core. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. + Notation Leibniz := (@eq _) (only parsing). (** * Specifications written using equivalences *) @@ -68,7 +71,7 @@ Section IffSpec. split; [| destruct 1; [apply add_1 | apply add_2]]; auto. destruct (Helt x y) as [E|E]. - intro. auto. - - intro H; right. eauto using add_3. + - intro H; right. eapply add_3;eauto. (* stopped working eauto using eauto *) Qed. Lemma add_other : x =/= y -> (In y (add x s) <-> In y s). @@ -122,7 +125,7 @@ Section IffSpec. Lemma exists_spec : exists_ f s = true <-> Exists (fun x => f x = true) s. Proof using HF Hf. split; [eapply exists_2 | eapply exists_1]; eauto. Qed. End ForFilter. - Arguments InA {A%type_scope} _ _ _. + Arguments InA {A%_type_scope} _ _ _. End IffSpec. @@ -293,14 +296,14 @@ Section BoolSpec. End BoolSpec. -Instance In_m `{HF : @FSetSpecs A St HA F} : +Global Instance In_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> iff) In. Proof using . intros x y H s s' H0. rewrite (In_eq_iff s H); auto. Qed. -Instance is_empty_m `{HF : @FSetSpecs A St HA F} : +Global Instance is_empty_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> @eq bool) is_empty. Proof using . intros s s' H. @@ -311,10 +314,10 @@ Proof using . - now rewrite H0, H, <- H1. Qed. -Instance Empty_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> iff) Empty. +Global Instance Empty_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> iff) Empty. Proof using . unfold Empty. intros ? ? H. now setoid_rewrite H. Qed. -Instance mem_m `{HF : @FSetSpecs A St HA F} : +Global Instance mem_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> @eq bool) mem. Proof using . intros x y H s s' H0. @@ -323,7 +326,7 @@ Proof using . destruct (mem x s); destruct (mem y s'); intuition. Qed. -Instance singleton_m `{HF : @FSetSpecs A St HA F} : +Global Instance singleton_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv) singleton. Proof using . intros x y H a. @@ -333,42 +336,42 @@ Proof using . - transitivity y; auto. Qed. -Instance add_m `{HF : @FSetSpecs A St HA F} : +Global Instance add_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> equiv) add. Proof using . intros x y H s s' H0 a. do 2 rewrite add_spec; rewrite H; rewrite H0; intuition. Qed. -Instance remove_m `{HF : @FSetSpecs A St HA F} : +Global Instance remove_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> equiv) remove. Proof using . intros x y H s s' H0 a. do 2 rewrite remove_spec; rewrite H; rewrite H0; intuition. Qed. -Instance union_m `{HF : @FSetSpecs A St HA F} : +Global Instance union_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> equiv) union. Proof using . intros s s' H s'' s''' H0 a. do 2 rewrite union_spec; rewrite H; rewrite H0; intuition. Qed. -Instance inter_m `{HF : @FSetSpecs A St HA F} : +Global Instance inter_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> equiv) inter. Proof using . intros s s' H s'' s''' H0 a. do 2 rewrite inter_spec; rewrite H; rewrite H0; intuition. Qed. -Instance diff_m `{HF : @FSetSpecs A St HA F} : +Global Instance diff_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> equiv) diff. Proof using . intros s s' H s'' s''' H0 a. do 2 rewrite diff_spec; rewrite H; rewrite H0; intuition. Qed. -Instance elements_compat `{HF : @FSetSpecs A St HA F} : +Global Instance elements_compat `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> PermutationA equiv) elements. Proof using . intros s s' Heq. apply NoDupA_equivlistA_PermutationA. @@ -378,13 +381,20 @@ intros s s' Heq. apply NoDupA_equivlistA_PermutationA. + intro x. rewrite 2 elements_spec. apply Heq. Qed. -Instance PermutationA_length {elt} `{Setoid elt} : Proper (PermutationA equiv ==> Logic.eq) (@length elt). -Proof using . clear. intros l1 l2 perm. induction perm; simpl; auto; congruence. Qed. - -Instance cardinal_compat `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> Logic.eq) cardinal. +Global Instance cardinal_compat `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> Logic.eq) cardinal. Proof using . intros ? ? Heq. rewrite 2 cardinal_spec. now do 2 f_equiv. Qed. -Instance Subset_m `{HF : @FSetSpecs A St HA F} : +Global Instance For_all_compat `{HF : @FSetSpecs A St HA F} : + Proper ((equiv ==> impl) ==> equiv ==> impl) For_all. +Proof using . +intros P Q HPQ s1 s2 Hs HP x Hin. +apply (HPQ x x ltac:(reflexivity)), HP. now rewrite Hs. +Qed. + +Global Instance PermutationA_length {elt} `{Setoid elt} : Proper (PermutationA equiv ==> Logic.eq) (@length elt). +Proof using . clear. intros l1 l2 perm. induction perm; simpl; auto; congruence. Qed. + +Global Instance Subset_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> iff) Subset. Proof using . unfold Subset; intros s s' H u u' H'; split; intros. @@ -392,7 +402,7 @@ Proof using . rewrite H'; apply H0; rewrite <- H; assumption. Qed. -Instance subset_m `{HF : @FSetSpecs A St HA F} : +Global Instance subset_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> @eq bool) subset. Proof using . intros s s' H s'' s''' H0. @@ -402,7 +412,7 @@ Proof using . rewrite H in H1; rewrite H0 in H1; intuition. Qed. -Instance equal_m `{HF : @FSetSpecs A St HA F} : +Global Instance equal_m `{HF : @FSetSpecs A St HA F} : Proper (equiv ==> equiv ==> @eq bool) equal. Proof using . intros s s' H s'' s''' H0. @@ -421,55 +431,55 @@ Lemma Subset_trans `{HF : @FSetSpecs A St HA F} : forall s s' s'', s[<=]s'->s'[<=]s''->s[<=]s''. Proof using . unfold Subset; eauto. Qed. -Instance SubsetSetoid `{@FSetSpecs A St HA F} : +Global Instance SubsetSetoid `{@FSetSpecs A St HA F} : PreOrder Subset := { PreOrder_Reflexive := Subset_refl; PreOrder_Transitive := Subset_trans }. (** * Set operations and morphisms *) -Instance In_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance In_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (equiv ==> Subset ++> impl) In | 1. Proof using . simpl_relation; apply H2; rewrite <- H1; auto. Qed. -Instance Empty_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance Empty_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (Subset --> impl) Empty. Proof using . simpl_relation; unfold Subset, Empty, impl; intros. exact (H2 a (H1 a H3)). Qed. -Instance add_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance add_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (equiv ==> Subset ++> Subset) add. Proof using . unfold Subset; intros x y H1 s s' H2 a. do 2 rewrite add_spec; rewrite H1; intuition. Qed. -Instance remove_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance remove_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (equiv ==> Subset ++> Subset) remove. Proof using . unfold Subset; intros x y H1 s s' H2 a. do 2 rewrite remove_spec; rewrite H1; intuition. Qed. -Instance union_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance union_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (Subset ++> Subset ++> Subset) union. Proof using . intros s s' H1 s'' s''' H2 a. do 2 rewrite union_spec; intuition. Qed. -Instance inter_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance inter_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (Subset ++> Subset ++> Subset) inter. Proof using . intros s s' H1 s'' s''' H2 a. do 2 rewrite inter_spec; intuition. Qed. -Instance diff_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance diff_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : Proper (Subset ++> Subset --> Subset) diff. Proof using . unfold Subset; intros s s' H1 s'' s''' H2 a. @@ -478,7 +488,7 @@ Qed. (** [fold], [filter], [for_all], [exists_] and [partition] require the additional hypothesis on [f]. *) -Instance filter_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance filter_m `{F : FSet A, @FSetSpecs _ _ _ F} : forall f `{Proper _ (equiv ==> @eq bool) f}, Proper (equiv ==> equiv) (filter f). Proof using . @@ -495,7 +505,7 @@ Proof using . red; repeat intro; rewrite <- 2 Hff'; auto. Qed. -Instance filter_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : +Global Instance filter_s_m `{F : FSet A, @FSetSpecs _ _ _ F} : forall f `{Proper _ (equiv ==> @eq bool) f}, Proper (Subset ==> Subset) (filter f). Proof using . @@ -673,7 +683,7 @@ Defined. (** * Map **) -Instance fold_compat {A B} `{FSetSpecs A} `{Setoid B} : +Global Instance fold_compat {A B} `{FSetSpecs A} `{Setoid B} : forall f : A -> B -> B, Proper (equiv ==> equiv ==> equiv) f -> transpose equiv f -> forall a, Proper (equiv ==> equiv) (fun x => fold f x a). Proof using . @@ -697,11 +707,11 @@ intros f Hf x l. induction l as [| e l]; intro acc; simpl. do 2 left. match goal with H : _ == e |- _ => now rewrite <- H end. Qed. -(* Instance elements_compat : Proper (equiv ==> PermutationA equiv) elements := elements_compat. *) +(* Global Instance elements_compat : Proper (equiv ==> PermutationA equiv) elements := elements_compat. *) Definition map {A B} `{FSet A} `{FSet B} (f : A -> B) s := fold (fun x acc => add (f x) acc) s empty. -Instance map_compat {A B} `{FSetSpecs A} `{FSetSpecs B} : forall f : A -> B, Proper (equiv ==> equiv) f -> +Global Instance map_compat {A B} `{FSetSpecs A} `{FSetSpecs B} : forall f : A -> B, Proper (equiv ==> equiv) f -> Proper (equiv ==> equiv) (map f). Proof using . intros f Hf mâ‚ mâ‚‚ Hm. unfold map. apply fold_compat; autoclass; [|]. @@ -735,7 +745,20 @@ split; intro Hin. + destruct Hin as [Hy | [? [Hin Hx]]]; eexists; set_iff; eauto. Qed. -Lemma map_injective_elements {A B} `{FSetSpecs A} `{FSetSpecs B} : forall f, +Lemma map_empty_iff_empty {A B} `{FSetSpecs A} `{FSetSpecs B} : + forall (f : A -> B), Proper (equiv ==> equiv) f -> + forall (s : set A), map f s == {} <-> s == {}. +Proof using . +intros f Hf s. split; intros Hempty. ++ destruct (choose_dec s) as [x Hx | Hs]. + - exfalso. specialize (Hempty (f x)). rewrite map_spec, empty_spec in Hempty; trivial; []. + rewrite <- Hempty. now exists x. + - intro x. specialize (Hs x). rewrite empty_spec. tauto. ++ rewrite Hempty. intro x. rewrite map_spec; trivial; []. setoid_rewrite empty_spec. firstorder. +Qed. +Arguments map_empty_iff_empty {A} {B} {_} {_} {_} {_} {_} {_} {_} {_} f Hf s. + +Lemma map_injective_elements {A B} `{FSetSpecs A} `{FSetSpecs B} : forall f : A -> B, Proper (equiv ==> equiv) f -> injective equiv equiv f -> forall s, PermutationA equiv (elements (map f s)) (List.map f (elements s)). @@ -750,3 +773,14 @@ apply NoDupA_equivlistA_PermutationA; autoclass. - exists x. rewrite elements_spec. auto. - exists x. rewrite <- elements_spec. now symmetry in Hx1. Qed. + +Lemma For_all_map `{FSetSpecs A} : forall P f (s : set A), + Proper (equiv ==> impl) P -> + Proper (equiv ==> equiv) f -> + For_all P (map f s) <-> For_all (fun x => P (f x)) s. +Proof using . +intros P f s HP Hf. split; intros Hall x Hin. ++ apply Hall. rewrite map_spec; trivial; []. now exists x. ++ rewrite map_spec in Hin; trivial; []. destruct Hin as [? [? Heq]]. + rewrite Heq. now apply Hall. +Qed. diff --git a/Util/FSets/FSetInterface.v b/Util/FSets/FSetInterface.v index 758c0ba555393cc851c3db7635e705a8f00dc2b0..0e5c1106efa3524a477d28d12645e92c7b46d223 100644 --- a/Util/FSets/FSetInterface.v +++ b/Util/FSets/FSetInterface.v @@ -94,7 +94,7 @@ Class FSet `{EqDec elt} := { SpecificOrderedType _ (Equal_pw set elt In) *) }. -Arguments set elt%type_scope {_} {_} {FSet}. +Arguments set elt%_type_scope {_} {_} {FSet}. (** Set notations (see below) are interpreted in scope [set_scope], delimited with elt [scope]. We bind it to the type [set] and to @@ -102,25 +102,25 @@ Arguments set elt%type_scope {_} {_} {FSet}. Declare Scope set_scope. Delimit Scope set_scope with set. Bind Scope set_scope with set. -Global Arguments In {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments is_empty {_%type_scope} {_} {_} {_} _%set_scope. -Global Arguments mem {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments add {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments remove {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments union {_%type_scope} {_} {_} {_} _%set_scope _%set_scope. -Global Arguments inter {_%type_scope} {_} {_} {_} _%set_scope _%set_scope. -Global Arguments diff {_%type_scope} {_} {_} {_} _%set_scope _%set_scope. -Global Arguments equal {_%type_scope} {_} {_} {_} _%set_scope _%set_scope. -Global Arguments subset {_%type_scope} {_} {_} {_} _%set_scope _%set_scope. -Global Arguments fold {_%type_scope} {_} {_} {_} {_} _ _%set_scope _. -Global Arguments for_all {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments exists_ {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments filter {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments partition {_%type_scope} {_} {_} {_} _ _%set_scope. -Global Arguments cardinal {_%type_scope} {_} {_} {_} _%set_scope. -Global Arguments elements {_%type_scope} {_} {_} {_} _%set_scope. -Global Arguments choose {_%type_scope} {_} {_} {_} _%set_scope. -(* Global Arguments min_elt {_%type_scope} {_} {_} {_} _%set_scope. *) +Global Arguments In {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments is_empty {_%_type_scope} {_} {_} {_} _%_set_scope. +Global Arguments mem {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments add {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments remove {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments union {_%_type_scope} {_} {_} {_} _%_set_scope _%_set_scope. +Global Arguments inter {_%_type_scope} {_} {_} {_} _%_set_scope _%_set_scope. +Global Arguments diff {_%_type_scope} {_} {_} {_} _%_set_scope _%_set_scope. +Global Arguments equal {_%_type_scope} {_} {_} {_} _%_set_scope _%_set_scope. +Global Arguments subset {_%_type_scope} {_} {_} {_} _%_set_scope _%_set_scope. +Global Arguments fold {_%_type_scope} {_} {_} {_} {_} _ _%_set_scope _. +Global Arguments for_all {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments exists_ {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments filter {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments partition {_%_type_scope} {_} {_} {_} _ _%_set_scope. +Global Arguments cardinal {_%_type_scope} {_} {_} {_} _%_set_scope. +Global Arguments elements {_%_type_scope} {_} {_} {_} _%_set_scope. +Global Arguments choose {_%_type_scope} {_} {_} {_} _%_set_scope. +(* Global Arguments min_elt {_%_type_scope} {_} {_} {_} _%_set_scope. *) (* Global Arguments max_elt {_%type_scope} {_} {_} {_} _%set_scope. *) (** All projections should be made opaque for tactics using [delta]-conversion, @@ -298,28 +298,28 @@ Class FSetSpecs_max_elt `(FSet A) := { }.*) Class FSetSpecs `(F : FSet A) := { - FFSetSpecs_In :> FSetSpecs_In F; - FFSetSpecs_mem :> FSetSpecs_mem F; - FFSetSpecs_equal :> FSetSpecs_equal F; - FFSetSpecs_subset :> FSetSpecs_subset F; - FFSetSpecs_empty :> FSetSpecs_empty F; - FFSetSpecs_is_empty :> FSetSpecs_is_empty F; - FFSetSpecs_add :> FSetSpecs_add F; - FFSetSpecs_remove :> FSetSpecs_remove F; - FFSetSpecs_singleton :> FSetSpecs_singleton F; - FFSetSpecs_union :> FSetSpecs_union F; - FFSetSpecs_inter :> FSetSpecs_inter F; - FFSetSpecs_diff :> FSetSpecs_diff F; - FFSetSpecs_fold :> FSetSpecs_fold F; - FFSetSpecs_cardinal :> FSetSpecs_cardinal F; - FFSetSpecs_filter :> FSetSpecs_filter F; - FFSetSpecs_for_all :> FSetSpecs_for_all F; - FFSetSpecs_exists :> FSetSpecs_exists F; - FFSetSpecs_partition :> FSetSpecs_partition F; - FFSetSpecs_elements :> FSetSpecs_elements F; - FFSetSpecs_choose :> FSetSpecs_choose F; - (* FFSetSpecs_min_elt :> FSetSpecs_min_elt F; *) - (* FFSetSpecs_max_elt :> FSetSpecs_max_elt F *) + #[global] FFSetSpecs_In :: FSetSpecs_In F; + #[global] FFSetSpecs_mem :: FSetSpecs_mem F; + #[global] FFSetSpecs_equal :: FSetSpecs_equal F; + FFSetSpecs_subset :: FSetSpecs_subset F; + #[global] FFSetSpecs_empty :: FSetSpecs_empty F; + #[global] FFSetSpecs_is_empty :: FSetSpecs_is_empty F; + #[global] FFSetSpecs_add :: FSetSpecs_add F; + #[global] FFSetSpecs_remove :: FSetSpecs_remove F; + #[global] FFSetSpecs_singleton :: FSetSpecs_singleton F; + #[global] FFSetSpecs_union :: FSetSpecs_union F; + #[global] FFSetSpecs_inter :: FSetSpecs_inter F; + #[global] FFSetSpecs_diff :: FSetSpecs_diff F; + #[global] FFSetSpecs_fold :: FSetSpecs_fold F; + #[global] FFSetSpecs_cardinal :: FSetSpecs_cardinal F; + #[global] FFSetSpecs_filter :: FSetSpecs_filter F; + #[global] FFSetSpecs_for_all :: FSetSpecs_for_all F; + #[global] FFSetSpecs_exists :: FSetSpecs_exists F; + #[global] FFSetSpecs_partition :: FSetSpecs_partition F; + #[global] FFSetSpecs_elements :: FSetSpecs_elements F; + #[global] FFSetSpecs_choose :: FSetSpecs_choose F; + (* #[global] FFSetSpecs_min_elt :: FSetSpecs_min_elt F; *) + (* #[global]FFSetSpecs_max_elt :: FSetSpecs_max_elt F *) }. (* About FSetSpecs. *) diff --git a/Util/FSets/FSetList.v b/Util/FSets/FSetList.v index 4405d949414dc3bd406adb1567d07f7911f4c48d..e1b5f083a08b0b985bdc79549ba6f97cc5981b3b 100644 --- a/Util/FSets/FSetList.v +++ b/Util/FSets/FSetList.v @@ -6,7 +6,8 @@ Require Import SetoidDec. Require Pactole.Util.Coqlib. Require Import Pactole.Util.FSets.FSetInterface. - +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Set Implicit Arguments. Open Scope signature_scope. @@ -257,9 +258,9 @@ Lemma list_add_3 : forall (m : t elt) x y e e', x =/= y -> InA (equiv@@1) (y, e) (list_add x e' m) -> InA (equiv@@1) (y, e) m. Proof. intros m x y e e' Hxy Hy. simpl. induction m as [| [z p] l]. -+ inversion_clear Hy; try inversion H0. now elim Hxy. ++ inversion_clear Hy; try inversion H0. now contradiction Hxy. + simpl in *. destruct (equiv_dec x z). - - right. inversion_clear Hy; trivial. now elim Hxy. + - right. inversion_clear Hy; trivial. now contradiction Hxy. - inversion_clear Hy; now left + (right; apply IHl). Qed.*) @@ -320,7 +321,7 @@ Arguments elements {_%type_scope} {_} _%set_scope. Arguments choose {_%type_scope} {_} _%set_scope. *) (** The full set of operations. *) -Instance SetList elt `{EqDec elt} : FSet := {| +Global Instance SetList elt `{EqDec elt} : FSet := {| set := sig (@NoDupA elt equiv); In := fun x s => InA equiv x (proj1_sig s); empty := exist _ nil (NoDupA_nil equiv); @@ -350,7 +351,7 @@ Instance SetList elt `{EqDec elt} : FSet := {| (** ** Proofs of the specifications **) (* -Instance MapListFacts_MapsTo key `{EqDec key} : FMapSpecs_MapsTo (MapList _). +Global Instance MapListFacts_MapsTo key `{EqDec key} : FMapSpecs_MapsTo (MapList _). Proof. split. intros elt [m Hm] x y e Hxy Hx. simpl in *. induction m; inversion_clear Hx. @@ -358,10 +359,10 @@ induction m; inversion_clear Hx. - right. inversion_clear Hm. now apply IHm. Qed. -Instance MapListFacts_mem key `{EqDec key} : FMapSpecs_mem (MapList _). +Global Instance MapListFacts_mem key `{EqDec key} : FMapSpecs_mem (MapList _). Proof. split. * intros elt [m Hm] x [e Hin]. simpl in *. induction m as [| [y n] l]; inversion_clear Hin. - + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. elim Hxy. now destruct H0. + + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. contradiction Hxy. now destruct H0. + simpl. destruct (equiv_dec x y) as [Hxy | Hxy]; trivial. inversion_clear Hm. auto. * intros elt [m Hm] x Hmem. unfold In. simpl in *. induction m as [| [y n] l]; simpl in Hmem. + discriminate. @@ -370,13 +371,13 @@ Proof. split. - inversion_clear Hm. apply IHl in Hmem; trivial; []. destruct Hmem as [e ?]. exists e. now right. Qed. -Instance MapListFacts_empty key `{EqDec key} : FMapSpecs_empty (MapList _). +Global Instance MapListFacts_empty key `{EqDec key} : FMapSpecs_empty (MapList _). Proof. split. intros elt x e Hin. inversion Hin. Qed. -Instance MapListFacts_is_empty key `{EqDec key} : FMapSpecs_is_empty (MapList _). +Global Instance MapListFacts_is_empty key `{EqDec key} : FMapSpecs_is_empty (MapList _). Proof. split. * intros elt [m Hm] Hm'. destruct m as [| [x n] l]; trivial. - elim Hm' with x n. now left. + contradiction Hm' with x n. now left. * intros elt [m Hm] Hm'. destruct m as [| [x n] l]; try discriminate. intros x n Hin. inversion Hin. Qed. @@ -387,7 +388,7 @@ Local Transparent union inter diff equal subset fold for_all exists_ filter partition cardinal elements choose. -Instance SetListFacts_In elt `{EqDec elt} : FSetSpecs_In SetList. +Global Instance SetListFacts_In elt `{EqDec elt} : FSetSpecs_In SetList. Proof using . split. simpl. intros s x y Heq. now rewrite Heq. Qed. Local Lemma mem_aux elt `{EqDec elt} : forall x l, list_mem x l = true <-> InA equiv x l. @@ -401,16 +402,16 @@ intros x l. induction l as [| e l]; simpl. - inversion_clear Hin; assumption || contradiction. Qed. -Instance SetListFacts_mem elt `{EqDec elt} : FSetSpecs_mem SetList. +Global Instance SetListFacts_mem elt `{EqDec elt} : FSetSpecs_mem SetList. Proof using . split. * intros [s Hs] x Hin. simpl in *. now rewrite mem_aux. * intros [s Hs] x Hin. simpl in *. now rewrite mem_aux in Hin. Qed. -Instance SetLIst_Facts_empty elt `{EqDec elt} : FSetSpecs_empty SetList. +Global Instance SetLIst_Facts_empty elt `{EqDec elt} : FSetSpecs_empty SetList. Proof using . split. intros x Hin; simpl in *. now rewrite InA_nil in Hin. Qed. -Instance SetListFacts_is_empty elt `{EqDec elt} : FSetSpecs_is_empty SetList. +Global Instance SetListFacts_is_empty elt `{EqDec elt} : FSetSpecs_is_empty SetList. Proof using . split. * intros [[| e s] Hs] Hempty; simpl in *. + reflexivity. @@ -434,7 +435,7 @@ destruct (list_mem e l) eqn:Hmem. - destruct Hin; now left + right. Qed. -Instance SetListFacts_add elt `{EqDec elt} : FSetSpecs_add SetList. +Global Instance SetListFacts_add elt `{EqDec elt} : FSetSpecs_add SetList. Proof using . split. * intros [s Hs] x y Heq. simpl. rewrite add_aux. now left. * intros [s Hs] x y Hin. simpl in *. rewrite add_aux. now right. @@ -460,14 +461,14 @@ intros x e l Hnodup. induction l as [| e' l]; simpl. - destruct Hin as [[] ?]; now left + right. Qed. -Instance SetListFacts_remove elt `{EqDec elt} : FSetSpecs_remove SetList. +Global Instance SetListFacts_remove elt `{EqDec elt} : FSetSpecs_remove SetList. Proof using . split. * intros [s Hs] x y Heq. simpl. symmetry in Heq. now rewrite remove_aux. * intros [s Hs] x y Hxy Hin. simpl in *. now rewrite remove_aux. * intros [s Hs] x y Hin. simpl in *. now rewrite remove_aux in Hin. Qed. -Instance SetListFacts_singleton elt `{EqDec elt} : FSetSpecs_singleton SetList. +Global Instance SetListFacts_singleton elt `{EqDec elt} : FSetSpecs_singleton SetList. Proof using . split. * intros x y Hin. simpl in *. unfold list_add in Hin. destruct (list_mem x nil); now inversion_clear Hin. @@ -484,7 +485,7 @@ intros s1. induction s1 as [| e1 s1]; simpl; intros s2 x Hin. unfold list_add. now destruct (list_mem e1 s2); try right. Qed. -Instance SetListFacts_union elt `{EqDec elt} : FSetSpecs_union SetList. +Global Instance SetListFacts_union elt `{EqDec elt} : FSetSpecs_union SetList. Proof using . split. * intros [s1 Hs1] [s2 Hs2]. simpl. unfold list_union, list_fold, flip. @@ -542,7 +543,7 @@ destruct (x == e), (y == e). - apply IHl. Qed. -Instance SetListFacts_inter elt `{EqDec elt} : FSetSpecs_inter SetList. +Global Instance SetListFacts_inter elt `{EqDec elt} : FSetSpecs_inter SetList. Proof using . split. * intros [s1 Hs1] [s2 Hs2] x. simpl. unfold list_inter, list_fold, flip. rewrite inter_aux. @@ -569,14 +570,14 @@ intros l1 l2 x. revert l1. induction l2 as [| e2 l2]; simpl; intros l1 Hnodup. + now apply t_remove_lemma. Qed. -Instance SetListFacts_diff elt `{EqDec elt} : FSetSpecs_diff SetList. +Global Instance SetListFacts_diff elt `{EqDec elt} : FSetSpecs_diff SetList. Proof using . split; intros [s1 Hs1] [s2 Hs2] x; simpl; unfold list_diff, list_fold, flip; now rewrite diff_aux. Qed. -Instance SetListFacts_subset elt `{EqDec elt} : FSetSpecs_subset SetList. +Global Instance SetListFacts_subset elt `{EqDec elt} : FSetSpecs_subset SetList. Proof using . split. * intros s1 s2 Hle. simpl. unfold list_subset. change (is_empty (diff s1 s2) = true). @@ -588,11 +589,11 @@ Proof using . split. apply is_empty_2 in Hle. intros x Hin. specialize (Hle x). destruct (mem x s2) eqn:Hmem. + now apply mem_2. - + elim Hle. apply diff_3; trivial; []. + + contradiction Hle. apply diff_3; trivial; []. intro Habs. apply mem_1 in Habs. congruence. Qed. -Instance SetListFacts_equal elt `{EqDec elt} : FSetSpecs_equal SetList. +Global Instance SetListFacts_equal elt `{EqDec elt} : FSetSpecs_equal SetList. Proof using . split. * intros s1 s2 Heq. simpl. change (subset s1 s2 && subset s2 s1 = true). @@ -605,13 +606,13 @@ Proof using . split. split; now apply subset_2. Qed. -Instance SetListFacts_cardinal elt `{EqDec elt} : FSetSpecs_cardinal SetList. +Global Instance SetListFacts_cardinal elt `{EqDec elt} : FSetSpecs_cardinal SetList. Proof using . split. reflexivity. Qed. -Instance SetListFacts_fold elt `{EqDec elt} : FSetSpecs_fold SetList. +Global Instance SetListFacts_fold elt `{EqDec elt} : FSetSpecs_fold SetList. Proof using . split. reflexivity. Qed. -Instance SetListFacts_filter elt `{EqDec elt} : FSetSpecs_filter SetList. +Global Instance SetListFacts_filter elt `{EqDec elt} : FSetSpecs_filter SetList. Proof using . split. * intros [s Hs] x f Hf Hin. induction s as [| e s]; simpl in *. @@ -639,7 +640,7 @@ Proof using . split. rewrite <- Hfe, <- Hfx. now apply Hf. Qed. -Instance SetListFacts_for_all elt `{EqDec elt} : FSetSpecs_for_all SetList. +Global Instance SetListFacts_for_all elt `{EqDec elt} : FSetSpecs_for_all SetList. Proof using . split. * intros [s Hs] f Hf Hall. unfold For_all in Hall. simpl in *. rewrite forallb_forall. intros x Hin. apply Hall. apply In_InA; Preliminary.autoclass. @@ -649,7 +650,7 @@ Proof using . split. rewrite Heq. auto. Qed. -Instance SetListFacts_exists elt `{EqDec elt} : FSetSpecs_exists SetList. +Global Instance SetListFacts_exists elt `{EqDec elt} : FSetSpecs_exists SetList. Proof using . split. * intros [s Hs] f Hf [x [Hin Hfx]]. simpl in *. rewrite existsb_exists. rewrite InA_alt in Hin. destruct Hin as [y [Heq Hin]]. @@ -658,7 +659,7 @@ Proof using . split. destruct Hex as [x [Hin Hfx]]. exists x. simpl. split; trivial; []. apply In_InA; Preliminary.autoclass. Qed. -Instance SetListFacts_partition elt `{EqDec elt} : FSetSpecs_partition SetList. +Global Instance SetListFacts_partition elt `{EqDec elt} : FSetSpecs_partition SetList. Proof using . split. * intros [s Hs] f Hf x. simpl. induction s as [| e s]; simpl. + reflexivity. @@ -678,14 +679,14 @@ Proof using . split. (progress rewrite IHs in * || rewrite <- IHs in *); auto. Qed. -Instance SetListFacts_elements elt `{EqDec elt} : FSetSpecs_elements SetList. +Global Instance SetListFacts_elements elt `{EqDec elt} : FSetSpecs_elements SetList. Proof using . split. * simpl. auto. * simpl. auto. * intros [s Hs]. simpl. assumption. Qed. -Instance SetListFacts_choose elt `{EqDec elt} : FSetSpecs_choose SetList. +Global Instance SetListFacts_choose elt `{EqDec elt} : FSetSpecs_choose SetList. Proof using . split. * intros [[| e s] Hs] x Hin; simpl in *. + discriminate. @@ -697,5 +698,5 @@ Qed. (** The full set of specifications. *) -Instance SetListFacts elt `{EqDec elt} : FSetSpecs SetList. +Global Instance SetListFacts elt `{EqDec elt} : FSetSpecs SetList. Proof using . split; auto with typeclass_instances. Qed. diff --git a/Util/FSets/OrderedType.v b/Util/FSets/OrderedType.v index 2dfbc12a7ed272a2b89d8d0ffd3347bc0cc849ad..fc271f16a98c06ad5d344ab4542a94a3c072757f 100644 --- a/Util/FSets/OrderedType.v +++ b/Util/FSets/OrderedType.v @@ -310,9 +310,9 @@ Ltac normalize_notations := end. Ltac abstraction := match goal with - | H : False |- _ => elim H - | H : ?x <<< ?x |- _ => elim (lt_antirefl H) - | H : ?x =/= ?x |- _ => elim (H (reflexivity x)) + | H : False |- _ => contradiction H + | H : ?x <<< ?x |- _ => contradiction (lt_antirefl H) + | H : ?x =/= ?x |- _ => contradiction (H (reflexivity x)) | H : ?x === ?x |- _ => clear H; abstraction | H : ~?x <<< ?x |- _ => clear H; abstraction | |- ?x === ?x => reflexivity @@ -703,7 +703,7 @@ Section KeyOrderedType. Proof using . intros; red; intros. destruct H1 as [e' H2]. - elim (@ltk_not_eqk (k,e) (k,e')). + contradiction (@ltk_not_eqk (k,e) (k,e')). eapply Sort_Inf_In; eauto. red; simpl; auto. Qed. diff --git a/Util/Fin.v b/Util/Fin.v new file mode 100644 index 0000000000000000000000000000000000000000..f73465d01bdcfa392d2bfe49bfafadb8ffbecd68 --- /dev/null +++ b/Util/Fin.v @@ -0,0 +1,1734 @@ +Require Import Utf8 Arith RelationPairs SetoidDec. +Require Import Lia. +Require Import Pactole.Util.Preliminary. +Require Import Pactole.Util.SetoidDefs. +Require Import Pactole.Util.NumberComplements. +Require Import Pactole.Util.Bijection. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(* + * + * The type fin u with u any nat ≠0 + * consisting of the nats smaller than u + * + *) + +Section Fin. + +Context (u : nat). + +Inductive fin : Type := Fin : ∀ j : nat, j < u -> fin. + +Global Instance fin_Setoid : Setoid fin := eq_setoid fin. + +Lemma fin_dec : ∀ f1 f2 : fin, {f1 = f2} + {f1 <> f2}. +Proof using . + intros [j1 H1] [j2 H2]. destruct (Nat.eq_dec j1 j2). + - subst. left. f_equal. apply le_unique. + - right. intro Habs. inv Habs. auto. +Qed. + +Global Instance fin_EqDec : EqDec fin_Setoid := fin_dec. + +Global Coercion fin2nat (f : fin) : nat := match f with @Fin m _ => m end. + +Definition fin2nat_compat : Proper (equiv ==> equiv) fin2nat := _. + +Lemma fin_lt : ∀ f : fin, f < u. +Proof using . intros. destruct f as [j p]. exact p. Qed. + +Lemma fin_le : ∀ f : fin, f <= Nat.pred u. +Proof using . intros. apply Nat.lt_le_pred, fin_lt. Qed. + +Lemma fin_between : ∀ f : fin, 0 <= f < u. +Proof using . intros. split. apply Nat.le_0_l. apply fin_lt. Qed. + +Lemma fin2natI : Preliminary.injective equiv equiv fin2nat. +Proof using . + intros [j1 H1] [j2 H2] Heq. cbn in Heq. subst. cbn. f_equal. apply le_unique. +Qed. + +Lemma lt_gt_cases : ∀ f1 f2 : fin, f1 ≠f2 <-> f1 < f2 ∨ f2 < f1. +Proof using . + intros. erewrite not_iff_compat. apply Nat.lt_gt_cases. + symmetry. apply (injective_eq_iff fin2natI). +Qed. + +Lemma le_lt_eq_dec : ∀ f1 f2 : fin, f1 <= f2 -> {f1 < f2} + {f1 = f2}. +Proof using . + intros * H. destruct (le_lt_eq_dec f1 f2 H) as [Hd | Hd]. + left. apply Hd. right. apply fin2natI, Hd. +Qed. + +Lemma lt_eq_lt_dec : ∀ f1 f2 : fin, {f1 < f2} + {f1 = f2} + {f2 < f1}. +Proof using . + intros. destruct (lt_eq_lt_dec f1 f2) as [[Hd | Hd] | Hd]. + left. left. apply Hd. left. right. apply fin2natI, Hd. right. apply Hd. +Qed. + +End Fin. + +Section mod2fin. + +Context {l u : nat} {ltc_l_u : l <c u}. + +(* + * + * The fin u whose underlying nat is 0 + * + *) + +Definition fin0 : fin u := Fin lt_0_u. + +Lemma fin02nat : fin0 = 0 :> nat. +Proof using . reflexivity. Qed. + +Lemma all_fin0 : ∀ f : fin u, u = 1 -> f = fin0. +Proof using . + intros f H. apply fin2natI. rewrite fin02nat. subst. + pose proof (fin_lt f) as H. inversion H as [H0 | n H0 H1]. + apply H0. inversion H0. +Qed. + +Lemma all_eq : ∀ f1 f2 : fin u, u = 1 -> f1 = f2. +Proof using ltc_l_u. + intros * H. rewrite (all_fin0 f1), (all_fin0 f2). reflexivity. all: apply H. +Qed. + +Lemma fin0_le : ∀ f : fin u, fin0 <= f. +Proof using . intros. rewrite fin02nat. apply Nat.le_0_l. Qed. + +Lemma neq_fin0_lt_fin0 : ∀ f : fin u, f ≠fin0 <-> fin0 < f. +Proof using . + intros. rewrite <- Nat.neq_0_lt_0, <- fin02nat. symmetry. + apply not_iff_compat, injective_eq_iff, fin2natI. +Qed. + +Lemma eq_fin0_lt_dec : ∀ f : fin u, {f = fin0} + {fin0 < f}. +Proof using . + intros. destruct (eq_0_lt_dec f) as [Hd | Hd]. + left. apply fin2natI, Hd. right. apply Hd. +Qed. + +Lemma le_fin0 : ∀ f : fin u, f <= fin0 <-> f = fin0. +Proof using . + intros. rewrite fin02nat, Nat.le_0_r, <- fin02nat. + apply injective_eq_iff, fin2natI. +Qed. + +Lemma lt_fin0 : ∀ f1 f2 : fin u, f1 < f2 -> fin0 < f2. +Proof using . intros * H. eapply Nat.lt_lt_0, H. Qed. + +(* + * + * The fin u whose underlying nat is Nat.pred u + * + *) + +Definition fin_max : fin u := Fin lt_pred_u. + +Lemma fin_max2nat : fin_max = Nat.pred u :> nat. +Proof using . reflexivity. Qed. + +Lemma le_max : ∀ f : fin u, f <= fin_max. +Proof using . + intros. apply Nat.lt_succ_r. rewrite fin_max2nat, S_pred_u. apply fin_lt. +Qed. + +Lemma max_le : ∀ f : fin u, fin_max <= f <-> f = fin_max. +Proof using . + intros. rewrite Nat.le_lteq, Nat.lt_nge. etransitivity. split. + - intros [H | H]. contradict H. apply le_max. apply H. + - intros H. right. apply H. + - rewrite Nat.eq_sym_iff. apply injective_eq_iff, fin2natI. +Qed. + +Lemma neq_max_lt_max : ∀ f : fin u, f ≠fin_max <-> f < fin_max. +Proof using . intros. rewrite <- max_le. apply Nat.nle_gt. Qed. + +Lemma eq_max_lt_dec : ∀ f : fin u, {f = fin_max} + {f < fin_max}. +Proof using . + intros. destruct (fin_dec f fin_max) as [Hd | Hd]. + left. apply Hd. right. apply neq_max_lt_max, Hd. +Qed. + +Lemma max_lt : ∀ f1 f2 : fin u, f2 < f1 -> f2 < fin_max. +Proof using . + intros * H. rewrite Nat.lt_nge, max_le. intros Ha. + subst. eapply Nat.lt_nge. apply H. apply le_max. +Qed. + +(* + * + * Building the fin u whose underlying nat + * is j mod u with j the input of the function + * + *) + +Lemma mod2fin_lt : ∀ j : nat, j mod u < u. +Proof using ltc_l_u. intros. apply Nat.mod_upper_bound, neq_u_0. Qed. + +Lemma mod2fin_le : ∀ j : nat, j mod u <= Nat.pred u. +Proof using ltc_l_u. intros. apply Nat.lt_le_pred, mod2fin_lt. Qed. + +Definition mod2fin (j : nat) : fin u := Fin (mod2fin_lt j). + +Definition mod2fin_compat : Proper (equiv ==> equiv) mod2fin := _. + +Lemma mod2fin2nat : ∀ j : nat, mod2fin j = j mod u :> nat. +Proof using . intros. reflexivity. Qed. + +Lemma mod2fin_mod : ∀ j : nat, mod2fin (j mod u) = mod2fin j. +Proof using . + intros. apply fin2natI. rewrite 2 mod2fin2nat. apply Nat.Div0.mod_mod. +Qed. + +Lemma mod2fin_mod2fin : ∀ j : nat, mod2fin (mod2fin j) = mod2fin j. +Proof using . intros. rewrite mod2fin2nat. apply mod2fin_mod. Qed. + +Lemma mod2finK : ∀ f : fin u, mod2fin f = f. +Proof using . intros. apply fin2natI, Nat.mod_small, fin_lt. Qed. + +Lemma mod2fin_small : ∀ j : nat, j < u -> mod2fin j = j :> nat. +Proof using . intros * H. rewrite mod2fin2nat. apply Nat.mod_small, H. Qed. + +Lemma mod2fin_bounded_diffI : ∀ j1 j2 : nat, j1 <= j2 + -> j2 - j1 < u -> mod2fin j1 = mod2fin j2 -> j1 = j2. +Proof using . + intros * Hleq Hsub Heq. eapply mod_bounded_diffI. apply neq_u_0. + exact Hleq. exact Hsub. rewrite <-2 mod2fin2nat, Heq. reflexivity. +Qed. + +Lemma mod2fin_betweenI : ∀ p j1 j2 : nat, p <= j1 < p + u + -> p <= j2 < p + u -> mod2fin j1 = mod2fin j2 -> j1 = j2. +Proof using . + intros p. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) mod2fin)). + intros * H1 H2. apply mod2fin_bounded_diffI. all: assumption. +Qed. + +Lemma mod2fin_muln : ∀ j : nat, mod2fin (u * j) = fin0. +Proof using . + intros. apply fin2natI. rewrite mod2fin2nat, fin02nat, Nat.mul_comm. + apply Nat.Div0.mod_mul. +Qed. + +Lemma mod2fin_le_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 <= j2 -> mod2fin j1 <= mod2fin j2. +Proof using . + intros * Hb1 Hb2 Hle. rewrite 2 mod2fin2nat. + eapply mod_le_between_compat. all: eassumption. +Qed. + +Lemma mod2fin_lt_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 < j2 -> mod2fin j1 < mod2fin j2. +Proof using . + intros * Hb1 Hb2 Hlt. rewrite 2 mod2fin2nat. + eapply mod_lt_between_compat. all: eassumption. +Qed. + +Lemma addn_mod2fin_idemp_l : ∀ j1 j2 : nat, + mod2fin (mod2fin j1 + j2) = mod2fin (j1 + j2). +Proof using . + intros. apply fin2natI. rewrite 3 mod2fin2nat. + apply Nat.Div0.add_mod_idemp_l. +Qed. + +Lemma addn_mod2fin_idemp_r : ∀ j1 j2 : nat, + mod2fin (j1 + mod2fin j2) = mod2fin (j1 + j2). +Proof using . + intros. apply fin2natI. rewrite 3 mod2fin2nat. + apply Nat.Div0.add_mod_idemp_r. +Qed. + +Lemma divide_fin0_mod2fin : ∀ j : nat, Nat.divide u j -> mod2fin j = fin0. +Proof using . + intros * H. apply fin2natI. rewrite mod2fin2nat. + apply Nat.Lcm0.mod_divide, H. +Qed. + +Lemma mod2fin0 : mod2fin 0 = fin0. +Proof using . apply divide_fin0_mod2fin, Nat.divide_0_r. Qed. + +Lemma mod2fin_max : mod2fin (Nat.pred u) = fin_max. +Proof using . apply fin2natI, mod2fin_small, fin_lt. Qed. + +Lemma mod2fin_u : mod2fin u = fin0. +Proof using . apply divide_fin0_mod2fin, Nat.divide_refl. Qed. + +(* + * + * The successor of fin0 + * + *) + +Definition fin1 : fin u := mod2fin 1. + +Lemma fin1E : fin1 = mod2fin 1. +Proof using . reflexivity. Qed. + +Lemma fin12nat {ltc_1_u : 1 <c u} : fin1 = 1 :> nat. +Proof using . rewrite fin1E. apply mod2fin_small, ltc_1_u. Qed. + +(* + * + * Adding either a nat (addm) or + * a fin u (addf) to a fin u + * + *) + +Definition addm (f : fin u) (j : nat) : fin u := mod2fin (f + j). + +Definition addm_compat : Proper (equiv ==> equiv ==> equiv) addm := _. + +Lemma addmE : ∀ (f : fin u) (j : nat), addm f j = mod2fin (f + j). +Proof using . reflexivity. Qed. + +Lemma addm2nat : ∀ (f : fin u) (j : nat), addm f j = (f + j) mod u :> nat. +Proof using . intros. rewrite addmE. apply mod2fin2nat. Qed. + +Lemma addm_mod : ∀ (f : fin u) (j : nat), addm f (j mod u) = addm f j. +Proof using . + intros. rewrite 2 addmE, <- mod2fin2nat. apply addn_mod2fin_idemp_r. +Qed. + +Lemma addm_mod2fin : ∀ (f : fin u) (j : nat), addm f (mod2fin j) = addm f j. +Proof using . intros. rewrite mod2fin2nat. apply addm_mod. Qed. + +Definition addf (f1 f2 : fin u) : fin u := addm f1 f2. + +Definition addf_compat : Proper (equiv ==> equiv ==> equiv) addf := _. + +Lemma addfE : ∀ f1 f2 : fin u, addf f1 f2 = mod2fin (f1 + f2). +Proof using . reflexivity. Qed. + +Lemma addf2nat : ∀ f1 f2 : fin u, addf f1 f2 = (f1 + f2) mod u :> nat. +Proof using . intros. apply addm2nat. Qed. + +Lemma addf_addm : ∀ f1 f2 : fin u, addf f1 f2 = addm f1 f2. +Proof using . reflexivity. Qed. + +Lemma addm_addf : ∀ (f : fin u) (j : nat), addm f j = addf f (mod2fin j). +Proof using . intros. rewrite addf_addm. symmetry. apply addm_mod2fin. Qed. + +Lemma addfC : ∀ f1 f2 : fin u, addf f1 f2 = addf f2 f1. +Proof using . intros. rewrite 2 addfE, Nat.add_comm. reflexivity. Qed. + +Lemma addmC : ∀ (f : fin u) (j : nat), addm f j = addm (mod2fin j) f. +Proof using . intros. rewrite 2 addm_addf, mod2finK. apply addfC. Qed. + +Lemma addmA : ∀ (f1 f2 : fin u) (j : nat), + addm f1 (addm f2 j) = addm (addm f1 f2) j. +Proof using . + intros. rewrite 4 addmE, addn_mod2fin_idemp_l, + addn_mod2fin_idemp_r, Nat.add_assoc. reflexivity. +Qed. + +Lemma addfA : ∀ f1 f2 f3 : fin u, + addf f1 (addf f2 f3) = addf (addf f1 f2) f3. +Proof using . intros. rewrite 2 (addf_addm _ f3). apply addmA. Qed. + +Lemma addmAC : ∀ (f : fin u) (j1 j2 : nat), + addm (addm f j1) j2 = addm (addm f j2) j1. +Proof using . + intros. rewrite 4 addm_addf, <-2 addfA, (addfC (mod2fin _)). reflexivity. +Qed. + +Lemma addfAC : ∀ f1 f2 f3 : fin u, + addf (addf f1 f2) f3 = addf (addf f1 f3) f2. +Proof using . intros. rewrite 4 addf_addm. apply addmAC. Qed. + +Lemma addfCA : ∀ f1 f2 f3 : fin u, + addf f1 (addf f2 f3) = addf f2 (addf f1 f3). +Proof using . + intros. rewrite (addfC f2 f3), addfA, (addfC (addf _ _) _). reflexivity. +Qed. + +Lemma addmCA : ∀ (f1 f2 : fin u) (j : nat), + addm f1 (addm f2 j) = addm f2 (addm f1 j). +Proof using . + intros. rewrite 2 (addm_addf _ j), <-2 addf_addm. apply addfCA. +Qed. + +Lemma addfACA : ∀ f1 f2 f3 f4 : fin u, + addf (addf f1 f2) (addf f3 f4) = addf (addf f1 f3) (addf f2 f4). +Proof using . intros. rewrite 2 addfA, (addfAC f1 f3 f2). reflexivity. Qed. + +Lemma addmACA : ∀ (f1 f2 f3 : fin u) (j : nat), + addm (addm f1 f2) (addm f3 j) = addm (addm f1 f3) (addm f2 j). +Proof using . + intros. rewrite 2 (addm_addf _ j), <- 4 addf_addm. apply addfACA. +Qed. + +Lemma addIm : ∀ j : nat, + Preliminary.injective equiv equiv (λ f : fin u, addm f j). +Proof using . + intros * f1 f2 H. eapply fin2natI, (Nat.add_cancel_r _ _ j), mod2fin_betweenI. + 1,2: rewrite Nat.add_comm. 1,2: split. 1,3: erewrite <- Nat.add_le_mono_l. + 1,2: apply Nat.le_0_l. 1,2: rewrite Nat.add_0_r, <- Nat.add_lt_mono_l. + 1,2: apply fin_lt. apply H. +Qed. + +Lemma addm_bounded_diffI : + ∀ (f : fin u) (j1 j2 : nat), j1 <= j2 -> j2 - j1 < u + -> addm f j1 = addm f j2 -> j1 = j2. +Proof using . + intros * Hle Hsu H. rewrite 2 addmE in H. + eapply Nat.add_cancel_l, mod2fin_bounded_diffI; try eassumption; lia. +Qed. + +Lemma addm_betweenI : ∀ (p : nat) (f : fin u) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> addm f j1 = addm f j2 -> j1 = j2. +Proof using . + intros p f. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) (λ x, addm f x))). + intros *. apply addm_bounded_diffI. +Qed. + +Lemma addIf : ∀ f1 : fin u, + Preliminary.injective equiv equiv (λ f2, addf f2 f1). +Proof using . intros f1 f2 f3. rewrite 2 addf_addm. apply addIm. Qed. + +Lemma addfI : ∀ f1 : fin u, Preliminary.injective equiv equiv (addf f1). +Proof using . + intros f1 f2 f3 H. eapply addIf. setoid_rewrite addfC. apply H. +Qed. + +Lemma addm0 : ∀ f : fin u, addm f 0 = f. +Proof using . intros. rewrite addmE, Nat.add_0_r. apply mod2finK. Qed. + +Lemma add0m : ∀ j : nat, addm fin0 j = mod2fin j. +Proof using . intros. rewrite addmE, Nat.add_0_l. reflexivity. Qed. + +Lemma add0f : ∀ f : fin u, addf fin0 f = f. +Proof using . intros. rewrite addf_addm, add0m. apply mod2finK. Qed. + +Lemma addf0 : ∀ f : fin u, addf f fin0 = f. +Proof using . intros. rewrite addf_addm. apply addm0. Qed. + +(* + * + * The successor of a fin u + * + *) + +Definition sucf (f : fin u) : fin u := addf f fin1. + +Definition sucf_compat : Proper (equiv ==> equiv) sucf := _. + +Lemma sucfE : ∀ f : fin u, sucf f = addf f fin1. +Proof using . reflexivity. Qed. + +Lemma sucfEmod : ∀ f : fin u, sucf f = mod2fin (S f). +Proof using . + intros. destruct (@eq_S_l_lt_dec 0 u _) as [Hd | Hd]. apply all_eq, Hd. + unshelve erewrite sucfE, addfE, fin12nat, Nat.add_1_r. apply Hd. reflexivity. +Qed. + +Lemma sucfI : Preliminary.injective equiv equiv sucf. +Proof using . apply addIf. Qed. + +Lemma mod2fin_S_sucf : ∀ j : nat, mod2fin (S j) = sucf (mod2fin j). +Proof using . + intros. rewrite sucfE, addfE, fin1E, addn_mod2fin_idemp_l, + addn_mod2fin_idemp_r, Nat.add_1_r. reflexivity. +Qed. + +Lemma sucf_max : sucf fin_max = fin0. +Proof using . rewrite sucfEmod, fin_max2nat, S_pred_u. apply mod2fin_u. Qed. + +Lemma sucf_addf : ∀ f1 f2 : fin u, sucf (addf f1 f2) = addf (sucf f1) f2. +Proof using . intros. rewrite 2 sucfE. apply addfAC. Qed. + +Lemma sucf_addm : ∀ (f : fin u) (j : nat), sucf (addm f j) = addm (sucf f) j. +Proof using . intros. rewrite 2 addm_addf. apply sucf_addf. Qed. + +Lemma addf_sucf : ∀ f1 f2 : fin u, addf (sucf f1) f2 = addf f1 (sucf f2). +Proof using . + intros. rewrite (addfC f1), <-2 sucf_addf, addfC. reflexivity. +Qed. + +Lemma addf_mod2fin_S : ∀ j1 j2 : nat, + addf (mod2fin (S j1)) (mod2fin j2) = addf (mod2fin j1) (mod2fin (S j2)). +Proof using . + intros. rewrite mod2fin_S_sucf, addf_sucf, mod2fin_S_sucf. reflexivity. +Qed. + +Lemma S_sucf : ∀ f : fin u, f < fin_max -> S f = sucf f. +Proof using . + intros * H. symmetry. rewrite sucfEmod. + apply mod2fin_small, Nat.lt_succ_lt_pred, H. +Qed. + +Lemma lt_sucf : ∀ f : fin u, f < fin_max -> f < sucf f. +Proof using . + intros * H. rewrite <- S_sucf. apply Nat.lt_succ_diag_r. apply H. +Qed. + +Lemma fin1_sucf_fin0 : fin1 = sucf fin0. +Proof using . rewrite sucfEmod, fin02nat. apply fin1E. Qed. + +Lemma lt_sucf_le : ∀ f1 f2 : fin u, f1 < sucf f2 -> f1 <= f2. +Proof using . + intros * H1. apply Nat.lt_succ_r. rewrite S_sucf. + apply H1. apply neq_max_lt_max. intros H2. subst. + eapply Nat.lt_irrefl, lt_fin0. rewrite <- sucf_max. apply H1. +Qed. + +(* + * + * The complement to fin_max of either j mod u + * (revm whose input is a nat j) or of a fin u (revf) + * + *) + +Lemma revf_subproof : ∀ f : fin u, Nat.pred u - f < u. +Proof using ltc_l_u. + intros. eapply Nat.le_lt_trans. apply Nat.le_sub_l. apply lt_pred_u. +Qed. + +Definition revf (f : fin u) : fin u := Fin (revf_subproof f). + +Definition revf_compat : Proper (equiv ==> equiv) revf := _. + +Lemma revf2nat : ∀ f : fin u, revf f = Nat.pred u - f :> nat. +Proof using . reflexivity. Qed. + +Lemma revfK : ∀ f : fin u, revf (revf f) = f. +Proof using . +intros. apply fin2natI. rewrite 2 revf2nat, sub_sub. +- hnf. lia. +- apply fin_le. +Qed. + +Lemma revfI : Preliminary.injective equiv equiv revf. +Proof using . + intros f1 f2 H. setoid_rewrite <- revfK. rewrite H. reflexivity. +Qed. + +Definition revm (j : nat) : fin u := revf (mod2fin j). + +Definition revm_compat : Proper (equiv ==> equiv) revm := _. + +Lemma revm_revf : ∀ j : nat, revm j = revf (mod2fin j). +Proof using . reflexivity. Qed. + +Lemma revf_revm : ∀ f : fin u, revf f = revm f. +Proof using . intros. rewrite revm_revf, mod2finK. reflexivity. Qed. + +Lemma revm2nat : ∀ j : nat, revm j = Nat.pred u - (j mod u) :> nat. +Proof using . + intros. rewrite revm_revf, revf2nat, mod2fin2nat. reflexivity. +Qed. + +Lemma revm_mod : ∀ j : nat, revm (j mod u) = revm j. +Proof using . intros. rewrite 2 revm_revf, mod2fin_mod. reflexivity. Qed. + +Lemma revm_mod2fin : ∀ j : nat, revm (mod2fin j) = revm j. +Proof using . intros. rewrite mod2fin2nat. apply revm_mod. Qed. + +Lemma revmK : ∀ j : nat, revm (revm j) = mod2fin j. +Proof using . intros. rewrite 2 revm_revf, mod2finK. apply revfK. Qed. + +Lemma revm_bounded_diffI : + ∀ j1 j2 : nat, j1 <= j2 -> j2 - j1 < u -> revm j1 = revm j2 -> j1 = j2. +Proof using . + intros. eapply mod_bounded_diffI, sub_cancel_l; + try apply mod2fin_lt || apply neq_u_0; trivial; []. + rewrite <- (Nat.lt_succ_pred l u) at 1 3. + rewrite <- Nat.add_1_r, 2 Nat.add_sub_swap, <-2 revm2nat, H1; trivial; + apply Nat.lt_le_pred, mod2fin_lt. apply lt_l_u. +Qed. + +Lemma revm_betweenI : ∀ p j1 j2 : nat, + p <= j1 < p + u -> p <= j2 < p + u -> revm j1 = revm j2 -> j1 = j2. +Proof using . + rewrite <- bounded_diff_between. apply revm_bounded_diffI. + apply (@inj_sym _ _ (eq_setoid _) (eq_setoid _)). +Qed. + +Lemma revm0 : revm 0 = fin_max. +Proof using . + apply fin2natI. rewrite revm2nat, fin_max2nat, Nat.Div0.mod_0_l. + apply Nat.sub_0_r. +Qed. + +Lemma revf0 : revf fin0 = fin_max. +Proof using . rewrite revf_revm, fin02nat. apply revm0. Qed. + +Lemma revf_le_compat : ∀ f1 f2 : fin u, f1 <= f2 -> revf f2 <= revf f1. +Proof using . intros * H. rewrite 2 revf2nat. apply Nat.sub_le_mono_l, H. Qed. + +Lemma revf_lt_compat : ∀ f1 f2 : fin u, f1 < f2 -> revf f2 < revf f1. +Proof using . + intros * H. rewrite 2 revf2nat. apply sub_lt_mono_l. apply fin_le. apply H. +Qed. + +Lemma revm_le_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 <= j2 -> revm j2 <= revm j1. +Proof using . + intros * Hb1 Hb2 Hle. rewrite 2 revm2nat. eapply Nat.sub_le_mono_l, + mod_le_between_compat. all: eassumption. +Qed. + +Lemma revm_lt_between_compat : ∀ p j1 j2 : nat, u * p <= j1 < u * p + u + -> u * p <= j2 < u * p + u -> j1 < j2 -> revm j2 < revm j1. +Proof using . + intros * Hb1 Hb2 Hle. rewrite 2 revm2nat. apply sub_lt_mono_l. + apply mod2fin_le. eapply mod_lt_between_compat. all: eassumption. +Qed. + +Lemma revf_fin_max : revf fin_max = fin0. +Proof using . symmetry. apply revfI. rewrite revfK. apply revf0. Qed. + +Lemma revm_fin_max : revm fin_max = fin0. +Proof using . rewrite <- revf_revm. apply revf_fin_max. Qed. + +Lemma addf_revf : ∀ f : fin u, addf (revf f) f = fin_max. +Proof using . + intros. rewrite addfE, revf2nat, Nat.sub_add. apply mod2fin_max. apply fin_le. +Qed. + +(* + * + * The complement to fin0 of either j mod u + * (oppm whose input is a nat j) or of a fin u (oppf) + * + *) + +Definition oppf (f : fin u) : fin u := sucf (revf f). + +Definition oppf_compat : Proper (equiv ==> equiv) oppf := _. + +Lemma oppfE : ∀ f : fin u, oppf f = sucf (revf f). +Proof using . reflexivity. Qed. + +Lemma oppfEmod : ∀ f : fin u, oppf f = mod2fin (u - f). +Proof using . + intros. rewrite oppfE, sucfEmod, revf2nat, <- Nat.sub_succ_l, S_pred_u; + auto using fin_le. +Qed. + +Lemma oppf2nat : ∀ f : fin u, oppf f = (u - f) mod u :> nat. +Proof using . intros. rewrite oppfEmod. apply mod2fin2nat. Qed. + +Definition oppm (j : nat) : fin u := oppf (mod2fin j). + +Definition oppm_compat : Proper (equiv ==> equiv) oppm := _. + +Lemma oppm_oppf : ∀ j : nat, oppm j = oppf (mod2fin j). +Proof using . reflexivity. Qed. + +Lemma oppf_oppm : ∀ f : fin u, oppf f = oppm f. +Proof using . intros. rewrite oppm_oppf, mod2finK. reflexivity. Qed. + +Lemma oppmE : ∀ j : nat, oppm j = sucf (revm j). +Proof using . intros. rewrite oppm_oppf, oppfE, revm_revf. reflexivity. Qed. + +Lemma oppmEmod : ∀ j : nat, oppm j = mod2fin (u - (j mod u)). +Proof using . + intros. rewrite oppm_oppf, oppfEmod, mod2fin2nat. reflexivity. +Qed. + +Lemma oppm2nat : ∀ j : nat, oppm j = (u - (j mod u)) mod u :> nat. +Proof using . intros. rewrite oppmEmod. apply mod2fin2nat. Qed. + +Lemma oppm_mod : ∀ j : nat, oppm (j mod u) = oppm j. +Proof using . intros. rewrite 2 oppm_oppf, mod2fin_mod. reflexivity. Qed. + +Lemma addfKoppf : ∀ f : fin u, addf (oppf f) f = fin0. +Proof using . + intros. rewrite oppfE, <- sucf_addf, addf_revf. apply sucf_max. +Qed. + +Lemma addmKoppm : ∀ j : nat, addm (oppm j) j = fin0. +Proof using . intros. rewrite oppm_oppf, addm_addf. apply addfKoppf. Qed. + +Lemma oppfKaddf : ∀ f : fin u, addf f (oppf f) = fin0. +Proof using . intros. rewrite addfC. apply addfKoppf. Qed. + +Lemma addfOoppf : ∀ f1 f2 : fin u, addf (oppf f1) (addf f1 f2) = f2. +Proof using . intros. rewrite addfA, addfKoppf, add0f. reflexivity. Qed. + +Lemma oppfOaddf : ∀ f1 f2 : fin u, addf (addf f2 f1) (oppf f1) = f2. +Proof using . intros. rewrite addfC, (addfC f2 f1). apply addfOoppf. Qed. + +Lemma addfOVoppf : ∀ f1 f2 : fin u, addf f1 (addf (oppf f1) f2) = f2. +Proof using . intros. rewrite (addfC _ f2), addfCA, oppfKaddf. apply addf0. Qed. + +Lemma oppfOVaddf : ∀ f1 f2 : fin u, addf (addf f2 (oppf f1)) f1 = f2. +Proof using . intros. rewrite addfC, (addfC f2 _). apply addfOVoppf. Qed. + +Lemma oppmOVaddm : ∀ (f : fin u) (j : nat), addm (addf f (oppm j)) j = f. +Proof using . intros. rewrite oppm_oppf, addm_addf. apply oppfOVaddf. Qed. + +Lemma oppfI : Preliminary.injective equiv equiv oppf. +Proof using . intros f1 f2 H. eapply revfI, sucfI, H. Qed. + +Lemma oppm_bounded_diffI : ∀ j1 j2 : nat, j1 <= j2 + -> j2 - j1 < u -> oppm j1 = oppm j2 -> j1 = j2. +Proof using . + intros * Hleq Hsub H. rewrite 2 oppmE in H. + eapply revm_bounded_diffI, sucfI. all: eassumption. +Qed. + +Lemma oppm_betweenI : ∀ (p : nat) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> oppm j1 = oppm j2 -> j1 = j2. +Proof using . + rewrite <- bounded_diff_between. apply oppm_bounded_diffI. + apply (@inj_sym _ _ (eq_setoid _) (eq_setoid _)). +Qed. + +Lemma oppf0 : oppf fin0 = fin0. +Proof using . rewrite oppfE, revf0. apply sucf_max. Qed. + +Lemma oppm0 : oppm 0 = fin0. +Proof using . rewrite oppm_oppf, mod2fin0. apply oppf0. Qed. + +Lemma addn_oppf : ∀ f : fin u, fin0 < f -> f + oppf f = u. +Proof using . +intros * H. rewrite oppf2nat, Nat.mod_small, Nat.add_comm, Nat.sub_add. +- reflexivity. +- apply Nat.lt_le_incl, fin_lt. +- apply lt_sub_u, H. +Qed. + +Lemma divide_addn_oppf : ∀ f : fin u, Nat.divide u (f + oppf f). +Proof using . + intros. destruct (eq_fin0_lt_dec f) as [Hd | Hd]. subst. rewrite oppf0, + Nat.add_0_r. apply Nat.divide_0_r. rewrite addn_oppf. reflexivity. apply Hd. +Qed. + +Lemma divide_addn_oppm : ∀ j : nat, Nat.divide u (j + oppm j). +Proof using . + intros. rewrite (Nat.div_mod j) at 1. rewrite <- Nat.add_assoc, + <- mod2fin2nat, oppm_oppf. apply Nat.divide_add_r. apply Nat.divide_factor_l. + apply divide_addn_oppf. apply neq_u_0. +Qed. + +(* + * + * Substracting either a nat (subm) or + * a fin u (subf) to a fin u + * + *) + +Definition subm (f : fin u) (j : nat) : fin u := addm f (oppm j). + +Definition subm_compat : Proper (equiv ==> equiv ==> equiv) subm := _. + +Definition subf (f1 f2 : fin u) : fin u := addf f1 (oppf f2). + +Definition subf_compat : Proper (equiv ==> equiv ==> equiv) subf := _. + +Lemma subfE : ∀ f1 f2 : fin u, subf f1 f2 = addf f1 (oppf f2). +Proof using . reflexivity. Qed. + +Lemma submE : ∀ (f : fin u) (j : nat), subm f j = addf f (oppm j). +Proof using . reflexivity. Qed. + +Lemma subf_subm : ∀ f1 f2 : fin u, subf f1 f2 = subm f1 f2. +Proof using . intros. rewrite submE, subfE, oppf_oppm. reflexivity. Qed. + +Lemma subm_subf : ∀ (f : fin u) (j : nat), subm f j = subf f (mod2fin j). +Proof using . intros. rewrite submE, subfE, oppm_oppf. reflexivity. Qed. + +Lemma submEmod : ∀ (f : fin u) (j : nat), + subm f j = mod2fin (f + (u - (j mod u))). +Proof using . + intros. rewrite submE, addfE, oppm2nat, + <- mod2fin2nat. apply addn_mod2fin_idemp_r. +Qed. + +Lemma subm2nat : ∀ (f : fin u) (j : nat), + subm f j = (f + (u - (j mod u))) mod u :> nat. +Proof using . intros. rewrite submEmod. apply mod2fin2nat. Qed. + +Lemma subfEmod : ∀ f1 f2 : fin u, subf f1 f2 = mod2fin (f1 + (u - f2)). +Proof using . + intros. rewrite subf_subm, submEmod, Nat.mod_small. reflexivity. apply fin_lt. +Qed. + +Lemma subf2nat : ∀ f1 f2 : fin u, + subf f1 f2 = (f1 + (u - f2)) mod u :> nat. +Proof using . intros. rewrite subfEmod. apply mod2fin2nat. Qed. + +Lemma subm_mod : ∀ (f : fin u) (j : nat), subm f (j mod u) = subm f j. +Proof using . intros. rewrite 2 subm_subf, mod2fin_mod. reflexivity. Qed. + +Lemma subm_mod2fin : ∀ (f : fin u) (j : nat), subm f (mod2fin j) = subm f j. +Proof using . intros. rewrite mod2fin2nat. apply subm_mod. Qed. + +Lemma subIf : ∀ f1 : fin u, + Preliminary.injective equiv equiv (λ f2, subf f2 f1). +Proof using . intros f1 f2 f3 H. eapply addIf, H. Qed. + +Lemma subfI : ∀ f1 : fin u, + Preliminary.injective equiv equiv (subf f1). +Proof using . intros f1 f2 f3 H. eapply oppfI, addfI, H. Qed. + +Lemma subIm : ∀ j : nat, + Preliminary.injective equiv equiv (λ f, subm f j). +Proof using . intros j f1 f2. rewrite 2 subm_subf. apply subIf. Qed. + +Lemma subm_bounded_diffI : ∀ (f : fin u) (j1 j2 : nat), + j1 <= j2 -> j2 - j1 < u -> subm f j1 = subm f j2 -> j1 = j2. +Proof using . + intros * Hle Hsu H. rewrite 2 submE in H. + eapply oppm_bounded_diffI, addfI. all: eassumption. +Qed. + +Lemma subm_betweenI : ∀ (p : nat) (f : fin u) (j1 j2 : nat), + p <= j1 < p + u -> p <= j2 < p + u -> subm f j1 = subm f j2 -> j1 = j2. +Proof using . + intros p f. eapply (bounded_diff_between _ + (inj_sym _ _ nat_Setoid (eq_setoid (fin u)) (λ x, subm f x))). + intros *. apply subm_bounded_diffI. +Qed. + +Lemma sub0f : ∀ f : fin u, subf fin0 f = oppf f. +Proof using . intros. rewrite subfE. apply add0f. Qed. + +Lemma sub0m : ∀ j : nat, subm fin0 j = oppm j. +Proof using . intros. rewrite submE. apply add0f. Qed. + +Lemma subf0 : ∀ f : fin u, subf f fin0 = f. +Proof using . intros. rewrite subfE, oppf0. apply addf0. Qed. + +Lemma subm0 : ∀ f : fin u, subm f 0 = f. +Proof using . intros. rewrite subm_subf, mod2fin0. apply subf0. Qed. + +Lemma subff : ∀ f : fin u, subf f f = fin0. +Proof using . intros. rewrite subfE. apply oppfKaddf. Qed. + +Lemma submm : ∀ j : nat, subm (mod2fin j) (mod2fin j) = fin0. +Proof using . intros. rewrite subm_subf, mod2finK. apply subff. Qed. + +Lemma subfAC : ∀ f1 f2 f3 : fin u, + subf (subf f1 f2) f3 = subf (subf f1 f3) f2. +Proof using . intros. rewrite 4 subfE. apply addfAC. Qed. + +Lemma submAC : ∀ (f : fin u) (j1 j2 : nat), + subm (subm f j1) j2 = subm (subm f j2) j1. +Proof using . intros. rewrite 4 submE. apply addmAC. Qed. + +Lemma addf_subf : ∀ f1 f2 f3 : fin u, + addf (subf f1 f2) f3 = subf (addf f1 f3) f2. +Proof using . intros. rewrite 2 subfE. apply addfAC. Qed. + +Lemma addm_subm : ∀ (f : fin u) (j1 j2 : nat), + addm (subm f j2) j1 = subm (addm f j1) j2. +Proof using . intros. rewrite 2 addm_addf, 2 subm_subf. apply addf_subf. Qed. + +Lemma addfKV : ∀ f1 f2 : fin u, subf (addf f1 f2) f1 = f2. +Proof using . intros. rewrite addfC, subfE. apply oppfOaddf. Qed. + +Lemma addfVKV : ∀ f1 f2 : fin u, subf (addf f2 f1) f1 = f2. +Proof using . intros. rewrite addfC. apply addfKV. Qed. + +Lemma addmVKV : ∀ (j : nat) (f : fin u), subm (addm f j) j = f. +Proof using . intros. rewrite subm_subf, addm_addf. apply addfVKV. Qed. + +Lemma subfVK : ∀ f1 f2 : fin u, addf f1 (subf f2 f1) = f2. +Proof using . intros. eapply subIf. rewrite addfKV. reflexivity. Qed. + +Lemma subfVKV : ∀ f1 f2 : fin u, addf (subf f2 f1) f1 = f2. +Proof using . intros. rewrite addfC. apply subfVK. Qed. + +Lemma submVKV : ∀ (j : nat) (f : fin u), addm (subm f j) j = f. +Proof using . intros. rewrite addm_addf, subm_subf. apply subfVKV. Qed. + +(* + * + * The predecessor of a fin u + * + *) + +Definition pref (f : fin u) : fin u := subf f fin1. + +Definition pref_compat : Proper (equiv ==> equiv) pref := _. + +Lemma prefE : ∀ f : fin u, pref f = subf f fin1. +Proof using . reflexivity. Qed. + +Lemma prefI : Preliminary.injective equiv equiv pref. +Proof using . apply subIf. Qed. + +Lemma sucfK : ∀ f : fin u, pref (sucf f) = f. +Proof using . intros. rewrite prefE, sucfE. apply addfVKV. Qed. + +Lemma prefK : ∀ f : fin u, sucf (pref f) = f. +Proof using . intros. rewrite sucfE, prefE. apply subfVKV. Qed. + +Lemma prefEmod : ∀ f : fin u, pref f = mod2fin (f + Nat.pred u). +Proof using . + intros. apply sucfI. symmetry. rewrite prefK, <- mod2fin_S_sucf, plus_n_Sm, + S_pred_u, <- addn_mod2fin_idemp_r, mod2fin_u, Nat.add_0_r. apply mod2finK. +Qed. + +Lemma revfE : ∀ f : fin u, revf f = pref (oppf f). +Proof using . intros. rewrite oppfE. symmetry. apply sucfK. Qed. + +Lemma pref0 : pref fin0 = fin_max. +Proof using . rewrite <- sucf_max. apply sucfK. Qed. + +Lemma pred_pref : ∀ f : fin u, fin0 < f -> Nat.pred f = pref f. +Proof using . + intros * H. symmetry. apply Nat.succ_inj. erewrite Nat.lt_succ_pred, S_sucf. + apply fin2nat_compat, prefK. rewrite <- neq_max_lt_max, <- pref0. + eapply not_iff_compat. 2: apply neq_fin0_lt_fin0. 2,3: apply H. + split. apply prefI. apply pref_compat. +Qed. + +Lemma pref_subf : ∀ f1 f2 : fin u, pref (subf f1 f2) = subf (pref f1) f2. +Proof using . intros. rewrite 2 prefE. apply subfAC. Qed. + +Lemma pref_subm : ∀ (f : fin u) (j : nat), pref (subm f j) = subm (pref f) j. +Proof using . intros. rewrite 2 subm_subf. apply pref_subf. Qed. + +Lemma sucf_subf : ∀ f1 f2 : fin u, sucf (subf f1 f2) = subf (sucf f1) f2. +Proof using . intros. rewrite 2 sucfE. apply addf_subf. Qed. + +Lemma sucf_subm : ∀ (f : fin u) (j : nat), sucf (subm f j) = subm (sucf f) j. +Proof using . intros. rewrite 2 subm_subf. apply sucf_subf. Qed. + +Lemma pref_addf : ∀ f1 f2 : fin u, pref (addf f1 f2) = addf (pref f1) f2. +Proof using . intros. rewrite 2 prefE. symmetry. apply addf_subf. Qed. + +Lemma pref_addm : ∀ (f : fin u) (j : nat), pref (addm f j) = addm (pref f) j. +Proof using . intros. rewrite 2 addm_addf. apply pref_addf. Qed. + +Lemma addf_pref : ∀ f1 f2 : fin u, addf (pref f1) f2 = addf f1 (pref f2). +Proof using . + intros. rewrite (addfC f1), <- 2 pref_addf, addfC. reflexivity. +Qed. + +Lemma lt_pref : ∀ f : fin u, fin0 < f -> pref f < f. +Proof using . + intros * H. rewrite <- pred_pref. eapply lt_pred. all: apply H. +Qed. + +Lemma fin0_pref_fin1 : fin0 = pref fin1. +Proof using. symmetry. apply sucfI. rewrite prefK. apply fin1_sucf_fin0. Qed. + +Lemma lt_pref_le : ∀ f1 f2 : fin u, pref f1 < f2 -> f1 <= f2. +Proof using . + intros * H1. apply Nat.lt_pred_le. rewrite pred_pref. + apply H1. apply neq_fin0_lt_fin0. intros H2. subst. + eapply Nat.lt_irrefl, max_lt. rewrite <- pref0. apply H1. +Qed. + +(* + * + * The complementaries (either to fin_max or fin0) + * or the output of several other functions + * (then used to prove other lemmas) + * + *) + +Lemma revf_addf : ∀ f1 f2 : fin u, revf (addf f1 f2) = subf (revf f1) f2. +Proof using . + intros. apply fin2natI. erewrite subf2nat, 2 revf2nat, addf2nat, + Nat.add_sub_assoc, <- Nat.add_sub_swap, <- (Nat.mod_small (Nat.pred _ - _)), + <- Nat.sub_add_distr, <- (Nat.Div0.mod_add (_ - ((_ + _) mod _)) 1), Nat.mul_1_l, + <- Nat.add_sub_swap, (Nat.Div0.mod_eq (_ + _)), sub_sub, + (Nat.add_sub_swap _ (_ * _)), Nat.mul_comm, Nat.Div0.mod_add. reflexivity. + - apply Nat.add_le_mono. + + apply Nat.lt_le_pred. + apply fin_lt. + + apply Nat.lt_le_incl. + apply fin_lt. + - apply Nat.Div0.mul_div_le. + - apply Nat.lt_le_pred. + apply mod2fin_lt. + - apply lt_sub_lt_add_l. + + apply Nat.lt_le_pred. + apply mod2fin_lt. + + apply Nat.lt_lt_add_l. + apply lt_pred_u. + - apply Nat.lt_le_pred. + apply fin_lt. + - apply Nat.lt_le_incl. + apply fin_lt. +Qed. + +Lemma revf_addm : ∀ (f : fin u) (j : nat), revf (addm f j) = subm (revf f) j. +Proof using . intros. rewrite addm_addf. apply revf_addf. Qed. + +Lemma revf_sucf : ∀ f : fin u, revf (sucf f) = pref (revf f). +Proof using . intros. rewrite sucfE, prefE. apply revf_addf. Qed. + +Lemma oppf_addf : ∀ f1 f2 : fin u, oppf (addf f1 f2) = subf (oppf f1) f2. +Proof using . intros. rewrite 2 oppfE, revf_addf. apply sucf_subf. Qed. + +Lemma oppf_addm : ∀ (f : fin u) (j : nat), oppf (addm f j) = subm (oppf f) j. +Proof using . intros. rewrite addm_addf, subm_subf. apply oppf_addf. Qed. + +Lemma oppf_sucf : ∀ f : fin u, oppf (sucf f) = pref (oppf f). +Proof using . intros. rewrite sucfE, prefE. apply oppf_addf. Qed. + +Lemma revf_subf : ∀ f1 f2 : fin u, revf (subf f1 f2) = addf (revf f1) f2. +Proof using . + intros. eapply subIf. rewrite <- revf_addf, addfVKV, subfVKV. reflexivity. +Qed. + +Lemma revf_subm : ∀ (f : fin u) (j : nat), revf (subm f j) = addm (revf f) j. +Proof using . intros. rewrite subm_subf, addm_addf. apply revf_subf. Qed. + +Lemma revf_pref : ∀ f : fin u, revf (pref f) = sucf (revf f). +Proof using . intros. rewrite sucfE, prefE. apply revf_subf. Qed. + +Lemma oppf_subf : ∀ f1 f2 : fin u, oppf (subf f2 f1) = subf f1 f2. +Proof using . + intros. rewrite oppfE, revf_subf, sucf_addf, + <- oppfE, addfC. symmetry. apply subfE. +Qed. + +Lemma oppf_subm : ∀ (f : fin u) (j : nat), + oppf (subm f j) = subf (mod2fin j) f. +Proof using . intros. rewrite subm_subf. apply oppf_subf. Qed. + +Lemma oppf_pref : ∀ f : fin u, oppf (pref f) = sucf (oppf f). +Proof using . intros. rewrite 2 oppfE, revf_pref. reflexivity. Qed. + +Lemma revf_oppf : ∀ f : fin u, revf (oppf f) = pref f. +Proof using . intros. rewrite oppfE, revf_sucf, revfK. reflexivity. Qed. + +Lemma oppfK : ∀ f : fin u, oppf (oppf f) = f. +Proof using . intros. rewrite oppfE, revf_oppf. apply prefK. Qed. + +Lemma oppmK : ∀ j : nat, oppf (oppm j) = mod2fin j. +Proof using . intros. rewrite oppm_oppf. apply oppfK. Qed. + +Lemma oppf_revf : ∀ f : fin u, oppf (revf f) = sucf f. +Proof using . intros. rewrite oppfE, revfK. reflexivity. Qed. + +Lemma addf_subf_oppf : ∀ f1 f2 : fin u, addf f1 f2 = subf f1 (oppf f2). +Proof using . intros. rewrite subfE, oppfK. reflexivity. Qed. + +Lemma addm_subf_oppm : ∀ (f : fin u) (j : nat), addm f j = subf f (oppm j). +Proof using . intros. rewrite subfE, oppmK. apply addm_addf. Qed. + +Lemma subf_subf : ∀ f1 f2 f3 : fin u, + subf f1 (subf f2 f3) = subf (addf f1 f3) f2. +Proof using . + intros. rewrite 3subfE, oppf_addf, subfE, oppfK, addfA. apply addfAC. +Qed. + +Lemma subm_subm : ∀ (f1 f2 : fin u) (j : nat), + subm f1 (subm f2 j) = subm (addm f1 j) f2. +Proof using . + intros. rewrite <-2 subf_subm, subm_subf, addm_addf. apply subf_subf. +Qed. + +Lemma subf_addf : ∀ f1 f2 f3 : fin u, + subf f1 (addf f2 f3) = subf (subf f1 f2) f3. +Proof using . + intros. rewrite 3 subfE, <- addfA, oppf_addf, subfE. reflexivity. +Qed. + +Lemma subf_addf_addf : ∀ f1 f2 f3 : fin u, + subf (addf f1 f3) (addf f2 f3) = subf f1 f2. +Proof using . intros. rewrite <- subf_subf, addfVKV. reflexivity. Qed. + +Lemma subf_addm_addm : ∀ (f1 f2 : fin u) (j : nat), + subf (addm f1 j) (addm f2 j) = subf f1 f2. +Proof using . intros. rewrite 2 addm_addf. apply subf_addf_addf. Qed. + +Lemma subf_addf_addfC : ∀ f1 f2 f3 : fin u, + subf (addf f1 f2) (addf f1 f3) = subf f2 f3. +Proof using . intros. rewrite 2 (addfC f1). apply subf_addf_addf. Qed. + +Lemma subfCAC : ∀ f1 f2 f3 f4 : fin u, + subf (subf f1 f2) (subf f3 f4) = subf (subf f1 f3) (subf f2 f4). +Proof using . + intros. rewrite <- subf_addf, addfC, addf_subf, addfC, + <- addf_subf, addfC, subf_addf. reflexivity. +Qed. + +Lemma subf_sucf : ∀ f1 f2 : fin u, subf (sucf f1) f2 = subf f1 (pref f2). +Proof using . intros. rewrite sucfE, prefE. symmetry. apply subf_subf. Qed. + +Lemma subf_pref : ∀ f1 f2 : fin u, subf (pref f1) f2 = subf f1 (sucf f2). +Proof using . + intros. rewrite sucfE, prefE, subfAC. symmetry. apply subf_addf. +Qed. + +Lemma oppf1 : oppf fin1 = fin_max. +Proof using . rewrite fin1_sucf_fin0, oppf_sucf, oppf0. apply pref0. Qed. + +Lemma oppm1 : oppm 1 = fin_max. +Proof using . rewrite oppm_oppf, <- fin1E. apply oppf1. Qed. + +Lemma oppf_max : oppf fin_max = fin1. +Proof using . rewrite <- oppf1. apply oppfK. Qed. + +(* + * + * The symmetric of f by the center c + * + *) + +Definition symf (c f : fin u) : fin u := addf c (subf c f). + +Definition symf_compat : Proper (equiv ==> equiv ==> equiv) symf := _. + +Lemma symfE : ∀ c f : fin u, symf c f = addf c (subf c f). +Proof using . reflexivity. Qed. + +Lemma symfEmod : ∀ c f : fin u, symf c f = mod2fin (u - f + c + c). +Proof using . + intros. rewrite symfE, addfE, subfEmod, addn_mod2fin_idemp_r, Nat.add_assoc, + (Nat.add_comm _ (_ - _)), Nat.add_assoc. reflexivity. +Qed. + +Lemma symf2nat : ∀ c f : fin u, symf c f = (u - f + c + c) mod u :> nat. +Proof using . intros. rewrite symfEmod. apply mod2fin2nat. Qed. + +Lemma symfI : ∀ c : fin u, Preliminary.injective equiv equiv (symf c). +Proof using . intros c f1 f2 H. eapply subfI, addfI, H. Qed. + +Lemma symfK : ∀ c f : fin u, symf c (symf c f) = f. +Proof using . + intros. rewrite 2 symfE, subf_addf, subff, sub0f, oppf_subf. apply subfVK. +Qed. + +Lemma sym0f : ∀ f : fin u, symf fin0 f = oppf f. +Proof using . intros. rewrite symfE, add0f. apply sub0f. Qed. + +Lemma symf0 : ∀ c : fin u, symf c fin0 = addf c c. +Proof using . intros. rewrite symfE, subf0. reflexivity. Qed. + +Lemma symff : ∀ f : fin u, symf f f = f. +Proof using . intros. rewrite symfE, subff. apply addf0. Qed. + +Lemma symf_addf : ∀ c f1 f2 : fin u, symf c (addf f1 f2) = subf (symf c f1) f2. +Proof using . + intros. rewrite 2 symfE, subf_addf, addfC, addf_subf, addfC. reflexivity. +Qed. + +Lemma symf_subf : ∀ c f1 f2 : fin u, symf c (subf f1 f2) = addf (symf c f1) f2. +Proof using . + intros. rewrite 2 symfE, subf_subf, <- addf_subf, addfA. reflexivity. +Qed. + +Lemma symf_sucf : ∀ c f : fin u, symf c (sucf f) = pref (symf c f). +Proof using . intros. rewrite sucfE, prefE. apply symf_addf. Qed. + +Lemma symf_pref : ∀ c f : fin u, symf c (pref f) = sucf (symf c f). +Proof using . intros. rewrite prefE, sucfE. apply symf_subf. Qed. + +(* + * + * The compatibility between various functions + * on fin u and both < and <= + * + *) + +Lemma addm_lt_large : ∀ (f : fin u) (j : nat), + fin0 < f -> oppf f <= j mod u -> addm f j < f. +Proof using . + intros * Hlt Hle. erewrite <- (Nat.mod_small f), <- Nat.Div0.mod_add, <- addm_mod, + addm2nat, Nat.mul_1_l. eapply mod_lt_between_compat; try rewrite Nat.mul_1_r. + - split. + + eapply Nat.le_trans. + all:swap 1 2. + * apply Nat.add_le_mono_l, Hle. + * rewrite addn_oppf. reflexivity. apply Hlt. + + apply Nat.add_lt_mono;try apply fin_lt. + apply mod2fin_lt. + - split. + + apply Nat.le_add_l. + + apply Nat.add_lt_mono_r. + apply fin_lt. + - apply Nat.add_lt_mono_l. + apply mod2fin_lt. + - apply fin_lt. +Qed. + +Lemma addm_le_small : + ∀ (f : fin u) (j : nat), j mod u < oppf f -> f <= addm f j. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f) as [| Hd]. subst. apply fin0_le. + erewrite <- (Nat.mod_small f), <- addm_mod, addm2nat. + eapply mod_le_between_compat. 1,2: rewrite Nat.mul_0_r. apply fin_between. + 2: apply Nat.le_add_r. 2: apply fin_lt. split. apply Nat.le_0_l. + eapply Nat.lt_le_trans. apply Nat.add_lt_mono_l, H. + rewrite addn_oppf. reflexivity. apply Hd. +Qed. + +Lemma addm_lt_small : + ∀ (f : fin u) (j : nat), fin0 < j mod u < oppf f -> f < addm f j. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_small, H2. + intros Habs. apply <- Nat.neq_0_lt_0. apply H1. eapply addm_betweenI. + 1,2: rewrite Nat.add_0_l. 1,2: split. 1,3: apply Nat.le_0_l. apply mod2fin_lt. + apply lt_0_u. symmetry. rewrite addm0, addm_mod. apply fin2natI, Habs. +Qed. + +Lemma addf_lt_large : + ∀ f1 f2 : fin u, fin0 < f1 -> oppf f1 <= f2 -> addf f1 f2 < f1. +Proof using . + intros * Hlt Hle. rewrite addf_addm. apply addm_lt_large. apply Hlt. + rewrite Nat.mod_small by apply fin_lt. apply Hle. +Qed. + +Lemma addf_le_small : + ∀ f1 f2 : fin u, f2 < oppf f1 -> f1 <= addf f1 f2. +Proof using . + intros * H. rewrite addf_addm. apply addm_le_small. + rewrite Nat.mod_small by apply fin_lt. apply H. +Qed. + +Lemma addf_lt_small : ∀ f1 f2 : fin u, fin0 < f2 < oppf f1 -> f1 < addf f1 f2. +Proof using . + intros * H. rewrite addf_addm. apply addm_lt_small. + rewrite Nat.mod_small. apply H. apply fin_lt. +Qed. + +Lemma addm_le_compat_large : ∀ (f : fin u) (j1 j2 : nat), + oppf f <= j1 mod u <= j2 mod u -> addm f j1 <= addm f j2. +Proof using . + intros * [H1 H2]. destruct (eq_fin0_lt_dec f) as [| Hd]. + - subst. rewrite (add0m j1), (add0m j2), 2 mod2fin2nat. apply H2. + - rewrite <- (addm_mod f j1), <- (addm_mod f j2), 2 addmE. + eapply mod2fin_le_between_compat. 3: apply Nat.add_le_mono_l, H2. + all: rewrite Nat.mul_1_r. all: split. 2,4: apply Nat.add_lt_mono, mod2fin_lt. + 2,3: apply fin_lt. all: eapply Nat.le_trans. 2,4: apply Nat.add_le_mono_l. + 3: eapply Nat.le_trans. 2,3: apply H1. 2: apply H2. all: rewrite addn_oppf. + 1,3: reflexivity. all: apply Hd. +Qed. + +Lemma addm_le_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j1 mod u <= j2 mod u < oppf f -> addm f j1 <= addm f j2. +Proof using . + intros * [H1 H2]. destruct (eq_fin0_lt_dec f) as [| Hd]. + - subst. rewrite (add0m j1), (add0m j2), 2 mod2fin2nat. apply H1. + - rewrite <- (addm_mod f j1), <- (addm_mod f j2), 2 addmE. + eapply mod2fin_le_between_compat. 3: apply Nat.add_le_mono_l, H1. + all: rewrite Nat.mul_0_r, Nat.add_0_l. all: split. 1,3: apply Nat.le_0_l. + eapply Nat.le_lt_trans. apply Nat.add_le_mono_l, H1. all: eapply Nat.lt_le_trans. + 1,3: apply Nat.add_lt_mono_l, H2. all: rewrite addn_oppf. + 1,3: reflexivity. all: apply Hd. +Qed. + +Lemma addf_le_compat_large : ∀ f1 f2 f3 : fin u, + oppf f1 <= f2 <= f3 -> addf f1 f2 <= addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_le_compat_large. + rewrite 2 Nat.mod_small by apply fin_lt. apply H. +Qed. + +Lemma addf_le_compat_small : ∀ f1 f2 f3 : fin u, + f2 <= f3 < oppf f1 -> addf f1 f2 <= addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_le_compat_small. + rewrite 2 Nat.mod_small by apply fin_lt. apply H. +Qed. + +Lemma addm_lt_compat_large : ∀ (f : fin u) (j1 j2 : nat), + oppf f <= j1 mod u < j2 mod u -> addm f j1 < addm f j2. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_compat_large. + split. apply H1. apply Nat.lt_le_incl, H2. rewrite <- (addm_mod _ j1), + <- (addm_mod _ j2). intros Habs. eapply Nat.lt_irrefl. + erewrite addm_betweenI. apply H2. 3: apply fin2natI, Habs. + all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. +Qed. + +Lemma addm_lt_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j1 mod u < j2 mod u < oppf f -> addm f j1 < addm f j2. +Proof using . + intros * [H1 H2]. apply Nat.le_neq. split. apply addm_le_compat_small. + split. apply Nat.lt_le_incl, H1. apply H2. rewrite <- (addm_mod _ j1), + <- (addm_mod _ j2). intros Habs. eapply Nat.lt_irrefl. + erewrite addm_betweenI. apply H1. 3: apply fin2natI, Habs. + all: split. 1,3: apply Nat.le_0_l. all: apply mod2fin_lt. +Qed. + +Lemma addf_lt_compat_large : + ∀ f1 f2 f3 : fin u, oppf f1 <= f2 < f3 -> addf f1 f2 < addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_lt_compat_large. + rewrite 2 Nat.mod_small. apply H. all: apply fin_lt. +Qed. + +Lemma addf_lt_compat_small : + ∀ f1 f2 f3 : fin u, f2 < f3 < oppf f1 -> addf f1 f2 < addf f1 f3. +Proof using . + intros * H. rewrite 2 addf_addm. apply addm_lt_compat_small. + rewrite 2 Nat.mod_small. apply H. all: apply fin_lt. +Qed. + +Lemma lt_fin0_oppf : ∀ f : fin u, fin0 < f -> fin0 < oppf f. +Proof using . + intros * H1. apply neq_fin0_lt_fin0 in H1. apply neq_fin0_lt_fin0. intros H2. + apply H1. apply oppfI. rewrite oppf0. apply H2. +Qed. + +Lemma oppf_le_compat : ∀ f1 f2 : fin u, fin0 < f1 <= f2 -> oppf f2 <= oppf f1. +Proof using . + intros * [H1 H2]. rewrite 2 oppfEmod. eapply mod2fin_le_between_compat. + 1,2: rewrite Nat.mul_0_r, Nat.add_0_l. 1,2: split. 1,3: apply Nat.le_0_l. + 3: apply Nat.sub_le_mono_l, H2. all: apply lt_sub_u. + eapply Nat.lt_le_trans. 1,3: apply H1. apply H2. +Qed. + +Lemma oppf_lt_compat : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> oppf f2 < oppf f1. +Proof using . + intros * [H1 H2]. rewrite <- (Nat.lt_succ_pred 0 (oppf f1)). apply Nat.lt_succ_r. + rewrite pred_pref, <- oppf_sucf. apply oppf_le_compat. rewrite <- S_sucf. + split. apply Nat.lt_0_succ. 2: eapply Nat.lt_le_trans. 1,2: apply H2. + apply le_max. all: apply lt_fin0_oppf, H1. +Qed. + +Lemma oppf_le_inj : ∀ f1 f2 : fin u, fin0 < oppf f2 <= oppf f1 -> f1 <= f2. +Proof using . + intros * H. rewrite <- (oppfK f1), <- (oppfK f2). apply oppf_le_compat, H. +Qed. + +Lemma oppf_lt_inj : ∀ f1 f2 : fin u, fin0 < oppf f2 < oppf f1 -> f1 < f2. +Proof using . + intros * H. rewrite <- (oppfK f1), <- (oppfK f2). apply oppf_lt_compat, H. +Qed. + +Lemma opmm_le_compat : + ∀ j1 j2 : nat, fin0 < j1 mod u <= j2 mod u -> oppm j2 <= oppm j1. +Proof using . + intros * H. rewrite 2 oppm_oppf. apply oppf_le_compat. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma oppm_lt_compat : + ∀ j1 j2 : nat, fin0 < j1 mod u < j2 mod u -> oppm j2 < oppm j1. +Proof using . + intros * H. rewrite 2 oppm_oppf. apply oppf_lt_compat. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma oppm_le_inj : + ∀ j1 j2 : nat, fin0 < oppm j2 <= oppm j1 -> j1 mod u <= j2 mod u. +Proof using . + intros * H. rewrite <-2 mod2fin2nat. apply oppf_le_inj. + rewrite <-2 oppm_oppf. apply H. +Qed. + +Lemma oppm_lt_inj : + ∀ j1 j2 : nat, fin0 < oppm j2 < oppm j1 -> j1 mod u < j2 mod u. +Proof using . + intros * H. rewrite <-2 mod2fin2nat. apply oppf_lt_inj. + rewrite <-2 oppm_oppf. apply H. +Qed. + +Lemma lt_fin0_subf : ∀ f1 f2 : fin u, f1 ≠f2 -> fin0 < subf f1 f2. +Proof using . + intros * H1. apply neq_fin0_lt_fin0. intros H2. apply H1. + eapply subIf. rewrite (subff f2). apply H2. +Qed. + +Lemma subf_lt_small : ∀ f1 f2 : fin u, fin0 < f2 <= f1 -> subf f1 f2 < f1. +Proof using . + intros * H. rewrite subfE. apply addf_lt_large. + 2: apply oppf_le_compat, H. eapply Nat.lt_le_trans. all: apply H. +Qed. + +Lemma subf_le_small : + ∀ f1 f2 : fin u, f2 <= f1 -> subf f1 f2 <= f1. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| H1]. subst. + rewrite sub0f, le_fin0, <- oppf0. f_equal. apply le_fin0, H. + destruct (eq_fin0_lt_dec f2) as [| H2]. subst. rewrite subf0. + reflexivity. rewrite subfE. apply Nat.lt_le_incl, addf_lt_large. + apply H1. apply oppf_le_compat. split. apply H2. apply H. +Qed. + +Lemma subf_lt_large : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> f1 < subf f1 f2. +Proof using . + intros * H. rewrite subfE. apply addf_lt_small. split. + 2: apply oppf_lt_compat, H. eapply lt_fin0_oppf, Nat.lt_trans. all: apply H. +Qed. + +Lemma subf_le_large : ∀ f1 f2 : fin u, f1 < f2 -> f1 <= subf f1 f2. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. apply fin0_le. + rewrite subfE. apply addf_le_small, oppf_lt_compat. split. apply Hd. apply H. +Qed. + +Lemma subm_lt_small : + ∀ (f : fin u) (j : nat), fin0 < j mod u <= f -> subm f j < f. +Proof using . + intros * H. rewrite subm_subf. apply subf_lt_small. + rewrite mod2fin2nat. apply H. +Qed. + +Lemma subm_le_small : + ∀ (f : fin u) (j : nat), j mod u <= f -> subm f j <= f. +Proof using . + intros * H. rewrite subm_subf. apply subf_le_small. + rewrite mod2fin2nat. apply H. +Qed. + +Lemma subm_lt_large : + ∀ (f : fin u) (j : nat), fin0 < f < j mod u -> f < subm f j. +Proof using . + intros * H. rewrite subm_subf. apply subf_lt_large. + rewrite mod2fin2nat. apply H. +Qed. + +Lemma subm_le_large : ∀ (f : fin u) (j : nat), f < j mod u -> f <= subm f j. +Proof using . + intros * H. rewrite subm_subf. apply subf_le_large. + rewrite mod2fin2nat. apply H. +Qed. + +Lemma subf_le_compat_small : + ∀ f1 f2 f3 : fin u, f3 <= f2 <= f1 -> subf f1 f2 <= subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f3) as [| Hd]. subst. + rewrite subf0. apply subf_le_small, H. rewrite 2 subfE. + apply addf_le_compat_large. split. all: apply oppf_le_compat. + all: split. eapply Nat.lt_le_trans. 1,4: apply Hd. all: apply H. +Qed. + +Lemma subf_le_compat_large : + ∀ f1 f2 f3 : fin u, f1 < f3 <= f2 -> subf f1 f2 <= subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. + rewrite 2 sub0f. apply oppf_le_compat, H. rewrite 2 subfE. + apply addf_le_compat_small. split. apply oppf_le_compat. split. + eapply lt_fin0. 1,2: apply H. apply oppf_lt_compat. split. apply Hd. apply H. +Qed. + +Lemma subf_lt_compat_small : + ∀ f1 f2 f3 : fin u, f3 < f2 <= f1 -> subf f1 f2 < subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f3) as [| Hd]. subst. + rewrite subf0. apply subf_lt_small, H. rewrite 2 subfE. + apply addf_lt_compat_large. split. apply oppf_le_compat. split. + eapply lt_fin0. 1,2: apply H. apply oppf_lt_compat. split. apply Hd. apply H. +Qed. + +Lemma subf_lt_compat_large : + ∀ f1 f2 f3 : fin u, f1 < f3 < f2 -> subf f1 f2 < subf f1 f3. +Proof using . + intros * H. destruct (eq_fin0_lt_dec f1) as [| Hd]. subst. + rewrite 2 sub0f. apply oppf_lt_compat, H. rewrite 2 subfE. + apply addf_lt_compat_small. split. all: apply oppf_lt_compat. + all: split. eapply lt_fin0. 1,2,4: apply H. apply Hd. +Qed. + +Lemma subm_le_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j2 mod u <= j1 mod u <= f -> subm f j1 <= subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_le_compat_small. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma subm_le_compat_large : ∀ (f : fin u) (j1 j2 : nat), + f < j2 mod u <= j1 mod u -> subm f j1 <= subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_le_compat_large. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma subm_lt_compat_small : ∀ (f : fin u) (j1 j2 : nat), + j2 mod u < j1 mod u <= f -> subm f j1 < subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_lt_compat_small. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma subm_lt_compat_large : ∀ (f : fin u) (j1 j2 : nat), + f < j2 mod u < j1 mod u -> subm f j1 < subm f j2. +Proof using . + intros * H. rewrite 2 subm_subf. apply subf_lt_compat_large. + rewrite 2 mod2fin2nat. apply H. +Qed. + +Lemma subf_lt_smallV : ∀ f1 f2 : fin u, fin0 < f2 < f1 -> subf f1 f2 < oppf f2. +Proof using . + intros * H. apply oppf_lt_inj. rewrite oppfK, oppf_subf. + split. apply H. apply subf_lt_large, H. +Qed. + +Lemma subf_le_smallV : + ∀ f1 f2 : fin u, fin0 < f2 < f1 -> subf f1 f2 <= oppf f2. +Proof using . + intros * H. apply oppf_le_inj. rewrite oppfK, oppf_subf. + split. apply H. apply subf_le_large, H. +Qed. + +Lemma subf_lt_largeV : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> oppf f2 < subf f1 f2. +Proof using . + intros * H. apply oppf_lt_inj. rewrite oppfK, oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_lt_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_le_largeV : ∀ f1 f2 : fin u, f1 < f2 -> oppf f2 <= subf f1 f2. +Proof using . + intros * H. apply oppf_le_inj. rewrite oppfK, oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_le_small. + 2: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_le_compat_largeV : + ∀ f1 f2 f3 : fin u, f1 <= f3 <= f2 -> subf f3 f1 <= subf f2 f1. +Proof using . + intros * [H1 H2]. destruct (le_lt_eq_dec H1) as [Hd |]. 2:{ subst. + rewrite subff. apply fin0_le. } apply oppf_le_inj. rewrite 2 oppf_subf. + split. apply lt_fin0_subf, lt_gt_cases. left. eapply Nat.lt_le_trans. + 3: apply subf_le_compat_large. 3: split. 1,3: apply Hd. all: apply H2. +Qed. + +Lemma subf_le_compat_smallV : + ∀ f1 f2 f3 : fin u, f3 <= f2 < f1 -> subf f3 f1 <= subf f2 f1. +Proof using . + intros * H. apply oppf_le_inj. rewrite 2 oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_le_compat_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma subf_lt_compat_largeV : + ∀ f1 f2 f3 : fin u, f1 <= f3 < f2 -> subf f3 f1 < subf f2 f1. +Proof using . + intros * [H1 H2]. destruct (le_lt_eq_dec H1) as [Hd |]. 2: subst. + 2: rewrite subff. apply oppf_lt_inj. rewrite 2 oppf_subf. split. + 1,3: apply lt_fin0_subf, lt_gt_cases. 3: apply subf_lt_compat_large. + left. 2: right. 3: split. eapply Nat.lt_trans. 1,4: apply Hd. all: apply H2. +Qed. + +Lemma subf_lt_compat_smallV : + ∀ f1 f2 f3 : fin u, f3 < f2 < f1 -> subf f3 f1 < subf f2 f1. +Proof using . + intros * H. apply oppf_lt_inj. rewrite 2 oppf_subf. split. + apply lt_fin0_subf, lt_gt_cases. right. 2: apply subf_lt_compat_small. + 2: split. 3: apply Nat.lt_le_incl. all: apply H. +Qed. + +Lemma sucf_le_compat : + ∀ f1 f2 : fin u, f1 <= f2 < fin_max -> sucf f1 <= sucf f2. +Proof using . + intros * H. rewrite 2 sucfE, (addfC f1), (addfC f2). + apply addf_le_compat_small. rewrite oppf1. apply H. +Qed. + +Lemma sucf_lt_compat : + ∀ f1 f2 : fin u, f1 < f2 < fin_max -> sucf f1 < sucf f2. +Proof using . + intros * H. rewrite 2 sucfE, (addfC f1), (addfC f2). + apply addf_lt_compat_small. rewrite oppf1. apply H. +Qed. + +Lemma pref_le_compat : + ∀ f1 f2 : fin u, fin0 < f1 <= f2 -> pref f1 <= pref f2. +Proof using . + intros * H. rewrite 2 prefE. apply subf_le_compat_largeV. split. + 2: apply H. apply lt_pref_le. rewrite <- fin0_pref_fin1. apply H. +Qed. + +Lemma pref_lt_compat : ∀ f1 f2 : fin u, fin0 < f1 < f2 -> pref f1 < pref f2. +Proof using . + intros * H. rewrite 2 prefE. apply subf_lt_compat_largeV. split. + 2: apply H. apply lt_pref_le. rewrite <- fin0_pref_fin1. apply H. +Qed. + +(* + * + * The bijections on fin u whose section adds + * and whose retraction substracts + * + *) + +Definition asbf (f : fin u) : bijection (fin u) := + cancel_bijection (λ f2 : fin u, subf f2 f) _ (@addIf f) + (addfVKV f) (subfVKV f). +Global Opaque asbf. + +Definition asbm (j : nat) : bijection (fin u) := + cancel_bijection (λ f : fin u, subm f j) _ (@addIm j) + (addmVKV j) (submVKV j). +Global Opaque asbm. + +Lemma asbfE : ∀ f1 : fin u, asbf f1 = (λ f2, addf f2 f1) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbfVE : + ∀ f1 : fin u, (asbf f1)â»Â¹ = (λ f2, subf f2 f1) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbmE : ∀ j : nat, asbm j = (λ f, addm f j) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbmVE : ∀ j : nat, (asbm j)â»Â¹ = (λ f, subm f j) :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbm_asbf : ∀ j : nat, asbm j == asbf (mod2fin j). +Proof using . intros j f. rewrite asbmE, asbfE. apply addm_addf. Qed. + +Lemma asbf_asbm : ∀ f1 : fin u, asbf f1 == asbm f1. +Proof using . intros f1 f2. rewrite asbmE, asbfE. apply addf_addm. Qed. + +Lemma asbmV_asbfV : ∀ j : nat, (asbm j)â»Â¹ == (asbf (mod2fin j))â»Â¹. +Proof using . intros j f. rewrite asbmVE, asbfVE. apply subm_subf. Qed. + +Lemma asbfV_asbmV : ∀ f1 : fin u, (asbf f1)â»Â¹ == (asbm f1)â»Â¹. +Proof using . intros f1 f2. rewrite asbmVE, asbfVE. apply subf_subm. Qed. + +Lemma asbm_mod : ∀ j : nat, asbm (j mod u) == asbm j. +Proof using . intros j f. rewrite 2 asbmE. apply addm_mod. Qed. + +Lemma asbmV_mod : ∀ j : nat, (asbm (j mod u))â»Â¹ == (asbm j)â»Â¹. +Proof using . intros j f. rewrite 2 asbmVE. apply subm_mod. Qed. + +Lemma asbm0 : asbm 0 == id. +Proof using . intros f. rewrite asbmE. apply addm0. Qed. + +Lemma asbf0 : asbf fin0 == id. +Proof using . intros f. rewrite asbfE. apply addf0. Qed. + +Lemma asbfAC : ∀ f1 f2 : fin u, asbf f2 ∘ (asbf f1) == asbf f1 ∘ (asbf f2). +Proof using . intros * f3. rewrite 2 compE, 2 asbfE. apply addfAC. Qed. + +Lemma asbmAC : ∀ j1 j2 : nat, asbm j2 ∘ (asbm j1) == asbm j1 ∘ (asbm j2). +Proof using . intros * f. rewrite 2 compE, 2 asbmE. apply addmAC. Qed. + +Lemma asbmVAC : + ∀ j1 j2 : nat, (asbm j2)â»Â¹ ∘ (asbm j1)â»Â¹ == (asbm j1)â»Â¹ ∘ (asbm j2)â»Â¹. +Proof using . intros * f. rewrite 2 compE, 2 asbmVE. apply submAC. Qed. + +Lemma asbfVAC : + ∀ f1 f2 : fin u, (asbf f2)â»Â¹ ∘ (asbf f1)â»Â¹ == (asbf f1)â»Â¹ ∘ (asbf f2)â»Â¹. +Proof using . intros * f3. rewrite 2 compE, 2 asbfVE. apply subfAC. Qed. + +Lemma asbmCV : ∀ j1 j2 : nat, asbm j1 ∘ (asbm j2)â»Â¹ == (asbm j2)â»Â¹ ∘ (asbm j1). +Proof using . intros * f. rewrite 2 compE, asbmVE, asbmE. apply addm_subm. Qed. + +Lemma asbfCV : + ∀ f1 f2 : fin u, asbf f1 ∘ (asbf f2)â»Â¹ == (asbf f2)â»Â¹ ∘ (asbf f1). +Proof using . intros * f3. rewrite 2 compE, asbfVE, asbfE. apply addf_subf. Qed. + +Lemma asbfA : ∀ f1 f2 : fin u, asbf (asbf f1 f2) == asbf f1 ∘ (asbf f2). +Proof using . intros * f3. rewrite compE, 3 asbfE. apply addfA. Qed. + +Lemma asbfAV : + ∀ f1 f2 : fin u, asbf ((asbf f1)â»Â¹ f2) == (asbf f1)â»Â¹ ∘ (asbf f2). +Proof using . + intros * f3. rewrite compE, 2 asbfE, + asbfVE, addfC, (addfC f3). apply addf_subf. +Qed. + +Lemma asbfVA : + ∀ f1 f2 : fin u, (asbf (asbf f1 f2))â»Â¹ == (asbf f1)â»Â¹ ∘ (asbf f2)â»Â¹. +Proof using . + intros * f3. rewrite compE, 3 asbfVE, asbfE. apply subf_addf. +Qed. + +Lemma asbfVAV : + ∀ f1 f2 : fin u, (asbf ((asbf f1)â»Â¹ f2))â»Â¹ == asbf f1 ∘ (asbf f2)â»Â¹. +Proof using . + intros * f3. rewrite compE, 3 asbfVE, asbfE, addf_subf. apply subf_subf. +Qed. + +Lemma asbfVf : ∀ f : fin u, (asbf f)â»Â¹ f = fin0. +Proof using . intros. rewrite asbfVE. apply subff. Qed. + +(* + * + * The bijection on fin u whose section is sucf + * and whose retraction is pref + * + *) + +Definition spbf : bijection (fin u) + := cancel_bijection pref _ sucfI sucfK prefK. +Global Opaque spbf. + +Lemma spbfE : spbf = sucf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma spbfVE : spbfâ»Â¹ = pref :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma asbf1 : asbf fin1 == spbf. +Proof using . intros f. rewrite asbfE, spbfE. symmetry. apply sucfE. Qed. + +Lemma asbm1 : asbm 1 == spbf. +Proof using . rewrite asbm_asbf, <- fin1E. apply asbf1. Qed. + +(* + * + * The bijection on fin u whose section and retraction are revf + * + *) + +Definition rebf : bijection (fin u) + := cancel_bijection revf _ revfI revfK revfK. +Global Opaque rebf. + +Lemma rebfE : rebf = revf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma rebfVE : rebfâ»Â¹ = revf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma rebfV : rebfâ»Â¹ == rebf. +Proof using . intros f. rewrite rebfE, rebfVE. reflexivity. Qed. + +(* + * + * The bijection on fin u whose section and retraction are oppf + * + *) + +Definition opbf : bijection (fin u) + := cancel_bijection oppf _ oppfI oppfK oppfK. +Global Opaque opbf. + +Lemma opbfE : opbf = oppf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma opbfVE : opbfâ»Â¹ = oppf :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma opbfV : opbfâ»Â¹ == opbf. +Proof using . intros f. rewrite opbfE, opbfVE. reflexivity. Qed. + +Lemma spbf_rebf : spbf ∘ rebf == opbf. +Proof using . + intros f. rewrite compE, spbfE, rebfE, opbfE. symmetry. apply oppfE. +Qed. + +Lemma spbfV_opbf : spbfâ»Â¹ ∘ opbf == rebf. +Proof using. + rewrite <- spbf_rebf, compose_assoc, compose_inverse_l. apply id_comp_l. +Qed. + +(* + * + * The bijection on fin u whose section and retraction are symf + * + *) + +Definition sybf (c : fin u) : bijection (fin u) := + cancel_bijection (symf c) _ (@symfI c) (symfK c) (symfK c). +Global Opaque sybf. + +Lemma sybfE : ∀ c : fin u, sybf c = symf c :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma sybfVE : ∀ c : fin u, sybf câ»Â¹ = symf c :> (fin u -> fin u). +Proof using . reflexivity. Qed. + +Lemma sybfV : ∀ c : fin u, sybf câ»Â¹ == sybf c. +Proof using . intros c f. rewrite sybfE, sybfVE. reflexivity. Qed. + +Lemma sybfK : ∀ c : fin u, sybf c ∘ sybf c == id. +Proof using . intros c f. rewrite compE, sybfE. apply symfK. Qed. + +Lemma sybff : ∀ c : fin u, sybf c c = c. +Proof using . intros. rewrite sybfE. apply symff. Qed. + +End mod2fin. diff --git a/Util/Lexprod.v b/Util/Lexprod.v index 1bd89797fb3be6e10eb952ab524614e918c88029..c468d2287fa46b8dca1f1dac7de35b211045c5f9 100644 --- a/Util/Lexprod.v +++ b/Util/Lexprod.v @@ -79,7 +79,7 @@ End WfLexicographic_Product. Global Arguments lexprod [A] [B] leA leB _ _. -Instance lexprod_compat: Proper (eq * eq ==> eq * eq ==> iff) (lexprod lt lt). +Global Instance lexprod_compat: Proper (eq * eq ==> eq * eq ==> iff) (lexprod lt lt). Proof using . intros (a, b) (a', b') (heqa, heqb) (c, d) (c', d') (heqc, heqd). hnf in *|-. simpl in *|-. now subst. diff --git a/Util/ListComplements.v b/Util/ListComplements.v index 742d86e22ea5e85604caf909351180f7f7a41ae2..52e9301990d5b7acf19344860b6111863ce93ac1 100644 --- a/Util/ListComplements.v +++ b/Util/ListComplements.v @@ -24,15 +24,17 @@ Require Import List SetoidList. Require Export SetoidPermutation. Require Import Sorting.Permutation. Require Import Bool. -Require Import Arith.Div2 Lia PeanoNat. +Require Import Lia PeanoNat. Require Import Psatz. Require Import SetoidDec. Require Import Pactole.Util.Preliminary. Require Import Pactole.Util.NumberComplements. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Set Implicit Arguments. -Arguments PermutationA {A}%type eqA%signature l1%list l2%list. +Arguments PermutationA {A}%_type eqA%_signature l1%_list l2%_list. (******************************) @@ -52,6 +54,15 @@ intros x l Hin. induction l. - destruct Hin. subst. now left. right. auto. Qed. *) +Lemma all_eq : forall (l : list A) (a1 : A), In a1 l ->length l = 1 + -> forall a2 : A, In a2 l -> a2 = a1. +Proof using . + intros * Hi1 H * Hi2. destruct l as [|h t]. + rewrite (proj2 (length_zero_iff_nil _)) in H by reflexivity. discriminate H. + destruct t as [|h1 t]. cbn in Hi1, Hi2. destruct Hi1, Hi2. subst. reflexivity. + 1-3: contradiction. exfalso. eapply Nat.neq_succ_0, eq_add_S, H. +Qed. + Lemma InA_Leibniz : forall (x : A) l, InA Logic.eq x l <-> In x l. Proof using . intros x l. split; intro Hl; induction l; inversion_clear Hl; @@ -64,14 +75,14 @@ Proof using . intros [| x l]; auto. right. exists x. now left. Qed. Lemma not_nil_In : forall l : list A, l <> nil -> exists x, In x l. Proof using . intros [| x l] Hl. -- now elim Hl. +- now contradiction Hl. - exists x. now left. Qed. Lemma not_nil_last : forall l, l <> nil -> exists (a : A) l', l = l' ++ a :: nil. Proof using . intros l Hl. induction l. -+ now elim Hl. ++ now contradiction Hl. + destruct l. - exists a, nil. reflexivity. - destruct (IHl ltac:(discriminate)) as [b [l'' Hl'']]. @@ -91,7 +102,7 @@ intros d x l Hin. induction l as [| e l]. + inversion_clear Hin. - exists 0, e. repeat split; trivial; simpl; lia. - destruct IHl as [n [y [Hn [Hy Hl]]]]; trivial; []. - apply Lt.lt_n_S in Hn. exists (S n), y. now repeat split. + exists (S n), y. repeat split; cbn; auto; lia. Qed. (* Already exists as [List.In_nth] but with a reversed argument order @@ -108,7 +119,7 @@ Qed. *) exists l1, exists l2, ~List.In x l1 /\ l = l1 ++ x :: l2. Proof. intros x l. induction l as [| a l]; intro Hin. - now elim (List.in_nil Hin). + now contradiction (List.in_nil Hin). destruct Hin. subst. exists nil. exists l. intuition. destruct (IHl H) as [l1 [l2 [Hnin Heq]]]. @@ -127,28 +138,28 @@ Qed. Lemma hd_indep : forall l (d d' : A), l <> nil -> hd d l = hd d' l. Proof using . intros [| x l] d d' Hl. -- now elim Hl. +- now contradiction Hl. - reflexivity. Qed. Lemma last_indep : forall l (d d' : A), l <> nil -> last l d = last l d'. Proof using . induction l as [| x l]; intros d d' Hl. -- now elim Hl. +- now contradiction Hl. - destruct l; trivial. apply IHl. discriminate. Qed. Lemma hd_In : forall (d : A) l, l <> nil -> In (hd d l) l. Proof using . intros d [| x l] Hl. -- now elim Hl. +- now contradiction Hl. - now left. Qed. Lemma last_In : forall l (d : A), l <> List.nil -> List.In (List.last l d) l. Proof using . induction l as [| x l]; intros d Hl. -- now elim Hl. +- now contradiction Hl. - destruct l; try now left. right. apply IHl. discriminate. Qed. @@ -167,7 +178,7 @@ Proof using . intros d l1 l2 Hl2. induction l1; simpl. reflexivity. assert (l1 ++ l2 <> nil). { intro Habs. apply Hl2. now destruct (app_eq_nil _ _ Habs). } - destruct (l1 ++ l2). now elim H. assumption. + destruct (l1 ++ l2). now contradiction H. assumption. Qed. Lemma rev_nil : forall l : list A, rev l = nil <-> l = nil. @@ -317,9 +328,9 @@ induction l. + inversion_clear Hfl. rewrite InA_cons in *. destruct Hx as [Hx | Hx], Hy as [Hy | Hy]. - now rewrite Hx, Hy. - - match goal with H : ~ InA _ _ _ |- _ => elim H end. + - match goal with H : ~ InA _ _ _ |- _ => contradiction H end. rewrite <- Hx, Hxy, InA_map_iff; firstorder. - - match goal with H : ~ InA _ _ _ |- _ => elim H end. + - match goal with H : ~ InA _ _ _ |- _ => contradiction H end. rewrite <- Hy, <- Hxy, InA_map_iff; firstorder. - auto. Qed. @@ -727,7 +738,7 @@ intros. split. destruct lâ‚‚ as [| x'' [| y'' [| ? ?]]]; discriminate Hlength || clear Hlength. destruct (IHperm1 _ _ _ _ ltac:(reflexivity) ltac:(reflexivity)) as [[H11 H12] | [H11 H12]], (IHperm2 _ _ _ _ ltac:(reflexivity) ltac:(reflexivity)) as [[H21 H22] | [H21 H22]]; - rewrite H11, H12, <- H21, <- H22; intuition. + rewrite H11, H12, <- H21, <- H22; intuition auto with *. + intros [[Heq1 Heq2] | [Heq1 Heq2]]; rewrite Heq1, Heq2. reflexivity. now constructor 3. Qed. @@ -829,7 +840,7 @@ Lemma NoDupA_2 : forall x y, ~eqA x y -> NoDupA eqA (x :: y :: nil). Proof using . intros x y Hdiff. constructor. intro Habs. inversion_clear Habs. - now elim Hdiff. + now contradiction Hdiff. inversion H. apply NoDupA_singleton. Qed. @@ -899,7 +910,7 @@ Lemma not_NoDupA : (forall x y, {eqA x y} + {~eqA x y} ) -> Proof using HeqA. intros eq_dec l. split; intro Hl. * induction l. - + elim Hl. now constructor. + + contradiction Hl. now constructor. + destruct (InA_dec eq_dec a l) as [Hin | Hnin]. - exists a. apply (PermutationA_split _) in Hin. destruct Hin as [l' Hperm]. exists l'. now rewrite Hperm. @@ -1058,7 +1069,7 @@ Qed. Lemma inclA_dec : forall l1 l2, {inclA eqA l1 l2} + {~inclA eqA l1 l2}. Proof. induction l1 as [| x1 l1 Hrec]; intro l2. -* left. abstract (intros x Habs; rewrite InA_nil in Habs; elim Habs). +* left. abstract (intros x Habs; rewrite InA_nil in Habs; contradiction Habs). * refine (match InA_dec eq_dec x1 l2 with | left in_x => match Hrec l2 with | left in_l => left _ @@ -1093,7 +1104,7 @@ Lemma not_inclA : forall l1 l2, ~inclA eqA l1 l2 <-> exists x, InA eqA x l1 /\ ~ Proof using HeqA eq_dec. intros l1 l2. split; intro H. * induction l1 as [| e l1]. - + elim H. intro. now rewrite InA_nil. + + contradiction H. intro. now rewrite InA_nil. + destruct (InA_dec eq_dec e l2). - assert (Hincl : ~ inclA eqA l1 l2). { intro Hincl. apply H. intros x Hin. inversion_clear Hin. @@ -1181,14 +1192,14 @@ intros a b Hab l. induction l as [| x l]; intros [| x' l'] Hl. apply (PermutationA_cons_inv _), IHl in Hl. simpl. rewrite Hl. repeat rewrite countA_occ_app. simpl. destruct (eq_dec x a) as [Hx | Hx], (eq_dec y b) as [Hy | Hy]; try lia. - - elim Hy. now rewrite <- Hxy, <- Hab. - - elim Hx. now rewrite Hxy, Hab. + - contradiction Hy. now rewrite <- Hxy, <- Hab. + - contradiction Hx. now rewrite Hxy, Hab. Qed. Lemma countA_occ_alls_in : forall x n, countA_occ x (alls x n) = n. Proof using HeqA. intros x n. induction n; simpl; trivial. -destruct (eq_dec x x) as [| Hneq]. now rewrite IHn. now elim Hneq. +destruct (eq_dec x x) as [| Hneq]. now rewrite IHn. now contradiction Hneq. Qed. Lemma countA_occ_alls_out : forall x y n, ~eqA x y -> countA_occ y (alls x n) = 0%nat. @@ -1199,11 +1210,11 @@ Proof using HeqA. intros x l. induction l as [| a l]; simpl. + split; intro Habs. - lia. - - rewrite InA_nil in Habs. elim Habs. + - rewrite InA_nil in Habs. contradiction Habs. + destruct (eq_dec a x) as [Heq | Heq]. - split; intro; lia || now left. - rewrite IHl. split; intro Hin; try now right; []. - inversion_clear Hin; trivial. now elim Heq. + inversion_clear Hin; trivial. now contradiction Heq. Qed. Lemma countA_occ_length_le : forall l (x : A), countA_occ x l <= length l. @@ -1259,7 +1270,7 @@ intros x l. induction l as [| a l]; intro n; simpl. destruct n as [| n]; try lia; []. f_equal. rewrite IHl. simpl in Hn. rewrite <- Heq in Hn at 1. apply PermutationA_cons_inv in Hn; autoclass. - + rewrite IHl. destruct (eq_dec x a); try (now elim Heq); []. + + rewrite IHl. destruct (eq_dec x a); try (now contradiction Heq); []. rewrite <- PermutationA_middle; autoclass; []. split; intro Hperm. - now constructor. @@ -1299,8 +1310,8 @@ Qed. Lemma firstn_In : forall (l : list A) n, incl (firstn n l) l. Proof using . induction l; intros n x Hin. -+ destruct n; elim Hin. -+ destruct n; try now elim Hin. simpl in Hin. destruct Hin. ++ destruct n; contradiction Hin. ++ destruct n; try now contradiction Hin. simpl in Hin. destruct Hin. - subst. now left. - right. now apply IHl in H. Qed. @@ -1313,7 +1324,7 @@ Lemma firstn_add_tl_In : forall n l (x a : A), Proof using . induction n; intros l x a Hin; simpl in *. - assumption. -- destruct l as [| b l]; simpl in *; solve [elim Hin | intuition]. +- destruct l as [| b l]; simpl in *; solve [contradiction Hin | intuition]. Qed. Lemma firstn_add_tl : forall l n (a : A), n <= length l -> firstn n (l ++ a :: nil) = firstn n l. @@ -1338,8 +1349,8 @@ Qed. Lemma skipn_In : forall (l : list A) n, incl (skipn n l) l. Proof using . induction l; intros n x Hin. -+ destruct n; elim Hin. -+ destruct n; try now elim Hin. simpl in Hin. apply IHl in Hin. now right. ++ destruct n; contradiction Hin. ++ destruct n; [apply Hin|]. simpl in Hin. apply IHl in Hin. now right. Qed. Lemma In_skipn : forall l l' (pt : A) n, n <= length l -> In pt (skipn n (l ++ pt :: l')). @@ -1360,7 +1371,7 @@ Lemma skipn_add_tl_In : forall n l (x a : A), In x (skipn n l) -> In x (skipn n Proof using . induction n; intros l x a Hin; simpl in *. - rewrite in_app_iff. now left. -- destruct l as [| b l]; simpl in *; solve [elim Hin | auto]. +- destruct l as [| b l]; simpl in *; solve [contradiction Hin | auto]. Qed. Lemma In_skipn_add : forall l (pt : A), In pt (skipn (Nat.div2 (length l)) (l ++ pt :: nil)). @@ -1402,7 +1413,7 @@ Proof using . intros B f k. induction k; intros [| x l]; simpl; now try rewrite Definition half1 (l : list A) := firstn (Nat.div2 (length l)) l. Definition half2 (l : list A) := skipn (Nat.div2 (length l)) l. -Lemma half1_length : forall l : list A, length (half1 l) = div2 (length l). +Lemma half1_length : forall l : list A, length (half1 l) = Nat.div2 (length l). Proof using . intros. unfold half1. @@ -1412,10 +1423,10 @@ Proof using . lia. Qed. -Corollary half2_length : forall l, length (half2 l) = length l - div2 (length l). +Corollary half2_length : forall l, length (half2 l) = length l - Nat.div2 (length l). Proof using . intros. unfold half2. now rewrite skipn_length. Qed. -Corollary half2_even_length : forall l, Nat.Even (length l) -> length (half2 l) = div2 (length l). +Corollary half2_even_length : forall l, Nat.Even (length l) -> length (half2 l) = Nat.div2 (length l). Proof using . intros l H. unfold half2. rewrite skipn_length. apply even_div2 in H. lia. Qed. Lemma merge_halves : forall l : list A, half1 l ++ half2 l = l. @@ -1644,7 +1655,7 @@ Lemma max_list_ex : forall f l, l <> nil -> exists pt, InA equiv pt l /\ max_list f l = Rmax (f pt) 0. Proof using . intros f l Hl. induction l as [| pt l]. -* now elim Hl. +* now contradiction Hl. * destruct (nil_or_In_dec l) as [? | Hin]. + subst l. exists pt. split; eauto; now left. + assert (Hnil : l <> nil). { intro. subst. destruct Hin as [? []]. } @@ -1767,7 +1778,7 @@ destruct (mem eq_dec x l) eqn:Hmem. Qed. Lemma odd_middle : forall l (d : A), Nat.Odd (length l) -> - nth (div2 (length l)) (rev l) d = nth (div2 (length l)) l d. + nth (Nat.div2 (length l)) (rev l) d = nth (Nat.div2 (length l)) l d. Proof using . intros l d. generalize (eq_refl (length l)). generalize (length l) at 2 3 4 5. intro n. revert l. induction n using nat_ind2; intros l Hl [m Hm]. @@ -1780,7 +1791,7 @@ induction n using nat_ind2; intros l Hl [m Hm]. rewrite app_length, Nat.add_comm in Hlen. simpl in Hlen. apply eq_add_S in Hlen. clear Hnil. destruct n as [| n]. - clear -Hm. lia. - - assert (div2 (S n) < length l'). { rewrite Hlen. apply Nat.lt_div2. lia. } + - assert (Nat.div2 (S n) < length l'). { rewrite Hlen. apply Nat.lt_div2. lia. } repeat rewrite app_nth1; trivial. apply IHn. assumption. destruct m as [| m]. lia. exists m. lia. @@ -1885,20 +1896,20 @@ Proof using . intros. now rewrite <- app_length, partition_Permutation. Qed. Corollary filter_length : forall {A} f (l : list A), length (filter f l) = length l - length (filter (fun x => negb (f x)) l). Proof using . -intros. apply plus_minus. -rewrite <- (partition_length f), partition_filter. -simpl. apply Nat.add_comm. +intros. symmetry. apply Nat.add_sub_eq_l. +rewrite <- (partition_length f l), partition_filter. +cbn. apply Nat.add_comm. Qed. (* Definition remove_Perm_properR := remove_Perm_proper Rdec. *) -Existing Instance Permutation_length_compat. -Existing Instance Permutation_NoDup_compat. +Global Existing Instance Permutation_length_compat. +Global Existing Instance Permutation_NoDup_compat. (* Existing Instance remove_Perm_properR. *) -Existing Instance In_perm_compat. -Existing Instance InA_impl_compat. -Existing Instance InA_compat. -Existing Instance InA_perm_compat. -Existing Instance PermutationA_length. -Existing Instance fold_left_start. -Existing Instance fold_left_symmetry_PermutationA. -Existing Instance PermutationA_map. +Global Existing Instance In_perm_compat. +Global Existing Instance InA_impl_compat. +Global Existing Instance InA_compat. +Global Existing Instance InA_perm_compat. +Global Existing Instance PermutationA_length. +Global Existing Instance fold_left_start. +Global Existing Instance fold_left_symmetry_PermutationA. +Global Existing Instance PermutationA_map. diff --git a/Util/MMultiset/MMultisetExtraOps.v b/Util/MMultiset/MMultisetExtraOps.v index 74734b588f51eb4c7a4d11f6ce305bef49524d04..df83277bc5452cb24281eb631e066c11b5237166 100644 --- a/Util/MMultiset/MMultisetExtraOps.v +++ b/Util/MMultiset/MMultisetExtraOps.v @@ -14,17 +14,20 @@ Require Import SetoidList. Require Import RelationPairs. Require Import SetoidDec. +Require Import Pactole.Util.SetoidDefs. Require Import Pactole.Util.MMultiset.Preliminary. Require Import Pactole.Util.MMultiset.MMultisetInterface. Require Import Pactole.Util.MMultiset.MMultisetFacts. Set Implicit Arguments. -Existing Instance multiplicity_compat. +Global Existing Instance multiplicity_compat. (* To have relation pairs unfolded *) Arguments RelationPairs.RelProd {A} {B} RA RB _ _ /. Arguments RelationPairs.RelCompFun {A} {B} R f a a' /. +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Section MMultisetExtra. @@ -35,74 +38,74 @@ Section MMultisetExtra. Hint Rewrite is_empty_spec nfilter_spec filter_spec npartition_spec_fst npartition_spec_snd : FMsetdec. Hint Rewrite partition_spec_fst partition_spec_snd for_all_spec exists_spec : FMsetdec. Hint Unfold In : FMsetdec. - + (* Include (MMultisetExtraOps E M). *) - + (** ** Function [remove_all] and its properties **) - + (** Remove all copies of [x] from [m] *) Definition remove_all x m := remove x m[x] m. - + Lemma remove_all_same : forall x m, (remove_all x m)[x] = 0. Proof using M. intros. unfold remove_all. rewrite remove_same. lia. Qed. - + Lemma remove_all_other : forall x y m, y =/= x -> (remove_all x m)[y] = m[y]. Proof using M. intros. unfold remove_all. now rewrite remove_other. Qed. - + Lemma remove_all_spec : forall x y m, (remove_all x m)[y] = if equiv_dec y x then 0 else m[y]. Proof using M. intros. unfold remove_all. msetdec. Qed. - + Instance remove_all_compat : Proper (equiv ==> equiv ==> equiv) remove_all. Proof using M. repeat intro. apply remove_compat; msetdec. Qed. - + Instance remove_all_sub_compat : Proper (equiv ==> Subset ==> Subset) remove_all. Proof using M. repeat intro. unfold remove_all. msetdec. Qed. - + Lemma remove_all_In : forall x y m, In x (remove_all y m) <-> In x m /\ x =/= y. Proof using M. intros. unfold remove_all. rewrite remove_In. intuition. msetdec. Qed. - + Lemma remove_all_subset : forall x m, remove_all x m [<=] m. Proof using M. intros x m y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_singleton_same : forall x n, remove_all x (singleton x n) == empty. Proof using M. intros x n y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_singleton_other : forall x y n, y =/= x -> remove_all y (singleton x n) == singleton x n. Proof using M. intros x y n Hxy z. unfold remove_all. msetdec. Qed. - + Lemma remove_all_add_same : forall x n m, remove_all x (add x n m) == remove_all x m. Proof using M. intros x n m y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_add_other : forall x y n m, x =/= y -> remove_all x (add y n m) == add y n (remove_all x m). Proof using M. intros x y n m Hxy z. unfold remove_all. msetdec. Qed. - + Lemma remove_all_remove : forall x n m, remove_all x (remove x n m) == remove_all x m. Proof using M. intros x n m y. unfold remove_all. msetdec. Qed. - + Lemma remove_remove_all : forall x n m, remove x n (remove_all x m) == remove_all x m. Proof using M. intros x n m y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_remove_other : forall x y n m, x =/= y -> remove_all y (remove x n m) == remove x n (remove_all y m). Proof using M. intros x y n m Hxy z. unfold remove_all. msetdec. Qed. - + Lemma remove_all_union : forall x mâ‚ mâ‚‚, remove_all x (union mâ‚ mâ‚‚) == union (remove_all x mâ‚) (remove_all x mâ‚‚). Proof using M. intros x n m y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_inter : forall x mâ‚ mâ‚‚, remove_all x (inter mâ‚ mâ‚‚) == inter (remove_all x mâ‚) (remove_all x mâ‚‚). Proof using M. intros x mâ‚ mâ‚‚ y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_diff : forall x mâ‚ mâ‚‚, remove_all x (diff mâ‚ mâ‚‚) == diff (remove_all x mâ‚) (remove_all x mâ‚‚). Proof using M. intros x mâ‚ mâ‚‚ y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_diff2 : forall x mâ‚ mâ‚‚, remove_all x (diff mâ‚ mâ‚‚) == diff (remove_all x mâ‚) mâ‚‚. Proof using M. intros x mâ‚ mâ‚‚ y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_lub : forall x mâ‚ mâ‚‚, remove_all x (lub mâ‚ mâ‚‚) == lub (remove_all x mâ‚) (remove_all x mâ‚‚). Proof using M. intros x mâ‚ mâ‚‚ y. unfold remove_all. msetdec. Qed. - + Lemma remove_all_for_all : forall f, compatb f -> forall x m, for_all f (remove_all x m) = for_all (fun y n => if equiv_dec y x then true else f y n) m. Proof using M. @@ -117,7 +120,7 @@ Section MMultisetExtra. - intros y y' Hy ? ? ?. subst. clear -Hf Hy. destruct (equiv_dec y x), (equiv_dec y' x); try apply Hf; trivial; rewrite Hy in *; contradiction. Qed. - + Lemma remove_all_exists : forall f, compatb f -> forall x m, exists_ f (remove_all x m) = exists_ (fun y n => if equiv_dec y x then false else f y n) m. Proof using M. @@ -132,47 +135,47 @@ Section MMultisetExtra. - intros y y' Hy ? ? ?. subst. clear -Hf Hy. destruct (equiv_dec y x), (equiv_dec y' x); try apply Hf; trivial; rewrite Hy in *; contradiction. Qed. - + Lemma remove_all_nfilter : forall f, compatb f -> forall x m, nfilter f (remove_all x m) == remove_all x (nfilter f m). Proof using M. repeat intro. unfold remove_all. msetdec. rewrite 2 Nat.sub_diag. now destruct_match. Qed. - + Lemma remove_all_filter : forall f, Proper (equiv ==> Logic.eq) f -> forall x m, filter f (remove_all x m) == remove_all x (filter f m). Proof using M. repeat intro. unfold remove_all. rewrite 2 filter_nfilter; trivial. apply remove_all_nfilter. repeat intro. auto. Qed. - + Lemma remove_all_filter_false : forall f, Proper (equiv ==> Logic.eq) f -> forall x m, f x = false -> filter f (remove_all x m) == filter f m. Proof using M. intros. rewrite remove_all_filter; trivial; []. apply remove_out. rewrite filter_In; intuition; congruence. Qed. - + Lemma remove_all_npartition_fst : forall f, compatb f -> forall x m, fst (npartition f (remove_all x m)) == remove_all x (fst (npartition f m)). Proof using M. intros. rewrite 2 npartition_spec_fst; trivial; []. now apply remove_all_nfilter. Qed. - + Lemma remove_all_npartition_snd : forall f, compatb f -> forall x m, snd (npartition f (remove_all x m)) == remove_all x (snd (npartition f m)). Proof using M. intros f Hf x m. rewrite 2 npartition_spec_snd; trivial; []. apply remove_all_nfilter. repeat intro. f_equal. now apply Hf. Qed. - + Lemma remove_all_partition_fst : forall f, Proper (equiv ==> Logic.eq) f -> forall x m, fst (partition f (remove_all x m)) == remove_all x (fst (partition f m)). Proof using M. intros. rewrite 2 partition_spec_fst; trivial; []. now apply remove_all_filter. Qed. - + Lemma remove_all_partition_snd : forall f, Proper (equiv ==> Logic.eq) f -> forall x m, snd (partition f (remove_all x m)) == remove_all x (snd (partition f m)). Proof using M. intros f Hf x m. rewrite 2 partition_spec_snd; trivial; []. apply remove_all_filter. repeat intro. f_equal. now apply Hf. Qed. - + Lemma remove_all_elements : forall x m, PermutationA eq_pair (elements (remove_all x m)) (removeA (fun x y => equiv_dec (fst x) (fst y)) (x, m[x]) (elements m)). @@ -191,43 +194,43 @@ Section MMultisetExtra. + destruct_match. - rewrite <- IHl. clear IHl. split; [intros [Hxy Hin] | intro Hin]; intuition. - inv Hin; try tauto; []. elim Hxy. hnf in *. simpl in *. now transitivity (fst e). + inv Hin; try tauto; []. contradiction Hxy. hnf in *. simpl in *. now transitivity (fst e). - split; [intros [Hxy Hin] | intro Hin]. -- inv Hin; try (now left); []. right. intuition. -- inv Hin; intuition. lazymatch goal with H : _ =/= _, H1 : eq_pair _ e |- False => apply H; rewrite <- H1; intuition end. Qed. - + Lemma remove_all_support : forall x m, PermutationA equiv (support (remove_all x m)) (removeA equiv_dec x (support m)). Proof using M. intros. unfold remove_all. rewrite support_remove. destruct_match; reflexivity || lia. Qed. - + Lemma remove_all_cardinal : forall x m, cardinal (remove_all x m) = cardinal m - m[x]. Proof using M. intros. unfold remove_all. now rewrite cardinal_remove, Nat.min_id. Qed. - + Lemma remove_all_size_in : forall x m, In x m -> size (remove_all x m) = size m - 1. Proof using M. intros. unfold remove_all. rewrite size_remove; trivial; []. destruct_match; lia. Qed. - + Lemma remove_all_size_out : forall x m, ~In x m -> size (remove_all x m) = size m. Proof using M. intros. unfold remove_all. now rewrite remove_out. Qed. - + Lemma remove_all_as_filter : forall x m, remove_all x m == filter (fun y => if equiv_dec y x then false else true) m. Proof using M. intros x m y. unfold remove_all. msetdec. repeat intro. do 2 destruct_match; trivial; exfalso; - match goal with H : _ =/= _ |- _ => apply H end; now etransitivity; eauto. + match goal with H : _ =/= _ |- _ => apply H end; now etransitivity; try eassumption; eauto. Qed. - + (** ** Function [map] and its properties **) - + Definition map f m := fold (fun x n acc => add (f x) n acc) m empty. - + Section map_results. Variable f : elt -> elt. Hypothesis Hf : Proper (equiv ==> equiv) f. - + Global Instance map_compat : Proper (equiv ==> equiv) (map f). Proof using Hf M. intros mâ‚ mâ‚‚ Hm. unfold map. apply (fold_compat _ _). @@ -236,7 +239,7 @@ Section MMultisetExtra. - assumption. - reflexivity. Qed. - + Lemma map_In : forall x m, In x (map f m) <-> exists y, x == (f y) /\ In y m. Proof using Hf M. intros x m. unfold In, map. apply fold_rect. @@ -248,10 +251,10 @@ Section MMultisetExtra. intros [? [? ?]]. msetdec. - rewrite Hrec. split; intros [z [Heq ?]]; exists z; split; msetdec. Qed. - + Lemma map_empty : map f empty == empty. Proof using M f. unfold map. now rewrite fold_empty. Qed. - + Lemma map_add : forall x n m, map f (add x n m) == add (f x) n (map f m). Proof using Hf M. intros x n m y. destruct n. @@ -263,7 +266,7 @@ Section MMultisetExtra. - repeat intro. apply add_merge. - lia. Qed. - + Lemma map_spec : forall x m, (map f m)[x] = cardinal (nfilter (fun y _ => if equiv_dec (f y) x then true else false) m). Proof using Hf M. @@ -274,7 +277,7 @@ Section MMultisetExtra. + intros * Hin Hrec. rewrite map_add, nfilter_add; trivial. unfold g at 2. msetdec. rewrite cardinal_add. lia. + now rewrite map_empty, nfilter_empty, cardinal_empty, empty_spec. Qed. - + Global Instance map_sub_compat : Proper (Subset ==> Subset) (map f). Proof using Hf M. intro m. pattern m. apply ind; clear m. @@ -282,32 +285,32 @@ Section MMultisetExtra. + intros m x n Hin Hn Hrec m' Hsub y. assert (Hx : m[x] = 0). { unfold In in Hin. lia. } rewrite <- (add_remove_cancel x m' (Hsub x)). do 2 rewrite (map_add _). msetdec. - - apply Plus.plus_le_compat; trivial; []. + - apply Nat.add_le_mono; trivial; []. repeat rewrite map_spec; trivial. apply add_subset_remove in Hsub. apply cardinal_sub_compat, nfilter_sub_compat; solve [lia | repeat intro; msetdec]. - now apply Hrec, add_subset_remove. + intros ? _. rewrite map_empty. apply subset_empty_l. Qed. - + Lemma map_singleton : forall x n, map f (singleton x n) == singleton (f x) n. Proof using Hf M. intros x n y. destruct n. + repeat rewrite singleton_0. now rewrite map_empty. + unfold map. rewrite fold_singleton; repeat intro; msetdec. Qed. - + Lemma map_remove1 : forall x n m, n <= m[x] -> map f (remove x n m) == remove (f x) n (map f m). Proof using Hf M. intros x n m Hle. rewrite <- (add_remove_cancel _ _ Hle) at 2. now rewrite (map_add _), remove_add_cancel. Qed. - + Lemma map_remove2 : forall x n m, m[x] <= n -> map f (remove x n m) == remove (f x) m[x] (map f m). Proof using Hf M. intros x n m Hle. rewrite <- (add_remove_id _ _ Hle) at 3. now rewrite (map_add _), remove_add_cancel. Qed. - + Lemma fold_map_union : forall mâ‚ mâ‚‚, fold (fun x n acc => add (f x) n acc) mâ‚ mâ‚‚ == union (map f mâ‚) mâ‚‚. Proof using Hf M. intros mâ‚ mâ‚‚. apply fold_rect with (m := mâ‚). @@ -315,7 +318,7 @@ Section MMultisetExtra. + now rewrite map_empty, union_empty_l. + intros * ? ? Heq. now rewrite Heq, map_add, union_add_comm_l. Qed. - + Theorem map_union : forall mâ‚ mâ‚‚, map f (union mâ‚ mâ‚‚) == union (map f mâ‚) (map f mâ‚‚). Proof using Hf M. intros mâ‚ mâ‚‚. unfold map at 1 2. rewrite (fold_union_additive _). @@ -324,7 +327,7 @@ Section MMultisetExtra. + repeat intro. apply add_comm. + repeat intro. apply add_merge. Qed. - + Theorem map_inter : forall mâ‚ mâ‚‚, map f (inter mâ‚ mâ‚‚) [<=] inter (map f mâ‚) (map f mâ‚‚). Proof using Hf M. intros m1 m2 x. destruct (map f (inter m1 m2))[x] eqn:Hfx. @@ -334,18 +337,18 @@ Section MMultisetExtra. destruct Hin as [Hin1 Hin2]. rewrite <- Hfx, Heq. rewrite inter_spec. apply Nat.min_glb; apply map_sub_compat; apply inter_subset_l + apply inter_subset_r. Qed. - + Theorem map_lub : forall mâ‚ mâ‚‚, lub (map f mâ‚) (map f mâ‚‚) [<=] map f (lub mâ‚ mâ‚‚). Proof using Hf M. intros m1 m2 x. destruct (lub (map f m1) (map f m2))[x] eqn:Hfx. + lia. + assert (Hin : In x (lub (map f m1) (map f m2))). { rewrite lub_spec in Hfx. rewrite lub_In. unfold In. - destruct (Max.max_dec (map f m1)[x] (map f m2)[x]) as [Heq | Heq]; + destruct (Nat.max_dec (map f m1)[x] (map f m2)[x]) as [Heq | Heq]; rewrite Heq in Hfx; left + right; lia. } rewrite lub_In in Hin. rewrite <- Hfx. destruct Hin as [Hin | Hin]; rewrite map_In in Hin; destruct Hin as [y [Heq Hin]]; rewrite Heq, lub_spec; - apply Max.max_lub; apply map_sub_compat; apply lub_subset_l + apply lub_subset_r. + apply Nat.max_lub; apply map_sub_compat; apply lub_subset_l + apply lub_subset_r. Qed. Lemma map_from_elements : @@ -384,7 +387,7 @@ Section MMultisetExtra. + intros m x n Hm Hn Hrec. rewrite map_add, size_add, size_add; trivial. destruct (In_dec x m) as [Hin | Hin], (In_dec (f x) (map f m)) as [Hinf | Hinf]. - apply Hrec. - - elim Hinf. rewrite map_In. now exists x. + - contradiction Hinf. rewrite map_In. now exists x. - lia. - lia. + now rewrite map_empty. @@ -451,7 +454,7 @@ Section MMultisetExtra. (map f (lub mâ‚ mâ‚‚))[x] = (map f mâ‚)[x]. Proof using Hf Hf2 M. intros x mâ‚ mâ‚‚ Hle. destruct (map f mâ‚)[x] eqn:Heq1. - - apply Le.le_n_0_eq in Hle. symmetry in Hle. destruct (map f (lub mâ‚ mâ‚‚))[x] eqn:Heq2; trivial. + - apply Nat.le_0_r in Hle. destruct (map f (lub mâ‚ mâ‚‚))[x] eqn:Heq2; trivial. assert (Hin : In x (map f (lub mâ‚ mâ‚‚))). { unfold In. lia. } rewrite map_In in Hin. destruct Hin as [y [Heq Hin]]. rewrite Heq in *. rewrite lub_In in Hin. rewrite map_injective_spec in *. unfold In in *. destruct Hin; lia. @@ -462,7 +465,7 @@ Section MMultisetExtra. Theorem map_injective_lub : forall mâ‚ mâ‚‚, map f (lub mâ‚ mâ‚‚) == lub (map f mâ‚) (map f mâ‚‚). Proof using Hf Hf2 M. - intros mâ‚ mâ‚‚ x. rewrite lub_spec. apply Max.max_case_strong; intro Hle. + intros mâ‚ mâ‚‚ x. rewrite lub_spec. apply Nat.max_case_strong; intro Hle. - now apply map_injective_lub_wlog. - rewrite lub_comm. now apply map_injective_lub_wlog. Qed. @@ -685,7 +688,7 @@ Section MMultisetExtra. + rewrite Hxy. rewrite add_spec, Hx. msetdec. + rewrite add_other; auto. transitivity (max_mult m). - apply Hrec. - - apply Max.le_max_r. + - apply Nat.le_max_r. * intro x. rewrite empty_spec. lia. Qed. @@ -700,7 +703,7 @@ Section MMultisetExtra. + destruct (equiv_dec z x) as [Hzx | Hzx]. - rewrite Hzx. rewrite add_spec, Hx. msetdec. - rewrite add_other; trivial; []. - transitivity (max_mult m); apply Hall || apply Max.le_max_r. + transitivity (max_mult m); apply Hall || apply Nat.le_max_r. + rewrite Hy. destruct (Compare_dec.le_dec n m[y]). - exists y. rewrite max_r; trivial; []. destruct (equiv_dec y x) as [Hyx | Hyx]. @@ -752,7 +755,7 @@ Section MMultisetExtra. Proof using M. intro m. split; intro Heq. + destruct (empty_or_In_dec m) as [? | [x Hin]]; trivial. - elim (Nat.lt_irrefl 0). apply Nat.lt_le_trans with m[x]. + contradiction (Nat.lt_irrefl 0). apply Nat.lt_le_trans with m[x]. - exact Hin. - rewrite <- Heq. apply max_mult_spec_weak. + rewrite Heq. apply max_mult_empty. @@ -793,7 +796,7 @@ Section MMultisetExtra. -- assert (p = 0) by lia. subst. reflexivity. -- repeat rewrite max_r; trivial; lia. - rewrite <- Hle, Nat.add_comm. apply Nat.le_max_l. - - apply Max.max_lub_r in Hle. rewrite max_l; lia. + - apply Nat.max_lub_r in Hle. rewrite max_l; lia. - rewrite <- Hle. apply Nat.max_le_compat; lia. + rewrite add_other in *; trivial; []. rewrite add_comm, max_mult_add; try (now rewrite add_In; intuition); []. @@ -846,7 +849,7 @@ Section MMultisetExtra. Nat.max (max_mult mâ‚) (max_mult mâ‚‚) <= max_mult (union mâ‚ mâ‚‚) <= max_mult mâ‚ + max_mult mâ‚‚. Proof using M. intros mâ‚ mâ‚‚. split. - + apply Max.max_lub; f_equiv; intro; msetdec. + + apply Nat.max_lub; f_equiv; intro; msetdec. + apply max_mult_upper_bound. intro. msetdec. apply Nat.add_le_mono; apply max_mult_spec_weak. Qed. @@ -909,7 +912,7 @@ Section MMultisetExtra. match Nat.compare n (fst acc) with | Lt => acc | Eq => (fst acc, add x n (snd acc)) - | gt => (n, singleton x n) + | Gt => (n, singleton x n) end. Definition max m := snd (fold max_aux m (0, empty)). @@ -1032,7 +1035,7 @@ Section MMultisetExtra. - rewrite Nat.eqb_eq, max_mult_add; trivial. rewrite Hm at 2. rewrite add_empty, singleton_spec. - msetdec. rewrite max_mult_empty. apply Max.max_0_r. + msetdec. rewrite max_mult_empty. apply Nat.max_0_r. + specialize (HI x'' Hx''). destruct HI as [y Hy]. unfold max. setoid_rewrite nfilter_In; auto; []. @@ -1042,12 +1045,12 @@ Section MMultisetExtra. destruct (Compare_dec.le_lt_dec n (m[y])). - exists y. split. -- msetdec. - -- rewrite Nat.eqb_eq, Heq, add_other, Max.max_r; trivial. + -- rewrite Nat.eqb_eq, Heq, add_other, Nat.max_r; trivial. Fail now msetdec. (* BUG? *) intro Habs. msetdec. - exists x. split. -- msetdec. - -- rewrite Nat.eqb_eq, Max.max_l; try lia. msetdec. - * intros x Hin. elim (In_empty Hin). + -- rewrite Nat.eqb_eq, Nat.max_l; try lia. msetdec. + * intros x Hin. contradiction (In_empty Hin). Qed. Lemma max_is_empty : forall m, max m == empty <-> m == empty. @@ -1133,7 +1136,7 @@ Section MMultisetExtra. split. - red. cut (m[x]<>0). lia. intro Habs. now rewrite Hx, max_mult_0 in Habs. - - now rewrite Hx, <- EqNat.beq_nat_refl. + - rewrite Hx. apply Nat.eqb_refl. Qed. Lemma max_max_mult_ex : forall m, ~m == empty -> exists x, max_mult m = m[x]. @@ -1145,10 +1148,10 @@ Section MMultisetExtra. + exists x. rewrite Hm, add_empty. rewrite max_mult_singleton. msetdec. + assert (Hempty : m =/= empty) by now rewrite not_empty_In. destruct (Hrec Hempty) as [max_m Hmax_m]. rewrite max_mult_add; trivial. - destruct (Max.max_spec n (max_mult m)) as [[Hmax1 Hmax2] | [Hmax1 Hmax2]]. + destruct (Nat.max_spec n (max_mult m)) as [[Hmax1 Hmax2] | [Hmax1 Hmax2]]. - exists max_m. msetdec. - exists x. msetdec. - * intro Habs. Fail now msetdec. (* BUG? *) now elim Habs. + * intro Habs. Fail now msetdec. (* BUG? *) now contradiction Habs. Qed. Lemma max_In_mult3 : forall m x, In x m -> (In x (max m) <-> m[x] = max_mult m). @@ -1302,7 +1305,7 @@ Section MMultisetExtra. forall m (i : A), eqA (fold f (max m) i) (fold (fun x n acc => if n =? max_mult m then f x n acc else acc) m i). Proof using M. intros A eqA HeqA f Hf Hf2 m i. - rewrite fold_compat; autoclass; try apply max_simplified; try reflexivity; []. + rewrite fold_compat; try eassumption; try apply max_simplified; try reflexivity; []. unfold simple_max. rewrite fold_nfilter; autoclass; []. apply fold_extensionality_compat; autoclass. - repeat intro. subst. now destruct_match; try apply Hf. diff --git a/Util/MMultiset/MMultisetFacts.v b/Util/MMultiset/MMultisetFacts.v index f74d46a808334ec6de3a9c8c8e81debb84f3468a..1ebe82158e024d826f3ff5d01e70f3004c7d281b 100644 --- a/Util/MMultiset/MMultisetFacts.v +++ b/Util/MMultiset/MMultisetFacts.v @@ -13,9 +13,9 @@ Require Import Bool. Require Import Lia PeanoNat. Require Import PArith. Require Import RelationPairs. -(* Require Import Equalities. *) Require Import SetoidList. Require Import SetoidDec. +Require Import Pactole.Util.SetoidDefs. Require Import Pactole.Util.MMultiset.Preliminary. Require Import Pactole.Util.MMultiset.MMultisetInterface. @@ -24,6 +24,8 @@ Set Implicit Arguments. Notation " x == y " := (equiv x y) (at level 70, no associativity). Notation " x =/= y " := (complement equiv x y) (at level 70, no associativity). +(* this will become non default soon. *) +Ltac Tauto.intuition_solver ::= auto with *. Section MMultisetFacts. @@ -78,7 +80,7 @@ Section MMultisetFacts. | H : ?x = ?x |- _ => clear H | H : ?x == ?x |- _ => clear H | H : ?x = ?y |- _ => subst x || rewrite H in * - | Hneq : ?x =/= ?x |- _ => now elim Hneq + | Hneq : ?x =/= ?x |- _ => now contradiction Hneq | Heq : @equiv elt _ ?x ?y |- _ => clear x Heq || rewrite Heq in * | Heq : @equiv (multiset _) _ ?x ?y, Hin : context[?x] |- _ => rewrite Heq in Hin | Heq : @equiv (multiset _) _ ?x ?y |- context[?x] => rewrite Heq @@ -120,7 +122,7 @@ Section MMultisetFacts. Lemma InA_pair_elt : forall x n p l, InA eq_pair (x, n) l -> InA eq_elt (x, p) l. Proof using . intros x n p l Hin. induction l as [| [y q] l]. - + rewrite InA_nil in Hin. elim Hin. + + rewrite InA_nil in Hin. contradiction Hin. + inversion_clear Hin. - destruct H as [Heq ?]. now left. - right. now apply IHl. @@ -129,7 +131,7 @@ Section MMultisetFacts. Lemma InA_elt_pair : forall x n l, InA eq_elt (x, n) l -> exists n', InA eq_pair (x, n') l. Proof using . intros x n l Hin. induction l as [| [y p] l]. - + rewrite InA_nil in Hin. elim Hin. + + rewrite InA_nil in Hin. contradiction Hin. + inversion_clear Hin. - compute in H. exists p. left. now rewrite H. - apply IHl in H. destruct H as [k Hin]. exists k. now right. @@ -304,7 +306,7 @@ Section MMultisetFacts. Proof using FMultisetsSpec. intros ? y ? ? ? Hs H. msetdec. specialize (Hs y). lia. Qed. Global Instance add_sub_compat : Proper (equiv ==> le ==> Subset ==> Subset) add. - Proof using FMultisetsSpec. repeat intro. msetdec. now apply Plus.plus_le_compat. Qed. + Proof using FMultisetsSpec. repeat intro. msetdec. now apply Nat.add_le_mono. Qed. Global Instance singleton_sub_compat : Proper (equiv ==> le ==> Subset) singleton. Proof using FMultisetsSpec. repeat intro. msetdec. Qed. @@ -536,7 +538,7 @@ Section MMultisetFacts. Proof using FMultisetsSpec. repeat intro. msetdec. Qed. Lemma remove_singleton_other : forall x y n p, ~y == x -> remove y n (singleton x p) == singleton x p. - Proof using FMultisetsSpec. repeat intro. msetdec. (* BUG?: saturate_Einequalities should work! *) now elim H. Qed. + Proof using FMultisetsSpec. repeat intro. msetdec. (* BUG?: saturate_Einequalities should work! *) now contradiction H. Qed. Theorem elements_singleton : forall x n, n > 0 -> eqlistA eq_pair (elements (singleton x n)) ((x, n) :: nil). Proof using FMultisetsSpec. @@ -744,7 +746,7 @@ Section MMultisetFacts. Lemma add_is_singleton : forall x y n p m, add x n m == singleton y p -> m == singleton y (p - n). Proof using FMultisetsSpec. intros x y n p m Hadd z. destruct n. - + rewrite add_0 in Hadd. now rewrite Hadd, <- Minus.minus_n_O. + + rewrite add_0 in Hadd. now rewrite Hadd, Nat.sub_0_r. + assert (Heq := Hadd x). msetdec. rewrite <- (add_other y z (S n)), Hadd; trivial. msetdec. Qed. @@ -987,7 +989,7 @@ Section MMultisetFacts. intros m1 m2. split; intro Hin. + intro x. destruct (In_dec x m1) as [Hin1 | Hin1], (In_dec x m2) as [Hin2 | Hin2]; auto. assert (Habs : In x (inter m1 m2)). { rewrite inter_In. auto. } - rewrite Hin in Habs. apply In_empty in Habs. elim Habs. + rewrite Hin in Habs. apply In_empty in Habs. contradiction Habs. + intro x. rewrite empty_spec, inter_spec. destruct (Hin x) as [[Hin1 Hin2] | [[Hin1 Hin2] | [Hin1 Hin2]]]; rewrite not_In in *; try rewrite Hin1; try rewrite Hin2; auto with arith. Qed. @@ -1195,7 +1197,7 @@ Section MMultisetFacts. - specialize (Heq y). msetdec. - specialize (Heq x). msetdec. + destruct Heq as [Hm1 [Hm2 Hn]]. rewrite Hm1, Hm2, lub_singleton_l, add_singleton_same. f_equiv. subst. - rewrite singleton_same. apply Max.max_case_strong; lia. + rewrite singleton_same. apply Nat.max_case_strong; lia. Qed. Lemma remove_lub : forall x n m1 m2, remove x n (lub m1 m2) == lub (remove x n m1) (remove x n m2). @@ -1275,7 +1277,7 @@ Section MMultisetFacts. { intros n Habs. rewrite elements_spec in Habs. destruct Habs. simpl in *. lia. } destruct (mâ‚[x]) eqn:Hmâ‚. reflexivity. specialize (Hin (S n)). rewrite <- Heq in Hin. rewrite elements_spec in Hin. - elim Hin. split; simpl. assumption. lia. + contradiction Hin. split; simpl. assumption. lia. - assert (Hin : InA eq_pair (x, S n) (elements mâ‚‚)). { rewrite elements_spec. split; simpl. assumption. lia. } rewrite <- Heq in Hin. rewrite elements_spec in Hin. now destruct Hin. + intros [x n]. now rewrite Heq. @@ -1398,7 +1400,7 @@ Section MMultisetFacts. + apply removeA_NoDupA; refine _. apply (NoDupA_strengthen _ (elements_NoDupA _)). + intros [y p]. rewrite removeA_InA_iff; refine _. rewrite elements_remove, elements_spec. simpl. intuition. - destruct H1. contradiction. - - destruct (equiv_dec y x) as [Heq | Heq]; auto. elim H1. split; msetdec. + - destruct (equiv_dec y x) as [Heq | Heq]; auto. contradiction H1. split; msetdec. Qed. Lemma elements_remove2 : forall x n m, n < m[x] -> @@ -1409,7 +1411,7 @@ Section MMultisetFacts. + apply (NoDupA_strengthen _ (elements_NoDupA _)). + constructor. - intro Habs. eapply InA_pair_elt in Habs. rewrite removeA_InA_iff in Habs; refine _. - destruct Habs as [_ Habs]. now elim Habs. + destruct Habs as [_ Habs]. now contradiction Habs. - eapply (NoDupA_strengthen subrelation_pair_elt). apply removeA_NoDupA, elements_NoDupA; refine _. + intros [y p]. rewrite elements_remove, elements_spec. simpl. intuition. - rewrite H. left. split. compute. reflexivity. assumption. @@ -1419,9 +1421,9 @@ Section MMultisetFacts. + inversion_clear H. - left. destruct H0. repeat split; auto. hnf in *. simpl in *. lia. - apply (InA_pair_elt (m[x])) in H0. rewrite Heq, removeA_InA in H0; refine _. - destruct H0 as [_ Habs]. elim Habs. reflexivity. + destruct H0 as [_ Habs]. contradiction Habs. reflexivity. + right. split; trivial. inversion_clear H. - - elim Heq. destruct H0. auto. + - contradiction Heq. destruct H0. auto. - apply removeA_InA_weak in H0. rewrite elements_spec in H0. assumption. } Qed. @@ -1559,7 +1561,7 @@ Section MMultisetFacts. + apply add_is_singleton in Hin. specialize (Hin z). msetdec. destruct Hl as [_ Hl]. inversion_clear Hl. inversion_clear H0. simpl in *. lia. } destruct Heq as [Heq1 Heq2]. destruct Hl as [Hl _]. inversion_clear Hl. - elim H. left. compute. now transitivity x. + contradiction H. left. compute. now transitivity x. - inversion_clear Hin. inversion_clear H0. Qed. @@ -1575,7 +1577,7 @@ Section MMultisetFacts. Lemma from_elements_in : forall x n l, NoDupA eq_elt l -> InA eq_pair (x, n) l -> (from_elements l)[x] = n. Proof using FMultisetsSpec. intros x n l Hl Hin. induction l as [| [y p] l]. - + rewrite InA_nil in Hin. elim Hin. + + rewrite InA_nil in Hin. contradiction Hin. + simpl. inversion_clear Hin. - destruct H as [Hx Hn]. compute in Hx, Hn. inversion Hl. now rewrite Hx, add_same, (@from_elements_out y p). - inversion_clear Hl. rewrite add_other. now apply IHl. @@ -1617,8 +1619,8 @@ Section MMultisetFacts. Proof using FMultisetsSpec. intros l x. induction l as [| [y p] l]. * simpl. split; intro Hin. - + elim (In_empty Hin). - + destruct Hin as [? [Hin _]]. rewrite InA_nil in Hin. elim Hin. + + contradiction (In_empty Hin). + + destruct Hin as [? [Hin _]]. rewrite InA_nil in Hin. contradiction Hin. * simpl. rewrite add_In, IHl; trivial. split; intros Hin. + destruct Hin as [[? Heq] | [n [Hin Hn]]]. - exists p. split; try (left; split); auto; lia. @@ -1653,10 +1655,10 @@ Section MMultisetFacts. - subst. now repeat left. - inversion_clear Hp. -- destruct H; auto. - -- inversion_clear Hnodup. elim H0. revert H. apply InA_pair_elt. + -- inversion_clear Hnodup. contradiction H0. revert H. apply InA_pair_elt. + rewrite add_other; trivial. destruct l as [| z l]. - simpl. rewrite empty_spec. intuition; try lia. - inversion_clear H. destruct H0; try contradiction. rewrite InA_nil in H0. elim H0. + inversion_clear H. destruct H0; try contradiction. rewrite InA_nil in H0. contradiction H0. - inversion_clear Hnodup. rewrite IHl; discriminate || trivial. intuition. inversion_clear H1; trivial. destruct H2. contradiction. Qed. @@ -1686,7 +1688,7 @@ Section MMultisetFacts. rewrite from_elements_cons, add_merge. rewrite elements_add_out. - constructor; try reflexivity. apply is_elements_cons_inv in Hl'. rewrite Hin, elements_from_elements; trivial. simpl. - destruct pair_dec as [? | Habs]; try now elim Habs. + destruct pair_dec as [? | Habs]; try now contradiction Habs. rewrite removeA_out; try reflexivity. intro Habs. apply Hout. revert Habs. apply InA_pair_elt. - apply proj2 in Hl'. inversion_clear Hl'. simpl in *. lia. - apply is_elements_cons_inv in Hl'. rewrite <- elements_In, elements_from_elements; eauto. @@ -1781,7 +1783,7 @@ Section MMultisetFacts. - rewrite <- fold_left_rev_right. rewrite rev_unit. simpl. rewrite <- Hfadd. f_equiv. rewrite fold_left_rev_right, <- fold_spec. etransitivity. symmetry. apply fold_add. lia. unfold In. rewrite remove_same. lia. - rewrite add_remove1; trivial. now rewrite Minus.minus_diag, add_0. + rewrite add_remove1; trivial. now rewrite Nat.sub_diag, add_0. - intros ? ? ? [? ?] [? ?] [Heq ?]. now apply Hf. - intros [? ?] [? ?] ?. simpl. apply Hf2. Qed. @@ -1923,7 +1925,7 @@ Section MMultisetFacts. + pattern m. apply ind; clear m. - intros m1 m2 Hm. setoid_rewrite Hm. reflexivity. - intros m x n Hm Hn Hrec _. exists x. apply add_In. left. split; lia || reflexivity. - - intro Habs. now elim Habs. + - intro Habs. now contradiction Habs. + intros [x Hin]. intro Habs. revert Hin. rewrite Habs. apply In_empty. Qed. @@ -1989,7 +1991,7 @@ Section MMultisetFacts. destruct (equiv_dec z x). exfalso. revert Hin. msetdec. split; msetdec. - destruct Hin. msetdec. (* BUG?: saturate_Einequalities shou work! *) now elim H0. + destruct Hin. msetdec. (* BUG?: saturate_Einequalities shou work! *) now contradiction H0. - do 2 rewrite support_spec. unfold In in *. msetdec. Qed. @@ -2068,7 +2070,7 @@ Section MMultisetFacts. intro m. split; intro Hm. + intro y. rewrite empty_spec, <- empty_spec with y. revert y. change (m == empty). rewrite <- elements_nil. destruct (elements m) as [| [x n] l] eqn:Helt. reflexivity. - simpl in Hm. elim (Nat.lt_irrefl 0). apply Nat.lt_le_trans with n. + simpl in Hm. contradiction (Nat.lt_irrefl 0). apply Nat.lt_le_trans with n. - apply elements_pos with x m. rewrite Helt. now left. - assert (Hn : m[x] = n). { eapply proj1. rewrite <- (elements_spec (x, n)), Helt. now left. } rewrite <- Hn, <- Hm. apply cardinal_lower. @@ -2104,10 +2106,10 @@ Section MMultisetFacts. Theorem cardinal_remove : forall x n m, cardinal (remove x n m) = cardinal m - min n (m[x]). Proof using FMultisetsSpec. intros x n m. destruct (Compare_dec.le_dec n (m[x])) as [Hle | Hlt]. - + setoid_rewrite <- (add_0 x) at 3. erewrite <- (Minus.minus_diag n). + + setoid_rewrite <- (add_0 x) at 3. erewrite <- (Nat.sub_diag n). rewrite <- (@add_remove1 x n n m), cardinal_add, min_l; trivial. lia. + assert (Hle : m[x] <= n) by lia. - setoid_rewrite <- (add_0 x) at 3. erewrite <- Minus.minus_diag. + setoid_rewrite <- (add_0 x) at 3. erewrite <- Nat.sub_diag. rewrite <- (@add_remove2 x _ n m Hle (Nat.le_refl _)), cardinal_add, min_r; trivial. lia. Qed. @@ -2119,7 +2121,7 @@ Section MMultisetFacts. assert (n <= m'[x]). { transitivity (n + m[x]). lia. specialize (Hsub x). msetdec. } assert (n <= cardinal m'). { etransitivity; try eassumption. apply cardinal_lower. } apply add_subset_remove in Hsub. apply Hrec in Hsub. rewrite cardinal_remove in Hsub. - etransitivity. apply Plus.plus_le_compat. reflexivity. apply Hsub. rewrite min_l; trivial. lia. + etransitivity. apply Nat.add_le_mono. reflexivity. apply Hsub. rewrite min_l; trivial. lia. + intros. rewrite cardinal_empty. lia. Qed. @@ -2129,11 +2131,11 @@ Section MMultisetFacts. * intros m1 m1' Heq. split; intros Hle m2; rewrite Heq || rewrite <- Heq; apply Hle. * intros m x n Hout Hn Hind m2. destruct (Compare_dec.le_lt_dec n (m2[x])) as [Hle | Hlt]. + rewrite inter_add_l1; trivial. rewrite <- (add_remove_cancel Hle) at 2. - do 3 rewrite cardinal_add. rewrite Min.plus_min_distr_l. apply Plus.plus_le_compat_l, Hind. + do 3 rewrite cardinal_add. rewrite Nat.add_min_distr_l. apply Nat.add_le_mono_l, Hind. + rewrite inter_add_l2; try lia. transitivity (Init.Nat.min (cardinal (add x (m2[x]) m)) (cardinal m2)). - rewrite <- (add_remove_cancel (reflexivity (m2[x]))) at 4. - do 3 rewrite cardinal_add. rewrite Min.plus_min_distr_l. apply Plus.plus_le_compat_l. + do 3 rewrite cardinal_add. rewrite Nat.add_min_distr_l. apply Nat.add_le_mono_l. rewrite remove_cap; try lia. apply Hind. - do 2 rewrite cardinal_add. apply Nat.min_le_compat_r. lia. * intro. rewrite inter_empty_l, cardinal_empty. lia. @@ -2151,9 +2153,9 @@ Section MMultisetFacts. replace (n + cardinal m - (n + cardinal(remove x n m2))) with (cardinal m - cardinal(remove x n m2)) by lia. apply Hind. - rewrite diff_add_l2; try lia. rewrite <- (add_remove_cancel (reflexivity (m2[x]))) at 1. - do 3 rewrite cardinal_add. rewrite <- (@remove_cap x n); try lia. - transitivity ((n - m2[x]) + (cardinal m - cardinal(remove x n m2))); try lia. - apply Plus.plus_le_compat_l, Hind. + do 3 rewrite cardinal_add. rewrite <- (@remove_cap x n); try lia; []. + transitivity ((n - m2[x]) + (cardinal m - cardinal(remove x n m2))); try lia; []. + apply Nat.add_le_mono_l, Hind. + intro. now rewrite diff_empty_l, cardinal_empty. Qed. @@ -2166,9 +2168,9 @@ Section MMultisetFacts. + intros m1 m1' Heq. now setoid_rewrite Heq. + intros m x n Hout Hn Hind m2. rewrite lub_add_l. do 2 rewrite cardinal_add. transitivity (n + Init.Nat.max (cardinal m) (cardinal (remove x n m2))). - - rewrite <- Max.plus_max_distr_l. apply Nat.max_le_compat_l. rewrite <- (cardinal_add x). + - rewrite <- Nat.add_max_distr_l. apply Nat.max_le_compat_l. rewrite <- (cardinal_add x). apply cardinal_sub_compat. intro. msetdec. - - apply Plus.plus_le_compat_l, Hind. + - apply Nat.add_le_mono_l, Hind. + intro. now rewrite lub_empty_l, cardinal_empty. Qed. @@ -2178,7 +2180,7 @@ Section MMultisetFacts. Lemma cardinal_from_elements : forall l, cardinal (from_elements l) = List.fold_left (fun acc xn => snd xn + acc) l 0. Proof using FMultisetsSpec. - intro l. rewrite <- Plus.plus_0_l at 1. generalize 0. induction l as [| [x n] l]; intro p; simpl. + intro l. rewrite <- Nat.add_0_l at 1. generalize 0. induction l as [| [x n] l]; intro p; simpl. - now rewrite cardinal_empty. - rewrite cardinal_add, Nat.add_assoc. rewrite (Nat.add_comm p n). apply IHl. Qed. @@ -2220,7 +2222,7 @@ Section MMultisetFacts. - inversion Hin. - exists x. rewrite <- support_spec, Heq. now left. + destruct Hin as [x Hin]. destruct (size m) eqn:Hsize. - - rewrite size_0 in Hsize. rewrite Hsize in Hin. elim (In_empty Hin). + - rewrite size_0 in Hsize. rewrite Hsize in Hin. contradiction (In_empty Hin). - auto with arith. Qed. @@ -2238,7 +2240,7 @@ Section MMultisetFacts. assert (Hnodup : NoDupA equiv (x :: l)). { rewrite <- Hin. apply support_NoDupA. } (* XXX: why does [rewrite Hin] fails here? *) rewrite removeA_Perm_compat; eauto; try reflexivity || apply setoid_equiv; []. rewrite Hin. - simpl. destruct (equiv_dec x x) as [_ | Hneq]; try now elim Hneq. + simpl. destruct (equiv_dec x x) as [_ | Hneq]; try now contradiction Hneq. inversion_clear Hnodup. now rewrite removeA_out. Qed. @@ -2248,7 +2250,7 @@ Section MMultisetFacts. Qed. Lemma size_union_lower : forall m1 m2, max (size m1) (size m2) <= size (union m1 m2). - Proof using FMultisetsSpec. intros. apply Max.max_case; apply size_sub_compat; apply union_subset_l || apply union_subset_r. Qed. + Proof using FMultisetsSpec. intros. apply Nat.max_case; apply size_sub_compat; apply union_subset_l || apply union_subset_r. Qed. (* TODO?: the most straigthforward way to express this would be by using set_union, hence requiring ListSetA. *) Lemma size_union_upper : forall m1 m2, size (union m1 m2) <= size m1 + size m2. @@ -2261,13 +2263,13 @@ Section MMultisetFacts. Qed. Lemma size_inter_upper : forall m1 m2, size (inter m1 m2) <= min (size m1) (size m2). - Proof using FMultisetsSpec. intros. apply Min.min_case; apply size_sub_compat; apply inter_subset_l || apply inter_subset_r. Qed. + Proof using FMultisetsSpec. intros. apply Nat.min_case; apply size_sub_compat; apply inter_subset_l || apply inter_subset_r. Qed. Lemma size_diff_upper : forall m1 m2, size (diff m1 m2) <= size m1. Proof using FMultisetsSpec. intros. apply size_sub_compat, diff_subset. Qed. Lemma size_lub_lower : forall m1 m2, max (size m1) (size m2) <= size (lub m1 m2). - Proof using FMultisetsSpec. intros. apply Max.max_case; apply size_sub_compat; apply lub_subset_l || apply lub_subset_r. Qed. + Proof using FMultisetsSpec. intros. apply Nat.max_case; apply size_sub_compat; apply lub_subset_l || apply lub_subset_r. Qed. Lemma size_lub_upper : forall m1 m2, size (lub m1 m2) <= size m1 + size m2. Proof using FMultisetsSpec. @@ -3488,7 +3490,7 @@ Section MMultisetFacts. - right. exists y. now split. + symmetry. rewrite orb_false_iff. rewrite exists_false in *; trivial. assert (Hxm : m[x] = 0) by (unfold In in Hin; lia). split. - - destruct (f x n) eqn:Hfxn; trivial. elim Hm. exists x. split; msetdec. + - destruct (f x n) eqn:Hfxn; trivial. contradiction Hm. exists x. split; msetdec. - intros [y [Hy Hfy]]. apply Hm. exists y. unfold In in *. split; msetdec. Qed. @@ -3502,7 +3504,7 @@ Section MMultisetFacts. rewrite <- (@add_remove_cancel x), exists_add; trivial. - apply (Hf2 _ _ (reflexivity x)) in Hle. simpl in Hle. rewrite Hall in Hle. simpl in Hle. now rewrite Hle. - lia. - - rewrite remove_In. intros [[_ Habs] | [Habs _]]; lia || now elim Habs. + - rewrite remove_In. intros [[_ Habs] | [Habs _]]; lia || now contradiction Habs. + setoid_rewrite Hall in Hrec. simpl in Hrec. apply Hrec. etransitivity; try eassumption. apply add_subset. * intros. rewrite exists_empty; trivial. intuition. Qed. @@ -3529,14 +3531,14 @@ Section MMultisetFacts. Proof using FMultisetsSpec Hf. intros m1 m2. repeat rewrite exists_spec; trivial. intros [x [Hin Hfx]]. rewrite inter_spec in Hfx. rewrite inter_In in Hin. destruct Hin. - destruct (Min.min_dec (m1[x]) (m2[x])) as [Hmin | Hmin]; + destruct (Nat.min_dec (m1[x]) (m2[x])) as [Hmin | Hmin]; rewrite Hmin in Hfx; left + right; now exists x. Qed. Lemma exists_lub : forall m1 m2, exists_ f (lub m1 m2) = true -> exists_ f m1 = true \/ exists_ f m2 = true. Proof using FMultisetsSpec Hf. intros m1 m2. repeat rewrite exists_spec; trivial. intros [x [Hin Hfx]]. unfold In in *. - rewrite lub_spec in Hin, Hfx. destruct (Max.max_dec (m1[x]) (m2[x])) as [Hmax | Hmax]; + rewrite lub_spec in Hin, Hfx. destruct (Nat.max_dec (m1[x]) (m2[x])) as [Hmax | Hmax]; rewrite Hmax in Hin, Hfx; left + right; now exists x. Qed. @@ -3569,7 +3571,7 @@ Section MMultisetFacts. - exists x. split. ++ rewrite add_In. left. split; lia || reflexivity. ++ rewrite not_In in Hm. rewrite add_same, Hm. simpl. now rewrite Hfxn. - + intro Habs. elim Habs. intros x Hin. elim (In_empty Hin). + + intro Habs. contradiction Habs. intros x Hin. contradiction (In_empty Hin). * intro Habs. destruct Hm as [x [Hin Hx]]. apply Habs in Hin. rewrite Hin in Hx. discriminate. Qed. @@ -3612,7 +3614,7 @@ Section MMultisetFacts. + destruct (for_all (fun (x : elt) (n : nat) => negb (f x n)) m) eqn:Hforall; trivial. rewrite for_all_false_exists, exists_spec in Hforall; trivial. destruct Hforall as [x [Hin Hfx]]. rewrite negb_involutive in Hfx. - elim (@In_empty x). rewrite <- Hall, nfilter_In; auto. + contradiction (@In_empty x). rewrite <- Hall, nfilter_In; auto. + rewrite for_all_spec in Hall; trivial. destruct (empty_or_In_dec (nfilter f m)) as [? | [x Hin]]; trivial. rewrite nfilter_In in Hin; trivial. destruct Hin as [Hin Hfx]. apply Hall in Hin. @@ -3704,7 +3706,7 @@ Ltac msetdec_step := | H : ?x = ?x |- _ => clear H | H : ?x == ?x |- _ => clear H | H : ?x = ?y |- _ => subst x || rewrite H in * - | Hneq : ?x =/= ?x |- _ => now elim Hneq + | Hneq : ?x =/= ?x |- _ => now contradiction Hneq | Heq : equiv ?x ?y |- _ => clear x Heq || rewrite Heq in * | Heq : @equiv (multiset _) _ ?x ?y, Hin : context[?x] |- _ => rewrite Heq in Hin | Heq : @equiv (multiset _) _ ?x ?y |- context[?x] => rewrite Heq diff --git a/Util/MMultiset/MMultisetInterface.v b/Util/MMultiset/MMultisetInterface.v index 32734f148ee0e9c3d06054a92e5d521fbc92c5ec..de832e4136f1b8697068ac1f1ce35d0e0619df4a 100644 --- a/Util/MMultiset/MMultisetInterface.v +++ b/Util/MMultiset/MMultisetInterface.v @@ -230,22 +230,22 @@ Class SizeSpec elt `(FMOps elt) := { (** *** Full specification **) Class FMultisetsOn elt `(Ops : FMOps elt) := { - FullMultiplicitySpec :> MultiplicitySpec elt Ops; - FullEmptySpec :> EmptySpec elt Ops; - FullSingletonSpec :> SingletonSpec elt Ops; - FullAddSpec :> AddSpec elt Ops; - FullRemoveSpec :> RemoveSpec elt Ops; - FullBinarySpec :> BinarySpec elt Ops; - FullFoldSpec :> FoldSpec elt Ops; - FullTestSpec: > TestSpec elt Ops; - FullElementsSpec :> ElementsSpec elt Ops; - FullSupportSpec :> SupportSpec elt Ops; - FullChooseSpec :> ChooseSpec elt Ops; - FullPartitionSpec :> PartitionSpec elt Ops; - FullNpartitionSpec :> NpartitionSpec elt Ops; - FullQuantifierSpec :> QuantifierSpec elt Ops; - FullSizeSpec :> SizeSpec elt Ops; - FullFilterSpec :> FilterSpec elt Ops}. + #[global] FullMultiplicitySpec :: MultiplicitySpec elt Ops; + #[global] FullEmptySpec :: EmptySpec elt Ops; + #[global] FullSingletonSpec :: SingletonSpec elt Ops; + #[global] FullAddSpec :: AddSpec elt Ops; + #[global] FullRemoveSpec :: RemoveSpec elt Ops; + #[global] FullBinarySpec :: BinarySpec elt Ops; + #[global] FullFoldSpec :: FoldSpec elt Ops; + #[global] FullTestSpec :: TestSpec elt Ops; + #[global] FullElementsSpec :: ElementsSpec elt Ops; + #[global] FullSupportSpec :: SupportSpec elt Ops; + #[global] FullChooseSpec :: ChooseSpec elt Ops; + #[global] FullPartitionSpec :: PartitionSpec elt Ops; + #[global] FullNpartitionSpec :: NpartitionSpec elt Ops; + #[global] FullQuantifierSpec :: QuantifierSpec elt Ops; + #[global] FullSizeSpec :: SizeSpec elt Ops; + #[global] FullFilterSpec :: FilterSpec elt Ops}. (* Global Notation "s [=] t" := (eq s t) (at level 70, no associativity, only parsing). *) diff --git a/Util/MMultiset/MMultisetWMap.v b/Util/MMultiset/MMultisetWMap.v index 4b0bf92ec8a13e3ef5e1e4a307a445f3bc05d5e9..184c77860ef5314baa4bed947ff3df651dfb7e8a 100644 --- a/Util/MMultiset/MMultisetWMap.v +++ b/Util/MMultiset/MMultisetWMap.v @@ -20,6 +20,7 @@ Require Import Pactole.Util.FMaps.FMapInterface. Require Import Pactole.Util.FMaps.FMapFacts. Require Import Pactole.Util.MMultiset.Preliminary. Require Import Pactole.Util.MMultiset.MMultisetInterface. +Require Import Pactole.Util.SetoidDefs. Require Import SetoidDec. @@ -149,9 +150,9 @@ Proof using . intros x y n m l Hl Hinx Hiny Heq. induction Hl as [| [z p] l]. inversion_clear Hiny. inversion_clear Hinx; inversion_clear Hiny. - compute in H0, H1. destruct H0 as [H0 ?], H1 as [H1 ?]. now subst p m. -- compute in H0. destruct H0 as [H0 ?]. subst p. elim H. +- compute in H0. destruct H0 as [H0 ?]. subst p. contradiction H. apply eq_pair_elt_weak_In with y m. now transitivity x. assumption. -- compute in H1. destruct H1 as [H1 ?]. subst p. elim H. +- compute in H1. destruct H1 as [H1 ?]. subst p. contradiction H. apply eq_pair_elt_weak_In with x n. now transitivity y; auto. assumption. - now apply IHHl. Qed. @@ -255,7 +256,7 @@ Lemma multiplicity_out : forall x m, m[x] = 0 <-> find x m = None. Proof using . simpl. unfold m_multiplicity. intros x m. split; intro Hm. + destruct (find x m) eqn:Hfind. - - elim (lt_irrefl 0). rewrite <- Hm at 2. now apply Pos2Nat.is_pos. + - contradiction (Nat.lt_irrefl 0). rewrite <- Hm at 2. now apply Pos2Nat.is_pos. - reflexivity. + now rewrite Hm. Qed. @@ -322,21 +323,21 @@ Proof using MapSpec. split. intros [n Habs]. now rewrite find_mapsto_iff, Hin in Habs. * intros x s s'. unfold inter. simpl. destruct (find x s') eqn:Hin. + setoid_rewrite (fold_add (fun x n => min (m_multiplicity x s) n) x p s' (FMapInterface.empty _)); trivial; [|]. - - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite empty_o, Hin, plus_0_r. + - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite empty_o, Hin, Nat.add_0_r. - intros ? ? Heq ? ? ?. subst. now rewrite Heq. + rewrite fold_add_out. - - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite Hin, Min.min_0_r, empty_o. + - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite Hin, Nat.min_0_r, empty_o. - intros [n Habs]. now rewrite find_mapsto_iff, Hin in Habs. * intros x s s'. unfold diff. simpl. destruct (find x s) eqn:Hin. + setoid_rewrite (fold_add (fun x n => n - multiplicity x s') x p s (FMapInterface.empty _)); trivial. - - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite empty_o, Hin, plus_0_r. + - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite empty_o, Hin, Nat.add_0_r. - intros ? ? Heq ? ? ?. subst. now rewrite Heq. + rewrite fold_add_out. - unfold multiplicity. simpl. unfold m_multiplicity. now rewrite Hin, empty_o. - intros [n Habs]. now rewrite find_mapsto_iff, Hin in Habs. * intros x s s'. unfold lub. simpl. replace (max (m_multiplicity x s) (m_multiplicity x s')) - with (m_multiplicity x s - m_multiplicity x s' + m_multiplicity x s') by (apply Max.max_case_strong; lia). + with (m_multiplicity x s - m_multiplicity x s' + m_multiplicity x s') by (apply Nat.max_case_strong; lia). destruct (find x s) eqn:Hin. + erewrite (fold_add); eauto. -unfold multiplicity. simpl. unfold m_multiplicity. now rewrite Hin. @@ -368,7 +369,7 @@ Proof using MapSpec. split. specialize (Hs x). cbn in Hs. unfold m_multiplicity in Hs. destruct (find x s1), (find x s2); solve [ apply Pos2Nat.inj in Hs; subst; reflexivity - | elim (lt_irrefl 0); rewrite Hs at 2 || rewrite <- Hs at 2; apply Pos2Nat.is_pos + | contradiction (Nat.lt_irrefl 0); rewrite Hs at 2 || rewrite <- Hs at 2; apply Pos2Nat.is_pos | split; intro; discriminate ]. + assumption. Qed. @@ -383,7 +384,7 @@ assert (His_empty_spec : forall s, is_empty s = true <-> s == empty). + rewrite <- is_empty_iff. intros x n Habs. rewrite find_mapsto_iff in Habs. specialize (H x). unfold m_multiplicity in H. rewrite Habs, empty_o in H. - apply (lt_irrefl 0). rewrite <- H at 2. apply Pos2Nat.is_pos. } + apply (Nat.lt_irrefl 0). rewrite <- H at 2. apply Pos2Nat.is_pos. } split; trivial. * intros s s'. unfold equal. simpl. destruct (FMapInterface.equal Pos.eqb s s') eqn:Heq. @@ -396,8 +397,8 @@ split; trivial. unfold m_multiplicity in Habs. specialize (Habs x). simpl in Habs. destruct (find x s) eqn:Hin1, (find x s') eqn:Hin2. - f_equal. now apply Pos2Nat.inj. - - elim (lt_irrefl 0). rewrite <- Habs at 2. now apply Pos2Nat.is_pos. - - elim (lt_irrefl 0). rewrite Habs at 2. now apply Pos2Nat.is_pos. + - contradiction (Nat.lt_irrefl 0). rewrite <- Habs at 2. now apply Pos2Nat.is_pos. + - contradiction (Nat.lt_irrefl 0). rewrite Habs at 2. now apply Pos2Nat.is_pos. - reflexivity. * intros s s'. unfold subset. simpl. cbn in His_empty_spec. rewrite His_empty_spec. clear His_empty_spec. @@ -406,16 +407,16 @@ split; trivial. + intro x. destruct (find x s) eqn:Hin. - cut (s[x] - s'[x] = 0). lia. rewrite <- diff_spec. simpl. erewrite fold_add; eauto; [|]. - ++ unfold multiplicity. simpl. unfold m_multiplicity at 2. rewrite empty_o, plus_0_r. + ++ unfold multiplicity. simpl. unfold m_multiplicity at 2. rewrite empty_o, Nat.add_0_r. specialize (Hle x). erewrite fold_add in Hle; eauto. -- unfold multiplicity in Hle. simpl in Hle. unfold m_multiplicity at 2 in Hle. - now rewrite empty_o, plus_0_r in Hle. + now rewrite empty_o, Nat.add_0_r in Hle. -- intros ? ? Heq ? ? ?. subst. now rewrite Heq. ++ intros ? ? Heq ? ? ?. subst. now rewrite Heq. - simpl. unfold m_multiplicity at 1. rewrite Hin. lia. + intro x. destruct (find x s) eqn:Hin. - erewrite fold_add; eauto. - -- unfold multiplicity. simpl. unfold m_multiplicity at 2. rewrite empty_o, plus_0_r. + -- unfold multiplicity. simpl. unfold m_multiplicity at 2. rewrite empty_o, Nat.add_0_r. specialize (Hle x). simpl in Hle. unfold m_multiplicity at 1 in Hle. rewrite Hin in Hle. lia. -- intros ? ? Heq ? ? ?. subst. now rewrite Heq. - erewrite fold_add_out; eauto. @@ -532,7 +533,7 @@ assert (Hs' : forall n, ~InA eq_key_elt (x, n) (FMapInterface.elements s)). revert o. induction (FMapInterface.elements s); simpl; intros o Hin. + assumption. + apply IHl in Hin. - - elim (Hs' (snd a)). left. split; simpl. now inversion Hin. reflexivity. + - contradiction (Hs' (snd a)). left. split; simpl. now inversion Hin. reflexivity. - intros n Habs. apply (Hs' n). now right. Qed. @@ -590,7 +591,7 @@ induction (FMapInterface.elements s); simpl; intro s'; simpl in Hs. -- clear. intros [] [] Hxy. now compute in *. - inversion_clear Hdup. rewrite NoDupA_inj_map in H0; solve [ eassumption | autoclass | now intros [] [] ]. + inversion_clear Hs. - - elim Hneq. destruct H as [H1 H2]. split; trivial; []. + - contradiction Hneq. destruct H as [H1 H2]. split; trivial; []. simpl in *. hnf in H2. cbn in H2. now rewrite H2, Pos2Nat.id. - inversion_clear Hdup. assert (Hxy : x =/= y) by (intro; eauto using eq_pair_elt_weak_In). @@ -625,7 +626,7 @@ assert (Hnfilter : forall (f : elt -> nat -> bool) (x : elt) (s : multiset elt), - unfold multiplicity. simpl. unfold m_multiplicity. rewrite empty_o. now destruct (f x 0). - unfold In. simpl. lia. + rewrite nfilter_spec_In. - - unfold multiplicity. simpl. unfold m_multiplicity at 3 4. now rewrite empty_o, Hin, plus_0_r. + - unfold multiplicity. simpl. unfold m_multiplicity at 3 4. now rewrite empty_o, Hin, Nat.add_0_r. - assumption. - unfold In. simpl. lia. * split; trivial. @@ -779,7 +780,7 @@ induction (FMapInterface.elements s); simpl; intros [s1 s2]; simpl in Hs. - inversion_clear Hdup. rewrite NoDupA_inj_map in H0; solve [ eassumption | autoclass | now intros [] [] ]. + inversion_clear Hs. - - elim Hneq. destruct H as [H1 H2]. split. assumption. hnf in *. simpl in *. now rewrite H2, Pos2Nat.id. + - contradiction Hneq. destruct H as [H1 H2]. split. assumption. hnf in *. simpl in *. now rewrite H2, Pos2Nat.id. - inversion_clear Hdup. assert (Hxy : x =/= y) by (intro; eauto using eq_pair_elt_weak_In). rewrite IHl; try assumption. simpl. unfold npartition_fun. @@ -853,7 +854,7 @@ induction (FMapInterface.elements s); simpl; intros [s1 s2]; simpl in Hs. - inversion_clear Hdup. rewrite NoDupA_inj_map in H0; solve [eassumption | autoclass | now intros [] [] ]. + inversion_clear Hs. - - elim Hneq. destruct H as [H1 H2]. split. assumption. hnf in *. simpl in *. now rewrite H2, Pos2Nat.id. + - contradiction Hneq. destruct H as [H1 H2]. split. assumption. hnf in *. simpl in *. now rewrite H2, Pos2Nat.id. - inversion_clear Hdup. assert (Hxy : x =/= y) by (intro; eauto using eq_pair_elt_weak_In). rewrite IHl; try assumption. simpl. @@ -889,7 +890,7 @@ Proof using MapSpec. split. - simpl. unfold multiplicity. simpl. unfold m_multiplicity. rewrite empty_o. now destruct (f x 0). - unfold In. simpl. lia. + setoid_rewrite npartition_spec_fst_In. - - simpl. unfold m_multiplicity at 3 4. now rewrite empty_o, Hin, plus_0_r. + - simpl. unfold m_multiplicity at 3 4. now rewrite empty_o, Hin, Nat.add_0_r. - assumption. - unfold In. simpl. lia. * intros f s Hf x. rewrite nfilter_spec. @@ -899,7 +900,7 @@ Proof using MapSpec. split. -- simpl. unfold m_multiplicity. rewrite empty_o. now destruct (f x 0). -- unfold In. simpl. lia. - rewrite npartition_spec_snd_In. - -- simpl. unfold m_multiplicity at 2 4. rewrite empty_o, Hin, plus_0_r. now destruct (f x (S n)). + -- simpl. unfold m_multiplicity at 2 4. rewrite empty_o, Hin, Nat.add_0_r. now destruct (f x (S n)). -- assumption. -- unfold In. simpl. lia. + intros ? ? Heq ? ? ?. subst. now rewrite Heq. diff --git a/Util/MMultiset/Preliminary.v b/Util/MMultiset/Preliminary.v index 06b3973b8a77cf8f872be5526d48b47b166a9a06..9bc2afcf5fb6d992ee52e0063bb4dc4647648950 100644 --- a/Util/MMultiset/Preliminary.v +++ b/Util/MMultiset/Preliminary.v @@ -10,7 +10,6 @@ Require Import Lia. -Require Import Arith.Div2. Require Import Reals. Require Import List. Require Import Morphisms. @@ -55,15 +54,15 @@ Definition transpose2 {A B C : Type} eqC (f : A -> B -> C -> C) := Definition additive2 {A B : Type} eqB (f : A -> nat -> B -> B) := forall x n p z, eqB (f x n (f x p z)) (f x (n + p) z). -Instance compose_compat {A B C : Type} eqA eqB eqC : forall (f : A -> B) (g : B -> C), +Global Instance compose_compat {A B C : Type} eqA eqB eqC : forall (f : A -> B) (g : B -> C), Proper (eqA ==> eqB) f -> Proper (eqB ==> eqC) g -> Proper (eqA ==> eqC) (fun x => g (f x)). Proof using . intros f g Hf Hg x y Heq. now apply Hg, Hf. Qed. -Instance compose2_compat {A B C D : Type} eqA eqB eqC eqD : forall (f : A -> B) (g : B -> C -> D), +Global Instance compose2_compat {A B C D : Type} eqA eqB eqC eqD : forall (f : A -> B) (g : B -> C -> D), Proper (eqA ==> eqB) f -> Proper (eqB ==> eqC ==> eqD) g -> Proper (eqA ==> eqC ==> eqD) (fun x => g (f x)). Proof using . intros f g Hf Hg x y Heq. now apply Hg, Hf. Qed. -Instance compose3_compat {A B C D E : Type} eqA eqB eqC eqD eqE : forall (f : A -> B) (g : B -> C -> D -> E), +Global Instance compose3_compat {A B C D E : Type} eqA eqB eqC eqD eqE : forall (f : A -> B) (g : B -> C -> D -> E), Proper (eqA ==> eqB) f -> Proper (eqB ==> eqC ==> eqD ==> eqE) g -> Proper (eqA ==> eqC ==> eqD ==> eqE) (fun x => g (f x)). Proof using . intros f g Hf Hg x y Heq. now apply Hg, Hf. Qed. @@ -101,7 +100,7 @@ Lemma NoDupA_2 : forall x y, ~eqA x y -> NoDupA eqA (x :: y :: nil). Proof using . intros x y Hdiff. constructor. intro Habs. inversion_clear Habs. - now elim Hdiff. + now contradiction Hdiff. inversion H. apply NoDupA_singleton. Qed. @@ -152,7 +151,7 @@ Proof using . intros x l Hin. induction l; simpl. + reflexivity. + destruct (eq_dec x a). - - elim Hin. now left. + - contradiction Hin. now left. - f_equal. apply IHl. intro Habs. apply Hin. now right. Qed. @@ -162,7 +161,7 @@ Proof using HeqA. intros x y l Hxy. induction l. reflexivity. simpl. destruct (eq_dec y a). subst. rewrite IHl. split; intro H. now right. inversion_clear H. - elim Hxy. now transitivity a. + contradiction Hxy. now transitivity a. assumption. split; intro H; inversion_clear H; (now left) || right; now rewrite IHl in *. Qed. @@ -174,7 +173,7 @@ Proof using HeqA. intros Hsub x y l. induction l as [| a l]. * split; intro Habs. inversion Habs. destruct Habs as [Habs _]. inversion Habs. * simpl. destruct (eq_dec y a) as [Heq | Hneq]. - + rewrite IHl. intuition. inversion_clear H2. apply Hsub in H0. now elim H3; transitivity a. assumption. + + rewrite IHl. intuition. inversion_clear H2. apply Hsub in H0. now contradiction H3; transitivity a. assumption. + split; intro Hin. - inversion_clear Hin. split. now left. rewrite H. intro. now apply Hneq. @@ -190,7 +189,7 @@ Proof using HeqA. apply (removeA_InA_iff_strong eq_dec). reflexivity. Qed. Lemma removeA_InA_weak eq_dec : forall x y l, InA eqA' x (@removeA A eqA eq_dec y l) -> InA eqA' x l. Proof using . intros x y l Hin. induction l; simpl in *. -+ rewrite InA_nil in Hin. elim Hin. ++ rewrite InA_nil in Hin. contradiction Hin. + destruct (eq_dec y a) as [Heq | Heq]. - auto. - inversion_clear Hin; auto. @@ -205,8 +204,8 @@ intros x y Hxy l l' ?. subst l'. induction l; simpl. + reflexivity. + destruct (eq_dec x a) as [Heqx | Hneqx], (eq_dec y a) as [Heqy | Hneqy]. - apply IHl. - - elim Hneqy. now rewrite <- Hxy. - - elim Hneqx. now rewrite Hxy. + - contradiction Hneqy. now rewrite <- Hxy. + - contradiction Hneqx. now rewrite Hxy. - f_equal. apply IHl. Qed. @@ -217,17 +216,17 @@ intros x y ? l1 l2 Hl. subst. induction Hl. + constructor. + simpl. destruct (eq_dec x xâ‚) as [Heqâ‚ | Hneqâ‚], (eq_dec y xâ‚‚) as [Heqâ‚‚ | Hneqâ‚‚]. - assumption. - - elim Hneqâ‚‚. now rewrite <- H, Heqâ‚. - - elim Hneqâ‚. now rewrite H, Heqâ‚‚. + - contradiction Hneqâ‚‚. now rewrite <- H, Heqâ‚. + - contradiction Hneqâ‚. now rewrite H, Heqâ‚‚. - now apply PermutationA_cons. + simpl. destruct (eq_dec x x0), (eq_dec y y0), (eq_dec y x0), (eq_dec x y0); - try (now elim n; rewrite H) || (now elim n; rewrite <- H). + try (now contradiction n; rewrite H) || (now contradiction n; rewrite <- H). - now erewrite removeA_eq_compat. - constructor. reflexivity. now erewrite removeA_eq_compat. - - elim n0. now rewrite <- H. + - contradiction n0. now rewrite <- H. - constructor. reflexivity. now erewrite removeA_eq_compat. - - elim n1. now rewrite H. - - elim n0. now rewrite <- H. + - contradiction n1. now rewrite H. + - contradiction n0. now rewrite <- H. - etransitivity. constructor 3. repeat constructor; reflexivity || now erewrite removeA_eq_compat. + constructor 4 with (removeA eq_dec y lâ‚‚). - assumption. @@ -243,7 +242,7 @@ Qed. Corollary removeA_inside_in eq_dec : forall (x : A) l1 l2, @removeA A eqA eq_dec x (l1 ++ x :: l2) = removeA eq_dec x l1 ++ removeA eq_dec x l2. -Proof using HeqA. intros x ? ?. rewrite removeA_app. simpl. destruct (eq_dec x x) as [| Hneq]; trivial. now elim Hneq. Qed. +Proof using HeqA. intros x ? ?. rewrite removeA_app. simpl. destruct (eq_dec x x) as [| Hneq]; trivial. now contradiction Hneq. Qed. Corollary removeA_inside_out eq_dec : forall (x y : A) l1 l2, ~eqA x y -> @removeA A eqA eq_dec x (l1 ++ y :: l2) = removeA eq_dec x l1 ++ y :: removeA eq_dec x l2. @@ -538,7 +537,7 @@ Lemma inclA_cons_inv : forall x y l1 l2, Proof using HeqA. intros x y l1 l2 Hx Heq Hincl z Hin. assert (Hin' : InA eqA z (x :: l1)) by now right. apply Hincl in Hin'. -inversion_clear Hin'; trivial. elim Hx. now rewrite Heq, <- H. +inversion_clear Hin'; trivial. contradiction Hx. now rewrite Heq, <- H. Qed. Lemma inclA_length : forall l1 l2, NoDupA eqA l1 -> inclA eqA l1 l2 -> length l1 <= length l2. @@ -556,7 +555,7 @@ Lemma not_NoDupA : (forall x y, {eqA x y} + {~eqA x y} ) -> Proof using HeqA. intros eq_dec l. split; intro Hl. + induction l. - elim Hl. now constructor. + contradiction Hl. now constructor. destruct (InA_dec eq_dec a l) as [Hin | Hnin]. exists a. apply PermutationA_split in Hin. destruct Hin as [l' Hperm]. exists l'. now rewrite Hperm. destruct IHl as [x [l' Hperm]]. @@ -600,7 +599,10 @@ Qed. Corollary filter_length : forall f (l : list A), length (filter f l) = length l - length (filter (fun x => negb (f x)) l). -Proof using . intros. apply plus_minus. rewrite <- (partition_length f), partition_filter. simpl. apply plus_comm. Qed. +Proof using . +intros. symmetry. apply Nat.add_sub_eq_l. +rewrite <- (partition_length f l), partition_filter. cbn. apply Nat.add_comm. +Qed. Lemma map_cond_Permutation : forall (f : A -> bool) (gâ‚ gâ‚‚ : A -> B) l, Permutation (map (fun x => if f x then gâ‚ x else gâ‚‚ x) l) @@ -737,13 +739,13 @@ End List_results. Lemma le_neq_lt : forall m n : nat, n <= m -> n <> m -> n < m. -Proof using . intros n m Hle Hneq. now destruct (le_lt_or_eq _ _ Hle). Qed. +Proof using . intros n m Hle Hneq. rewrite Nat.lt_eq_cases in Hle. intuition auto with *. Qed. Lemma min_is_0 : forall n m, min n m = 0 <-> n = 0 \/ m = 0. -Proof using . intros [| n] [| m]; intuition;discriminate. Qed. +Proof using . intros [| n] [| m]; intuition auto with *. Qed. Lemma max_is_0 : forall n m, max n m = 0 <-> n = 0 /\ m = 0. -Proof using . intros [| n] [| m]; intuition; discriminate. Qed. +Proof using . intros [| n] [| m]; intuition auto with *; discriminate. Qed. Lemma Bleb_refl : forall x, Bool.le x x. Proof using . intros [|]; simpl; auto. Qed. diff --git a/Util/NumberComplements.v b/Util/NumberComplements.v index be488364305d749cabcd9d8dcd27d05630bc679c..ae022b589c53129db4c6cf69afbc98e03c6bcc2c 100644 --- a/Util/NumberComplements.v +++ b/Util/NumberComplements.v @@ -18,18 +18,14 @@ *) (**************************************************************************) - -Require Import SetoidDec. -Require Import Reals. -Require Import Lia Psatz. - +Require Import Utf8 SetoidDec Reals Lia Psatz. +Require Import Pactole.Util.SetoidDefs. +Set Implicit Arguments. (* ******************************** *) (** * Results about real numbers **) (* ******************************** *) - -Set Implicit Arguments. Open Scope R_scope. (* Should be in Reals from the the std lib! *) @@ -46,15 +42,17 @@ intros x y. unfold Basics.flip. cbn. split; intro Hxy. - destruct Hxy. now apply Rle_antisym. Qed. -Lemma Rdec : forall x y : R, {x = y} + {x <> y}. -Proof. -intros x y. destruct (Rle_dec x y). destruct (Rle_dec y x). - left. now apply Rle_antisym. - right; intro; subst. contradiction. - right; intro; subst. pose (Rle_refl y). contradiction. +Global Instance Rle_partialorder_equiv : PartialOrder equiv Rle := Rle_partialorder. + +Global Instance Rlt_SO : StrictOrder Rlt. +Proof. split. ++ intro. apply Rlt_irrefl. ++ intros ? ? ?. apply Rlt_trans. Qed. -Instance R_EqDec : @EqDec R _ := Rdec. +Global Instance R_EqDec : @EqDec R _ := Req_dec_T. +(* #[export] *) +Notation Rdec := R_EqDec. Lemma Rdiv_le_0_compat : forall a b, 0 <= a -> 0 < b -> 0 <= a / b. Proof. intros a b ? ?. now apply Fourier_util.Rle_mult_inv_pos. Qed. @@ -104,7 +102,7 @@ Proof. intros k r r'. destruct (Rdec_bool (k + r) (k + r')) eqn:Heq1, (Rdec_bool r r') eqn:Heq2; trivial; rewrite ?Rdec_bool_true_iff, ?Rdec_bool_false_iff in *. -- elim Heq2. eapply Rplus_eq_reg_l; eassumption. +- contradiction Heq2. eapply Rplus_eq_reg_l; eassumption. - subst. auto. Qed. @@ -114,7 +112,7 @@ intros k r r' Hk. destruct (Rdec_bool r r') eqn:Heq1, (Rdec_bool (k * r) (k * r')) eqn:Heq2; trivial; rewrite ?Rdec_bool_true_iff, ?Rdec_bool_false_iff in *. - subst. auto. -- elim Heq1. eapply Rmult_eq_reg_l; eassumption. +- contradiction Heq1. eapply Rmult_eq_reg_l; eassumption. Qed. Corollary Rdec_bool_plus_r : forall k r r', Rdec_bool (r + k) (r' + k) = Rdec_bool r r'. @@ -214,7 +212,6 @@ Qed. Close Scope R_scope. - (** * Results about integers **) Lemma nat_compare_Eq_comm : forall n m, Nat.compare n m = Eq <-> Nat.compare m n = Eq. @@ -241,15 +238,225 @@ Qed. Lemma even_div2 : forall n, Nat.Even n -> Nat.div2 n + Nat.div2 n = n. Proof. intros n Hn. replace (Nat.div2 n + Nat.div2 n) with (2 * Nat.div2 n) by lia. -rewrite <- Nat.double_twice. symmetry. apply Div2.even_double. now rewrite Even.even_equiv. +rewrite <- Nat.double_twice. symmetry. now apply Nat.Even_double. +Qed. + +Lemma eq_0_lt_dec : ∀ n : nat, {n = 0} + {0 < n}. +Proof using . + intros. destruct (le_lt_dec n 0) as [Hd | Hd]. + left. apply Nat.le_0_r, Hd. right. apply Hd. +Qed. + +Lemma sub_sub : ∀ m n p : nat, p <= n → m - (n - p) = m + p - n. +Proof using . lia. Qed. + +Lemma sub_cancel_l : ∀ p m n : nat, m < p -> p - m = p - n <-> m = n. +Proof using . lia. Qed. + +Lemma add_mul_lt : ∀ a b c d : nat, a < d -> b < c -> d * b + a < d * c. +Proof using . nia. Qed. + +Lemma le_neq_lt : forall m n : nat, n <= m -> n ≠m -> n < m. +Proof using . lia. Qed. + +Lemma lt_sub_lt_add_l : ∀ n m p : nat, m <= n -> n < m + p -> n - m < p. +Proof using . lia. Qed. + +Lemma lt_sub_lt_add_r : ∀ n m p : nat, p <= n -> n < m + p -> n - p < m. +Proof using . lia. Qed. + +Lemma sub_le_mono : ∀ n m p q : nat, n <= m -> p <= q -> n - q <= m - p. +Proof using . lia. Qed. + +Lemma lt_sub : ∀ n m : nat, 0 < n -> 0 < m -> n - m < n. +Proof using . lia. Qed. + +Lemma sub_lt_mono : ∀ n m p q : nat, p <= q -> p < m -> n < m -> n - q < m - p. +Proof using . lia. Qed. + +(* Statement improvable? Weaker precondition than m <= p? *) +Lemma sub_lt_mono_l : ∀ m n p : nat, m <= p -> n < m -> p - m < p - n. +Proof using . lia. Qed. + +Lemma add_sub_le_sub_add : ∀ n m p : nat, n + m - p <= n - p + m. +Proof using . lia. Qed. + +Lemma S_pred2 : ∀ n : nat, 1 < n -> n = S (S (Nat.pred (Nat.pred n))). +Proof using . lia. Qed. + +Lemma S_pred3 : ∀ n : nat, + 2 < n -> n = S (S (S (Nat.pred (Nat.pred (Nat.pred n))))). +Proof using . lia. Qed. + +Lemma mod_bounded_diffI : ∀ p m n : nat, p ≠0 -> m <= n + -> n - m < p -> m mod p = n mod p -> m = n. +Proof using . +intros * Hne Hle Hsu Heq. +rewrite (Nat.div_mod m p), (Nat.div_mod n p), Heq; trivial; []. +apply Nat.add_cancel_r, Nat.mul_cancel_l; trivial; []. +apply Nat.le_antisymm. ++ apply Nat.Div0.div_le_mono. exact Hle. ++ apply Nat.sub_0_le, Nat.lt_1_r. rewrite (Nat.mul_lt_mono_pos_l p); try lia; []. + rewrite Nat.mul_1_r, Nat.mul_sub_distr_l. + assert (Hmod : ∀ x y : nat, y ≠0 -> x - x mod y = y * (x / y)). + { intros. rewrite Nat.Div0.mod_eq; trivial; []. + cut (y * (x / y) <= x); try lia; []. + rewrite (Nat.div_mod_eq x y) at 2. assert (HH := Nat.Div0.mod_le x y). lia. } + setoid_rewrite <- Hmod; lia. +Qed. + +Lemma between_1 : ∀ m : nat, m <= m < m +1. +Proof using . lia. Qed. + +Lemma between_neq_0 : ∀ n p m : nat, p <= m < p + n -> n ≠0. +Proof using . lia. Qed. + +Lemma between_addn : ∀ n1 n2 p1 p2 m1 m2 : nat, p1 <= m1 < p1 + n1 + -> p2 <= m2 < p2 + n2 -> p1 + p2 <= m1 + m2 < p1 + p2 + Nat.pred (n1 + n2). +Proof using . lia. Qed. + +Lemma between_subn : ∀ n1 n2 p1 p2 m1 m2 : nat, p1 <= m1 < p1 + n1 + -> p2 <= m2 < p2 + n2 -> p1 - Nat.pred (p2 + n2) + <= m1 - m2 < p1 - Nat.pred (p2 + n2) + Nat.pred (n1 + n2). +Proof using . lia. Qed. + +Lemma bounded_diff_between : ∀ P : nat -> nat -> Prop, Symmetric P + -> ∀ n : nat, (∀ m1 m2 : nat, m1 <= m2 -> m2 - m1 < n -> P m1 m2) + <-> ∀ (p m1 m2: nat), p <= m1 < p + n -> p <= m2 < p + n -> P m1 m2. +Proof using . +intros * HS *. split. ++ intros H * [H11 H12] [H21 H22]. + destruct (le_ge_dec m1 m2) as [Hd | Hd]; [| symmetry]; apply H; trivial; lia. ++ intros H m1 m2 Hle Hlt. apply (H m1); split; lia. +Qed. + +Lemma inj_sym : ∀ (A B : Type) (SA : Setoid A) (SB : Setoid B) (f : A -> B), + Symmetric (λ a1 a2, f a1 == f a2 -> a1 == a2). +Proof using . + intros * a1 a2 H1 H2. symmetry. apply H1. symmetry. exact H2. +Qed. + +Lemma mod_betweenI : ∀ p n m1 m2 : nat, p <= m1 < p + n -> p <= m2 < p + n + -> m1 mod n = m2 mod n -> m1 = m2. +Proof using . + intros * H. pose proof (between_neq_0 _ _ _ H) as Hn. revert m1 m2 H. + eapply (bounded_diff_between _ (inj_sym _ _ (eq_setoid nat) (eq_setoid nat) + (λ x, x mod n))). intros * H1 H2 H3. eapply mod_bounded_diffI. all: eassumption. +Qed. + +Lemma between_muln : ∀ p n m : nat, n * p <= m < n * p + n -> m / n = p. +Proof using . intros * []. symmetry. apply Nat.div_unique with (m - n * p); nia. Qed. + +Lemma mod_le_between_compat : ∀ p n m1 m2 : nat, n * p <= m1 < n * p + n + -> n * p <= m2 < n * p + n -> m1 <= m2 -> m1 mod n <= m2 mod n. +Proof using . + intros * H1 H2 Hle. erewrite 2 Nat.Div0.mod_eq, 2 between_muln, + Nat.add_le_mono_r, 2 Nat.sub_add;try assumption;try eassumption. + - apply H2. + - apply H1. +Qed. + +Lemma mod_lt_between_compat : ∀ p n m1 m2 : nat, n * p <= m1 < n * p + n + -> n * p <= m2 < n * p + n -> m1 < m2 -> m1 mod n < m2 mod n. +Proof using . + intros * H1 H2 Hlt. erewrite 2 Nat.Div0.mod_eq, 2 between_muln, + Nat.add_lt_mono_r, 2 Nat.sub_add;try assumption; try eassumption. + - apply H2. + - apply H1. +Qed. + +Lemma divide_neq : ∀ a b : nat, b ≠0 -> Nat.divide a b -> a ≠0. +Proof using . intros * Hn Hd H. subst. apply Hn, Nat.divide_0_l, Hd. Qed. + +Lemma divide_div_neq : ∀ a b : nat, b ≠0 -> Nat.divide a b -> b / a ≠0. +Proof using . + intros * Hn Hd H. apply Hn. unshelve erewrite (Nat.div_mod b a), + (proj2 (Nat.Lcm0.mod_divide b a)), Nat.add_0_r, H; try eassumption. + - apply Nat.mul_0_r. + - eapply divide_neq; eassumption. + Qed. + +Lemma addn_muln_divn : ∀ n q1 q2 r1 r2 : nat, n ≠0 -> r1 < n -> r2 < n + -> n * q1 + r1 = n * q2 + r2 -> q1 = q2 ∧ r1 = r2. +Proof using . + intros * Hn H1 H2 H. apply Nat.div_unique in H as H'. rewrite Nat.mul_comm, + Nat.div_add_l, Nat.div_small, Nat.add_0_r in H'. subst. split. reflexivity. + eapply Nat.add_cancel_l. all: eassumption. +Qed. + +Lemma divide_mul : ∀ a b : nat, b ≠0 -> Nat.divide a b -> b = a * (b / a). +Proof using . + intros * Hn Hd. erewrite (Nat.div_mod b a) at 1. + unshelve erewrite (proj2 (Nat.Lcm0.mod_divide b a )); try eassumption. + - apply Nat.add_0_r. + - eapply divide_neq;eassumption. +Qed. + +Lemma divide_mod : + ∀ a b c : nat, b ≠0 -> Nat.divide a b -> (c mod b) mod a = c mod a. +Proof using . + intros * Hn Hd. rewrite (divide_mul a b), Nat.Div0.mod_mul_r, Nat.mul_comm, + Nat.Div0.mod_add; try assumption. + apply Nat.Div0.mod_mod. +Qed. + +Lemma mul_add : ∀ a b : nat, + b ≠0 -> a * b = a + Nat.pred b * a. +Proof using . + intros * Hb. destruct (Nat.eq_dec a 0) as [?|Ha]. + - subst. rewrite Nat.mul_0_r. apply Nat.mul_0_l. + - rewrite <- (Nat.mul_1_l a) at 2. rewrite <- Nat.mul_add_distr_r, + Nat.add_1_l, Nat.succ_pred. apply Nat.mul_comm. apply Hb. +Qed. + +Lemma pred_mul_add : ∀ a b : nat, + b ≠0 -> Nat.pred (a * b) = Nat.pred a + Nat.pred b * a. +Proof using . + intros * Hb. destruct (Nat.eq_dec a 0) as [?|Ha]. + - subst. rewrite Nat.mul_0_r. apply Nat.mul_0_l. + - rewrite mul_add by apply Hb. symmetry. apply Nat.add_pred_l, Ha. +Qed. + +Lemma lt_pred_mul : ∀ a b c d : nat, + a < b -> c < d -> Nat.pred ((S a) * (S c)) < b * d. +Proof using . + intros * Hab Hcd. erewrite <- Nat.le_succ_l, Nat.lt_succ_pred. + - apply Nat.mul_le_mono. apply Hab. apply Hcd. + - apply Nat.mul_pos_pos. all: apply Nat.lt_0_succ. +Qed. + +Fact mul_le_lt_compat : ∀ (n m p q : nat), + 0 < n -> n <= m -> p < q -> n * p < m * q. +Proof using. + intros. + apply Nat.le_lt_trans with (m * p). + - apply Nat.mul_le_mono_r; assumption. + - apply Nat.mul_lt_mono_pos_l; try assumption. + apply Nat.lt_le_trans with n; assumption. +Qed. + +Corollary mul_lt_compat : ∀ (m p q : nat), 0 < m -> p < q -> p < m * q. +Proof using. + intros. + rewrite <- (Nat.mul_1_l p). + apply (@mul_le_lt_compat 1); try assumption. + apply Nat.lt_0_1. Qed. -Lemma le_neq_lt : forall m n : nat, n <= m -> n <> m -> n < m. -Proof. intros n m Hle Hneq. now destruct (le_lt_or_eq _ _ Hle). Qed. +Lemma neq_lt : ∀ a b : nat, a < b -> b ≠a. +Proof using . + intros * Hlt Heq. apply (Nat.lt_irrefl a). rewrite <- Heq at 2. exact Hlt. +Qed. + +Lemma lt_pred : ∀ a b : nat, a < b -> Nat.pred b < b. +Proof using . lia. Qed. + +Lemma neq_pred : ∀ a b c : nat, a < b -> b < c -> Nat.pred c ≠a. +Proof using . lia. Qed. Open Scope Z. -Instance Z_EqDec : @EqDec Z _ := Z.eq_dec. +Global Instance Z_EqDec : @EqDec Z _ := Z.eq_dec. Lemma Zincr_mod : forall k n, 0 < n -> (k + 1) mod n = k mod n + 1 \/ (k + 1) mod n = 0 /\ k mod n = n - 1. @@ -291,3 +498,115 @@ Lemma Zmin_bounds : forall n m, n < m -> Z.min n (m - n) <= m / 2. Proof. intros. apply Z.min_case_strong; intro; apply Zdiv.Zdiv_le_lower_bound; lia. Qed. Close Scope Z. + +Class ltc (l u : nat) := lt_l_u : l < u. +Infix "<c" := ltc (at level 70, no associativity). + +Section ltc_l_u. + +Context {l u : nat} {ltc_l_u : l <c u}. + +Lemma neq_u_l : u ≠l. +Proof using ltc_l_u. apply neq_lt, lt_l_u. Qed. + +Lemma le_l_pred_u : l <= Nat.pred u. +Proof using ltc_l_u. apply Nat.lt_le_pred, lt_l_u. Qed. + +Lemma lt_pred_u : Nat.pred u < u. +Proof using ltc_l_u. eapply lt_pred, lt_l_u. Qed. + +Lemma S_pred_u : S (Nat.pred u) = u. +Proof using ltc_l_u. eapply Nat.lt_succ_pred, lt_l_u. Qed. + +End ltc_l_u. + +Section ltc_s_u. + +Context {l u : nat} {ltc_l_u : l <c u}. + +Lemma lt_s_u : ∀ s : nat, s <= l -> s < u. +Proof using ltc_l_u. + intros * H. eapply Nat.le_lt_trans. exact H. exact lt_l_u. +Qed. + +Lemma lt_s_pred_u : ∀ s : nat, s < l -> s < Nat.pred u. +Proof using ltc_l_u. + intros * H. eapply Nat.lt_le_trans. exact H. exact le_l_pred_u. +Qed. + +Lemma le_s_pred_u : ∀ s : nat, s <= l -> s <= Nat.pred u. +Proof using ltc_l_u. + intros * H. eapply Nat.le_trans. exact H. exact le_l_pred_u. +Qed. + +Lemma neq_pred_u_s : ∀ s : nat, s < l -> Nat.pred u ≠s. +Proof using ltc_l_u. intros * H. eapply neq_pred. apply H. apply lt_l_u. Qed. + +Lemma lt_S_l_u : u ≠S l -> S l < u. +Proof using ltc_l_u. + intros H. apply le_neq_lt. apply ltc_l_u. apply not_eq_sym, H. +Qed. + +Lemma eq_S_l_lt_dec : {u = S l} + {S l < u}. +Proof using ltc_l_u. + destruct (Nat.eq_dec u (S l)) as [Hd | Hd]. + left. apply Hd. right. apply lt_S_l_u, Hd. +Qed. + +Global Instance lt_0_u : 0 <c u. +Proof using ltc_l_u. apply lt_s_u, Nat.le_0_l. Qed. + +Lemma neq_u_0 : u ≠0. +Proof using ltc_l_u. apply neq_u_l. Qed. + +Lemma le_0_pred_u : 0 <= Nat.pred u. +Proof using ltc_l_u. apply le_l_pred_u. Qed. + +Lemma lt_l_g : ∀ g : nat, u <= g -> l < g. +Proof using ltc_l_u. + intros * H. eapply Nat.lt_le_trans. exact lt_l_u. exact H. +Qed. + +Lemma lt_sub_u : ∀ s : nat, 0 < s -> u - s < u. +Proof using ltc_l_u. intros * H. apply lt_sub, H. apply lt_0_u. Qed. + +End ltc_s_u. + +Section lt_pred_mul_ul_ur. + +Context {ll lr ul ur : nat} {ltc_ll_ul : ll <c ul} {ltc_lr_ur : lr <c ur}. + +Lemma lt_pred_mul_ul_ur : Nat.pred ((S ll) * (S lr)) < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + apply lt_pred_mul. all: apply lt_l_u. +Qed. + +Lemma lt_ll_mul_ul_ur : ll < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + rewrite Nat.mul_comm. apply mul_lt_compat. apply lt_0_u. apply lt_l_u. +Qed. + +Lemma lt_lr_mul_ul_ur : lr < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + apply mul_lt_compat. apply lt_0_u. apply lt_l_u. +Qed. + +End lt_pred_mul_ul_ur. + +Section lt_pred_mul_ul_ur. + +Context {ll lr ul ur : nat} {ltc_ll_ul : ll <c ul} {ltc_lr_ur : lr <c ur}. + +Lemma lt_pred_ul_mul_ul_ur : Nat.pred ul < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + eapply @lt_s_u. eapply @lt_pred_mul_ul_ur. apply lt_pred_u. + apply lt_0_u. rewrite Nat.mul_1_r. reflexivity. +Qed. + +Lemma lt_pred_ur_mul_ul_ur : Nat.pred ur < ul * ur. +Proof using ltc_ll_ul ltc_lr_ur. + eapply @lt_s_u. eapply @lt_pred_mul_ul_ur. apply lt_0_u. + apply lt_pred_u. rewrite Nat.mul_1_l. reflexivity. +Qed. + +End lt_pred_mul_ul_ur. diff --git a/Util/Preliminary.v b/Util/Preliminary.v index da5a5ad2af0470420bc6a4eaeee3df40de353362..5a579429aa0606cad4205eefc98496071e9d3c8d 100644 --- a/Util/Preliminary.v +++ b/Util/Preliminary.v @@ -19,13 +19,10 @@ (**************************************************************************) -Require Import Relations. -Require Import Morphisms. -Require Import SetoidClass. -Require Pactole.Util.FMaps.FMapInterface. (* for prod_Setoid and prod_EqDec *) +Require Import Utf8 Relations Morphisms SetoidClass. +Require Logic.Decidable. Set Implicit Arguments. - Ltac autoclass := eauto with typeclass_instances. Ltac inv H := inversion H; subst; clear H. Hint Extern 1 (equiv ?x ?x) => reflexivity : core. @@ -84,23 +81,109 @@ Definition monotonic {A B : Type} (RA : relation A) (RB : relation B) (f : A -> Definition full_relation {A : Type} := fun _ _ : A => True. -Global Hint Extern 0 (full_relation _ _) => exact I : core. +#[export] Hint Extern 0 (full_relation _ _) => exact I : core. -Instance relation_equivalence_subrelation {A} : +Global Instance relation_equivalence_subrelation {A} : forall R R' : relation A, relation_equivalence R R' -> subrelation R R'. Proof. intros R R' Heq x y Hxy. now apply Heq. Qed. +Global Instance neq_Symmetric {A} : Symmetric (fun x y : A => x <> y). +Proof. intros ? ? Hneq ?. apply Hneq. now symmetry. Qed. + Global Hint Extern 3 (relation_equivalence _ _) => symmetry : core. (** Notations for composition and inverse *) + Class Composition T `{Setoid T} := { compose : T -> T -> T; - compose_compat :> Proper (equiv ==> equiv ==> equiv) compose }. + #[global]compose_compat :: Proper (equiv ==> equiv ==> equiv) compose }. Infix "∘" := compose (left associativity, at level 40). Arguments Composition T {_}. +(* TODO: generic properties of inverse should be proven here, not in each file + (Bijection, Similarity, Isometry, etc.) *) Class Inverse T `{Setoid T} := { inverse : T -> T; - inverse_compat :> Proper (equiv ==> equiv) inverse }. + #[global]inverse_compat :: Proper (equiv ==> equiv) inverse }. Notation "bij â»Â¹" := (inverse bij) (at level 39). Arguments Inverse T {_}. + +(* pick_spec can be useful when you want to test whether some 'oT : option T' is + 'Some t' or 'None'. Destructing 'pick_spec P oT' gives you directly 'P t' + in the first case and '∀ t : T, ¬ (P t)' in the other one *) +Variant pick_spec (T : Type) (P : T -> Prop) : option T -> Type := + Pick : ∀ x : T, P x → pick_spec P (Some x) + | Nopick : (∀ x : T, ¬ (P x)) → pick_spec P None. +Arguments pick_spec [T] _ _. +Arguments Pick [T P x] _. +Arguments Nopick [T P] _. + +Lemma injective_compat_iff : + ∀ {A B : Type} {eqA : relation A} {eqB : relation B} (f : A -> B) + {compat : Proper (eqA ==> eqB) f}, injective eqA eqB f + -> ∀ a1 a2 : A, eqB (f a1) (f a2) <-> eqA a1 a2. +Proof using . + intros * Hcompat Hinj *. split. apply Hinj. apply Hcompat. +Qed. + +Lemma injective_eq_iff : ∀ {A B : Type} (f : A -> B), + injective eq eq f -> ∀ a1 a2 : A, f a1 = f a2 <-> a1 = a2. +Proof using . intros *. apply injective_compat_iff. autoclass. Qed. + +Lemma eq_sym_iff : ∀ (A : Type) (x y : A), x = y <-> y = x. +Proof using . intros. split. all: apply eq_sym. Qed. + +Lemma decidable_not_and_iff : + ∀ P Q : Prop, Decidable.decidable P -> ¬ (P ∧ Q) <-> ¬ P ∨ ¬ Q. +Proof using . + intros * Hd. split. apply Decidable.not_and, Hd. intros Ho Ha. + destruct Ho as [Ho | Ho]. all: apply Ho, Ha. +Qed. + +Lemma decidable_not_not_iff : ∀ P : Prop, Decidable.decidable P -> ¬ ¬ P <-> P. +Proof using . + intros * Hd. split. apply Decidable.not_not, Hd. intros H Hn. apply Hn, H. +Qed. + +Lemma sumbool_iff_compat : + ∀ P Q R S: Prop, P <-> R -> Q <-> S -> {P} + {Q} -> {R} + {S}. +Proof using . + intros * H1 H2 [Hd | Hd]. left. apply H1, Hd. right. apply H2, Hd. +Qed. + +Lemma sumbool_not_iff_compat : + ∀ P Q : Prop, P <-> Q -> {P} + {¬ P} -> {Q} + {¬ Q}. +Proof using . + intros * H Hd. eapply sumbool_iff_compat. + 2: apply not_iff_compat. 1,2: apply H. apply Hd. +Qed. + +Lemma sumbool_and_compat : + ∀ P Q R S: Prop, {P} + {Q} -> {R} + {S} -> {P ∧ R} + {Q ∨ S}. +Proof using . + intros * [H1 | H1] [H2 | H2]. left. split. apply H1. apply H2. + all: right. right. apply H2. all: left. all: apply H1. +Qed. + +Lemma sumbool_decidable : ∀ P : Prop, {P} + {¬ P} -> Decidable.decidable P. +Proof using . intros P [H | H]. left. apply H. right. apply H. Qed. + +Lemma pair_eqE : + ∀ {A B : Type} (p1 p2 : A * B), fst p1 = fst p2 ∧ snd p1 = snd p2 <-> p1 = p2. +Proof using . + intros. rewrite (surjective_pairing p1), (surjective_pairing p2). + symmetry. apply pair_equal_spec. +Qed. + +Lemma pair_dec : ∀ {A B : Type}, + (∀ a1 a2 : A, {a1 = a2} + {a1 ≠a2}) -> (∀ b1 b2 : B, {b1 = b2} + {b1 ≠b2}) + -> ∀ p1 p2 : A * B, {p1 = p2} + {p1 ≠p2}. +Proof using . + intros * Ha Hb *. eapply sumbool_not_iff_compat. apply pair_eqE. + eapply sumbool_iff_compat. reflexivity. symmetry. apply decidable_not_and_iff. + 2: apply sumbool_and_compat. 2: apply Ha. 2: apply Hb. + apply sumbool_decidable, Ha. +Qed. + +Lemma and_cancel : ∀ P : Prop, P -> P ∧ P. +Proof using . intros * H. split. all: apply H. Qed. diff --git a/Util/Ratio.v b/Util/Ratio.v index 6926488394af4c3c635bb556c5b3eada5dd1d81f..496d3492c9ba76aab8379d5422cb7373733fd8f7 100644 --- a/Util/Ratio.v +++ b/Util/Ratio.v @@ -32,12 +32,15 @@ Require Import Pactole.Util.Coqlib. (** A ratio (of some quantity), as a real number between [0] and [1]. *) Definition ratio := {x : R | 0 <= x <= 1}%R. -Instance ratio_Setoid : Setoid ratio := sig_Setoid _. -Instance ratio_EqDec : EqDec ratio_Setoid := @sig_EqDec _ _ _ _. +Global Instance ratio_Setoid : Setoid ratio := sig_Setoid _. +Global Instance ratio_EqDec : EqDec ratio_Setoid := @sig_EqDec _ _ _ _. Definition proj_ratio : ratio -> R := @proj1_sig _ _. -Instance proj_ratio_compat : Proper (equiv ==> equiv) proj_ratio. +Lemma proj_ratio_inj : injective equiv equiv proj_ratio. +Proof using . apply proj1_sig_inj. Qed. + +Global Instance proj_ratio_compat : Proper (equiv ==> eq) proj_ratio. Proof. intros ? ? Heq. apply Heq. Qed. Coercion proj_ratio : ratio >-> R. @@ -75,22 +78,25 @@ abstract (split; solve [ apply Rplus_le_le_0_compat; apply ratio_bounds | now apply Rlt_le, Rnot_le_lt ]). Defined. -Instance add_ratio_compat : Proper (equiv ==> equiv ==> equiv) add_ratio. +Global Instance add_ratio_compat : Proper (equiv ==> equiv ==> equiv) add_ratio. Proof. intros [] [] ? [] [] ?. unfold add_ratio. simpl in *. subst. destruct_match; reflexivity. Qed. (** A strict ratio is a [ratio] that is neither [0] nor [1]. *) Definition strict_ratio := {x : R | 0 < x < 1}%R. -Instance strict_ratio_Setoid : Setoid ratio := sig_Setoid _. -Instance strict_ratio_EqDec : EqDec strict_ratio_Setoid := @sig_EqDec _ _ _ _. +Global Instance strict_ratio_Setoid : Setoid strict_ratio := sig_Setoid _. +Global Instance strict_ratio_EqDec : EqDec strict_ratio_Setoid := @sig_EqDec _ _ _ _. Definition proj_strict_ratio (x : strict_ratio) : ratio := let '(exist _ v Hv) := x in exist _ v (conj (Rlt_le _ _ (proj1 Hv)) (Rlt_le _ _ (proj2 Hv))). -Instance proj_strict_ratio_compat : Proper (equiv ==> equiv) proj_strict_ratio. +Global Instance proj_strict_ratio_compat : Proper (equiv ==> equiv) proj_strict_ratio. Proof. intros [] [] Heq. apply Heq. Qed. +Lemma proj_strict_ratio_inj : injective equiv equiv proj_strict_ratio. +Proof using . intros [][] H. cbn in H. subst. reflexivity. Qed. + Coercion proj_strict_ratio : strict_ratio >-> ratio. Lemma strict_ratio_bounds : forall r : strict_ratio, (0 < r < 1)%R. @@ -112,6 +118,9 @@ Notation "n '/sr' m" := (mk_strict_ratio n m ltac:(clear; abstract lia) ltac:(cl (only parsing, at level 10). Notation "n '/sr' m" := (mk_strict_ratio n m _ _) (at level 10, only printing). +Lemma proj_ratio_strict_ratio : forall x, proj_ratio (proj_strict_ratio x) = proj1_sig x. +Proof. intros []. reflexivity. Qed. + (** ** Trajectory **) (** Trajectories are paths inside the space. *) @@ -122,15 +131,15 @@ Record path T `{Setoid T}:= { path_compat :> Proper (equiv ==> equiv) path_f }. Arguments path_f {T} {_} _ _. -Instance path_Setoid T {S : Setoid T} : Setoid (path T). -simple refine {| equiv := fun x y => path_f x == y |}; try apply fun_equiv; auto; []. +Global Instance path_Setoid T {S : Setoid T} : Setoid (path T). +simple refine {| equiv := fun x y => path_f x == y |}; try apply fun_Setoid; auto; []. Proof. split. + intro. reflexivity. + intros ? ? ?. now symmetry. + intros ? ? ? ? ?. etransitivity; eauto. Defined. -Instance path_full_compat {T} `(Setoid T): Proper (equiv ==> equiv ==> equiv) path_f. +Global Instance path_full_compat {T} `(Setoid T): Proper (equiv ==> equiv ==> equiv) path_f. Proof. intros p p' Hp x y Hxy. transitivity (path_f p y). - now apply path_compat. @@ -145,7 +154,7 @@ refine (Build_path _ _ (fun x => f (p x)) _). Proof. intros x y Hxy. now apply Hf, path_compat. Defined. Arguments lift_path T U _ _ f _ p /. -Instance lift_path_compat {T U} {HT : Setoid T} {HU : Setoid U} : +Global Instance lift_path_compat {T U} {HT : Setoid T} {HU : Setoid U} : forall f (Hf : Proper (equiv ==> equiv) f), Proper (equiv ==> equiv) (@lift_path T U HT HU f Hf). Proof. repeat intro. simpl. auto. Qed. diff --git a/Util/SetoidDefs.v b/Util/SetoidDefs.v index 3a2f9953e291754e8393784ac8c1050f5adec479..7308e14c9e63574c2ae9873dc518588678f34624 100644 --- a/Util/SetoidDefs.v +++ b/Util/SetoidDefs.v @@ -19,9 +19,7 @@ (**************************************************************************) -Require Import Rbase. -Require Import RelationPairs. -Require Import SetoidDec. +Require Import Utf8 Rbase RelationPairs SetoidDec. Require Import Pactole.Util.Preliminary. Set Implicit Arguments. @@ -29,31 +27,32 @@ Set Implicit Arguments. (* To avoid infinite loops, we use a breadth-first search... *) Typeclasses eauto := (bfs) 20. (* but we need to remove [eq_setoid] as it matches everything... *) -Global Remove Hints eq_setoid : Setoid. +#[export] Remove Hints eq_setoid : Setoid. (* while still declaring it for the types for which we still want to use it. *) -Instance R_Setoid : Setoid R := eq_setoid R. -Instance Z_Setoid : Setoid Z := eq_setoid Z. -Instance nat_Setoid : Setoid nat := eq_setoid nat. -Instance bool_Setoid : Setoid bool := eq_setoid bool. -Instance unit_Setoid : Setoid unit := eq_setoid unit. - -Instance R_EqDec : EqDec R_Setoid := Req_EM_T. -Instance Z_EqDec : EqDec Z_Setoid := Z.eq_dec. -Instance nat_EqDec : EqDec nat_Setoid := Nat.eq_dec. -Instance bool_EqDec : EqDec bool_Setoid := Bool.bool_dec. -Instance unit_EqDec : EqDec unit_Setoid := fun x y => match x, y with tt, tt => left eq_refl end. +#[export]Instance R_Setoid : Setoid R := eq_setoid R. +#[export]Instance Z_Setoid : Setoid Z := eq_setoid Z. +#[export]Instance nat_Setoid : Setoid nat := eq_setoid nat. +#[export]Instance bool_Setoid : Setoid bool := eq_setoid bool. +#[export]Instance unit_Setoid : Setoid unit := eq_setoid unit. + +#[export]Instance R_EqDec : EqDec R_Setoid := Req_dec_T. +#[export]Instance Z_EqDec : EqDec Z_Setoid := Z.eq_dec. +#[export]Instance nat_EqDec : EqDec nat_Setoid := Nat.eq_dec. +#[export]Instance bool_EqDec : EqDec bool_Setoid := Bool.bool_dec. +#[export]Instance unit_EqDec : EqDec unit_Setoid := fun x y => match x, y with tt, tt => left eq_refl end. Notation "x == y" := (equiv x y). +Notation "x == y :> A" := (@equiv A _ x y) (at level 70, y at next level, no associativity, only parsing). Arguments complement A R x y /. -Arguments Proper {A}%type R%signature m. +Arguments Proper {A}%_type R%_signature m. Lemma equiv_dec_refl {T U} {S : Setoid T} {E : EqDec S} : forall (e : T) (A B : U), (if equiv_dec e e then A else B) = A. -Proof using . intros. destruct_match; intuition. Qed. +Proof using . intros. destruct_match; intuition auto with *. Qed. (** ** Setoid Definitions **) -Instance fun_equiv A B `(Setoid B) : Setoid (A -> B) | 4. +Instance fun_Setoid A B `(Setoid B) : Setoid (A -> B) | 4. Proof. exists (fun f g : A -> B => forall x, f x == g x). split. + repeat intro. reflexivity. @@ -61,10 +60,23 @@ split. + repeat intro. etransitivity; eauto. Defined. -Instance fun_equiv_pointwise_compat A B `{Setoid B} : - subrelation (@equiv _ (fun_equiv A _)) (pointwise_relation _ equiv). +#[export]Instance fun_Setoid_respectful_compat (A B : Type) (SB : Setoid B) : + subrelation (@equiv (A->B) (fun_Setoid A SB)) + (respectful (@eq A) (@equiv B SB)). +Proof using . intros ?? H ???. subst. apply H. Qed. + +#[export]Instance fun_Setoid_pointwise_compat A B `{Setoid B} : + subrelation (@equiv _ (fun_Setoid A _)) (pointwise_relation _ equiv). Proof using . intros f g Hfg x. apply Hfg. Qed. +#[export]Instance eq_Setoid_eq_compat (T : Type) : + subrelation (@equiv T (eq_setoid T)) (@eq T). +Proof using . apply subrelation_refl. Qed. + +#[export]Instance eq_eq_Setoid_compat (T : Type) : + subrelation (@eq T) (@equiv T (eq_setoid T)). +Proof using . apply subrelation_refl. Qed. + Notation "x =?= y" := (equiv_dec x y) (at level 70, no associativity). (** Lifting an equivalence relation to an option type. *) @@ -75,34 +87,125 @@ Definition opt_eq {T} (eqT : T -> T -> Prop) (xo yo : option T) := | Some x, Some y => eqT x y end. -Instance opt_eq_refl : forall T (R : T -> T -> Prop), Reflexive R -> Reflexive (opt_eq R). +#[export]Instance opt_eq_refl : ∀ T (R : T -> T -> Prop), Reflexive R -> Reflexive (opt_eq R). Proof using . intros T R HR [x |]; simpl; auto. Qed. -Instance opt_eq_sym : forall T (R : T -> T -> Prop), Symmetric R -> Symmetric (opt_eq R). +#[export]Instance opt_eq_sym : ∀ T (R : T -> T -> Prop), Symmetric R -> Symmetric (opt_eq R). Proof using . intros T R HR [x |] [y |]; simpl; auto. Qed. -Instance opt_eq_trans : forall T (R : T -> T -> Prop), Transitive R -> Transitive (opt_eq R). +#[export]Instance opt_eq_trans : ∀ T (R : T -> T -> Prop), Transitive R -> Transitive (opt_eq R). Proof using . intros T R HR [x |] [y |] [z |]; simpl; intros; eauto; contradiction. Qed. -Instance opt_equiv T eqT (HeqT : @Equivalence T eqT) : Equivalence (opt_eq eqT). +#[export]Instance opt_equiv T eqT (HeqT : @Equivalence T eqT) : Equivalence (opt_eq eqT). Proof using . split; auto with typeclass_instances. Qed. -Instance opt_Setoid T (S : Setoid T) : Setoid (option T) := {| equiv := opt_eq equiv |}. +#[export]Instance opt_Setoid T (S : Setoid T) : Setoid (option T) := {| equiv := opt_eq equiv |}. -Instance Some_compat `(Setoid) : Proper (equiv ==> @equiv _ (opt_Setoid _)) Some. +#[export]Instance Some_compat `(Setoid) : Proper (equiv ==> @equiv _ (opt_Setoid _)) Some. Proof using . intros ? ? Heq. apply Heq. Qed. -Instance prod_Setoid : forall A B, Setoid A -> Setoid B -> Setoid (A * B) := - Pactole.Util.FMaps.FMapInterface.prod_Setoid. -Instance prod_EqDec A B `(EqDec A) `(EqDec B) : EqDec (@prod_Setoid A B _ _) := - Pactole.Util.FMaps.FMapInterface.prod_EqDec _ _. -Arguments prod_EqDec [A] [B] {_} _ {_} _. +Lemma Some_inj {T : Type} {ST : Setoid T} : + injective (@equiv _ ST) (@equiv _ (opt_Setoid ST)) Some. +Proof using . intros t1 t2 H. exact H. Qed. + +Lemma Some_eq_inj {T : Type} : @injective T _ eq eq Some. +Proof using . intros t1 t2 H1. inversion_clear H1. reflexivity. Qed. + +Lemma not_None {T : Type} {ST : Setoid T} : + ∀ ot : option T, ¬ (@equiv _ (opt_Setoid ST) ot None) + -> {t : T | @equiv _ (opt_Setoid ST) ot (Some t)}. +Proof using . + intros * H. destruct ot as [t|]. constructor 1 with t. + reflexivity. contradict H. reflexivity. +Qed. + +Lemma not_Some {T : Type} {ST : Setoid T} : ∀ ot : option T, + (∀ t : T, ¬ (@equiv _ (opt_Setoid ST) ot (Some t))) + -> @equiv _ (opt_Setoid ST) ot None. +Proof using . + intros * H. destruct ot as [t|]. exfalso. apply (H t). all: reflexivity. +Qed. + +Lemma not_None_iff {T : Type} {ST : Setoid T} : + ∀ ot : option T, ¬ (@equiv _ (opt_Setoid ST) ot None) + <-> ∃ t : T, @equiv _ (opt_Setoid ST) ot (Some t). +Proof using . + intros. etransitivity. 2: split. 2: apply inhabited_sig_to_exists. + 2: apply exists_to_inhabited_sig. split. + - intros H. constructor. apply not_None, H. + - intros [[t H]] Hn. rewrite H in Hn. inversion Hn. +Qed. + +Lemma option_decidable {T : Type} {ST : Setoid T} : ∀ ot : option T, + Decidable.decidable (@equiv _ (opt_Setoid ST) ot None). +Proof using . + intros. destruct ot as [t|]. right. intros H. inversion H. left. reflexivity. +Qed. + +Lemma pick_Some_None : + ∀ {T : Type} {ST : Setoid T} (P : T -> Prop) (oT : option T), + (∀ t : T, @equiv _ (opt_Setoid ST) oT (Some t) <-> P t) + -> (@equiv _ (opt_Setoid ST) oT None <-> (∀ t : T, ¬ (P t))) + -> pick_spec P oT. +Proof using . + intros T ST * Hs Hn. destruct oT as [t|]. apply Pick, Hs. + 2: apply Nopick, Hn. all: reflexivity. +Qed. + +Lemma opt_Setoid_eq : ∀ {T : Type} {ST : Setoid T} (o1 o2 : option T), + (∀ t1 t2 : T, @equiv _ ST t1 t2 <-> t1 = t2) + -> @equiv _ (opt_Setoid ST) o1 o2 <-> o1 = o2. +Proof using . + intros * H1. destruct o1 as [t1|], o2 as [t2|]. + 4: split; intros; reflexivity. 2,3: split; intros H2; inversion H2. + rewrite (injective_compat_iff Some_inj), (injective_compat_iff Some_eq_inj). + apply H1. +Qed. + +(* This Setoid could be written using RelProd in which case, + revealing the "and" inside would be more troublesome than + with this simple definition. *) +#[export]Instance prod_Setoid A B (SA : Setoid A) (SB : Setoid B) : Setoid (A * B). +simple refine {| equiv := fun xn yp => fst xn == fst yp /\ snd xn == snd yp |}; auto; []. +Proof. split. ++ repeat intro; now split. ++ repeat intro; split; now symmetry. ++ intros ? ? ? [? ?] [? ?]; split; etransitivity; eauto. +Defined. + +#[export]Instance prod_EqDec A B (SA : Setoid A) (SB : Setoid B) (EDA : EqDec SA) (EDB : EqDec SB) + : EqDec (prod_Setoid SA SB). +refine (fun xn yp => if equiv_dec (fst xn) (fst yp) then + if equiv_dec (snd xn) (snd yp) then left _ else right _ + else right _). +Proof. +- now split. +- abstract (intros [? ?]; contradiction). +- abstract (intros [? ?]; contradiction). +Defined. + +#[export]Instance fst_compat (A B : Type) (SA : Setoid A) (SB : Setoid B) : + Proper (@equiv (A*B) (prod_Setoid SA SB) ==> @equiv A SA) fst. +Proof using . intros ?? [H _]. exact H. Qed. -Instance fst_compat {A B} : forall R S, Proper (R * S ==> R) (@fst A B) := fst_compat. -Instance snd_compat {A B} : forall R S, Proper (R * S ==> S) (@snd A B) := snd_compat. +#[export]Instance snd_compat (A B : Type) (SA : Setoid A) (SB : Setoid B) : + Proper (@equiv (A*B) (prod_Setoid SA SB) ==> @equiv B SB) snd. +Proof using . intros ?? [_ H]. exact H. Qed. + +#[export]Instance pair_compat (A B : Type) (SA : Setoid A) (SB : Setoid B) : + Proper (@equiv A SA ==> @equiv B SB ==> @equiv (A*B) (prod_Setoid SA SB)) pair. +Proof using . intros ?? Hf ?? Hs. split; cbn. exact Hf. exact Hs. Qed. + +Lemma prod_Setoid_eq : ∀ {A B : Type} {SA : Setoid A} {SB : Setoid B} + (p1 p2 : A * B), (∀ a1 a2 : A, @equiv _ SA a1 a2 <-> a1 = a2) + -> (∀ b1 b2 : B, @equiv _ SB b1 b2 <-> b1 = b2) + -> @equiv _ (prod_Setoid SA SB) p1 p2 <-> p1 = p2. +Proof using . + intros * Ha Hb. rewrite <- pair_eqE, <- Ha, <- Hb. reflexivity. +Qed. (* Setoid over [sig] types *) -Instance sig_Setoid {T} (S : Setoid T) {P : T -> Prop} : Setoid (sig P). +#[export]Instance sig_Setoid {T} (S : Setoid T) {P : T -> Prop} : Setoid (sig P). simple refine {| equiv := fun x y => proj1_sig x == proj1_sig y |}; auto; []. Proof. split. + intro. reflexivity. @@ -110,10 +213,18 @@ Proof. split. + intros ? ? ? ? ?. etransitivity; eauto. Defined. -Instance sig_EqDec {T} {S : Setoid T} (E : EqDec S) (P : T -> Prop) : EqDec (@sig_Setoid T S P). +#[export]Instance sig_EqDec {T} {S : Setoid T} (E : EqDec S) (P : T -> Prop) : EqDec (@sig_Setoid T S P). Proof. intros ? ?. simpl. apply equiv_dec. Defined. -Instance sigT_Setoid {T} (S : Setoid T) {P : T -> Type} : Setoid (sigT P). +#[export]Instance proj1_sig_compat {T} {S : Setoid T} (E : EqDec S) (P : T -> Prop) : + Proper (@equiv _ (sig_Setoid S) ==> equiv) (@proj1_sig T P). +Proof using . intros ?? H. apply H. Qed. + +Lemma proj1_sig_inj : ∀ {T : Type} {S : Setoid T} (P : T -> Prop), + injective (@equiv (sig P) (sig_Setoid S)) (@equiv T S) (@proj1_sig T P). +Proof using . intros * ?? H. cbn. apply H. Qed. + +#[export]Instance sigT_Setoid {T} (S : Setoid T) {P : T -> Type} : Setoid (sigT P). simple refine {| equiv := fun x y => projT1 x == projT1 y |}; auto; []. Proof. split. + intro. reflexivity. @@ -121,9 +232,13 @@ Proof. split. + intros ? ? ? ? ?. etransitivity; eauto. Defined. -Instance sigT_EqDec {T} {S : Setoid T} (E : EqDec S) (P : T -> Type) : EqDec (@sigT_Setoid T S P). +#[export]Instance sigT_EqDec {T} {S : Setoid T} (E : EqDec S) (P : T -> Type) : EqDec (@sigT_Setoid T S P). Proof. intros ? ?. simpl. apply equiv_dec. Defined. +#[export]Instance projT1_compat {T} {S : Setoid T} (E : EqDec S) (P : T -> Type) : + Proper (@equiv _ (sigT_Setoid S) ==> equiv) (@projT1 T P). +Proof using . intros ?? H. apply H. Qed. + (** The intersection of equivalence relations is still an equivalence relation. *) Lemma inter_equivalence T R1 R2 (E1 : Equivalence R1) (E2 : Equivalence R2) : Equivalence (fun x y : T => R1 x y /\ R2 x y). @@ -140,7 +255,7 @@ Definition inter_Setoid {T} (S1 : Setoid T) (S2 : Setoid T) : Setoid T := {| Definition inter_EqDec {T} {S1 S2 : Setoid T} (E1 : EqDec S1) (E2 : EqDec S2) : EqDec (inter_Setoid S1 S2). -Proof. +Proof using . intros x y. destruct (E1 x y), (E2 x y); (now left; split) || (right; intros []; contradiction). Defined. diff --git a/Util/Stream.v b/Util/Stream.v index ed419ec827036305277a186ff95653fdd0186a44..8538667ae9773b7f0d17d3071f2a6f4ce11f865a 100644 --- a/Util/Stream.v +++ b/Util/Stream.v @@ -64,34 +64,25 @@ Inductive eventually (P : t A -> Prop) (s : t A) : Prop := Global Arguments Now [P] [s] _. Global Arguments Later [P] [s] _. -(** Logical operators on relations over streams. *) - -Definition instant2 (P : A -> A -> Prop) := fun s1 s2 => P (hd s1) (hd s2). -Global Arguments instant2 P s1 s2 /. - -CoInductive forever2 (P : t A -> t A -> Prop) (s1 s2 : t A) : Prop := - Always2 : P s1 s2 -> forever2 P (tl s1) (tl s2) -> forever2 P s1 s2. -Global Arguments Always2 [P] [s1] [s2] _ _. - -Inductive eventually2 (P : t A -> t A -> Prop) (s1 s2 : t A) : Prop := - | Now2 : P s1 s2 -> eventually2 P s1 s2 - | Later2 : eventually2 P (tl s1) (tl s2) -> eventually2 P s1 s2. -Global Arguments Now2 [P] [s1] [s2] _. -Global Arguments Later2 [P] [s1] [s2] _. +Lemma instant_impl_compat : forall P Q : A -> Prop, + (forall s, P s -> Q s) -> forall s, instant P s -> instant Q s. +Proof using . unfold instant. auto. Qed. +Lemma forever_impl_compat : forall P Q : t A -> Prop, + (forall s, P s -> Q s) -> forall s, forever P s -> forever Q s. +Proof using . +cofix Hrec. intros P Q HPQ s HP. constructor. +- inv HP. auto. +- inv HP. eapply Hrec; eauto. +Qed. -(** Extensional equality on streams, up to a setoid equality on the stream elements. *) -Global Instance stream_Setoid : Setoid (t A). -Proof. -exists (forever2 (instant2 equiv)). split. -+ coinduction Heq. simpl. reflexivity. -+ coinduction Heq; simpl. - - symmetry. match goal with H : forever2 _ _ _ |- _ => now inv H end. - - match goal with H : forever2 _ _ _ |- _ => now inv H end. -+ coinduction Heq; simpl. - - transitivity (hd y); match goal with H : forever2 _ _ _ |- _ => now inv H end. - - transitivity (tl y); match goal with H : forever2 _ _ _ |- _ => now inv H end. -Defined. +Lemma eventually_impl_compat : forall P Q : t A -> Prop, + (forall s, P s -> Q s) -> forall s, eventually P s -> eventually Q s. +Proof using . +intros P Q HPQ s HP. induction HP as [e HP | HP]. +- apply Now. auto. +- now apply Later. +Qed. (** Some sanity check on [constant] and [alternate]. *) Lemma constant_hd : forall c : A, hd (constant c) = c. @@ -109,11 +100,122 @@ Proof using . reflexivity. Qed. Lemma alternate_tl_hd : forall c1 c2 : A, hd (tl (alternate c1 c2)) = c2. Proof using . reflexivity. Qed. -(** Alternative caracterisation of [alternate]. *) -Lemma alternate_tl : forall c1 c2 : A, tl (alternate c1 c2) == alternate c2 c1. -Proof using . cofix alt. constructor; [| constructor]; simpl; try reflexivity; []. apply alt. Qed. +Lemma and_forever: forall (s : t A) P Q, + forever P s -> + forever Q s -> + forever (fun x => P x /\ Q x) s. +Proof using Type. + cofix cof. + intros s P Q HforeverP HforeverQ. + destruct HforeverP. + destruct HforeverQ. + constructor. + - split;assumption. + - apply cof;assumption. +Qed. + +Lemma instant_and_forever: forall (s : t A) P Q, + forever (instant P) s -> + forever (instant Q) s -> + forever (instant (fun x => P x /\ Q x)) s. +Proof using Type. + cofix cof. + intros s P Q HforeverP HforeverQ. + destruct HforeverP. + destruct HforeverQ. + constructor. + - split;assumption. + - apply cof;assumption. +Qed. + +Inductive until (P Q : t A -> Prop) (s : t A) : Prop := + | NotYet : P s -> until P Q (tl s) -> until P Q s + | YesNow : Q s -> until P Q s. + +Definition weak_until P Q s := forever P s \/ until P Q s. + +End Streams. + +(** Logical operators on relations over streams *) + +Definition instant2 {A B : Type} (P : A -> B -> Prop) := fun s1 s2 => P (hd s1) (hd s2). +Global Arguments instant2 {A} {B} P s1 s2 /. + +CoInductive forever2 {A B : Type} (P : t A -> t B -> Prop) (s1 : t A) (s2 : t B) : Prop := + Always2 : P s1 s2 -> forever2 P (tl s1) (tl s2) -> forever2 P s1 s2. +Global Arguments Always2 {A} {B} [P] [s1] [s2] _ _. + +Inductive eventually2 {A B : Type} (P : t A -> t B -> Prop) (s1 : t A) (s2 : t B) : Prop := + | Now2 : P s1 s2 -> eventually2 P s1 s2 + | Later2 : eventually2 P (tl s1) (tl s2) -> eventually2 P s1 s2. +Global Arguments Now2 {A} {B} [P] [s1] [s2] _. +Global Arguments Later2 {A} {B} [P] [s1] [s2] _. + +(** Relation properties lifted to steams. *) +Section RelationOperators. +Context {A B : Type}. +Context {SA : Setoid A}. +Context {SB : Setoid B}. + +Global Instance instant2_refl (R : A -> A -> Prop) `{Reflexive _ R} : Reflexive (instant2 R). +Proof using . intro. simpl. reflexivity. Qed. + +Global Instance instant2_sym (R : A -> A -> Prop) `{Symmetric _ R} : Symmetric (instant2 R). +Proof using . intros x y HR. simpl in *. now symmetry. Qed. + +Global Instance instant2_trans (R : A -> A -> Prop) `{Transitive _ R} : Transitive (instant2 R). +Proof using . intros x y z ? ?. simpl in *. now transitivity (hd y). Qed. + +Global Instance forever2_refl (R : t A -> t A -> Prop) `{Reflexive _ R} : Reflexive (forever2 R). +Proof using . coinduction Hrec. reflexivity. Qed. + +Global Instance forever2_sym (R : t A -> t A -> Prop) `{Symmetric _ R} : Symmetric (forever2 R). +Proof using . +coinduction Hrec; match goal with H : forever2 _ _ _ |- _ => inv H end. +- now symmetry. +- assumption. +Qed. + +Global Instance forever2_trans (R : t A -> t A -> Prop) `{Transitive _ R} : Transitive (forever2 R). +Proof using . +coinduction Hrec; match goal with H : forever2 _ _ _, H' : forever2 _ _ _ |- _ => inv H; inv H' end. +- now transitivity y. +- now transitivity (tl y). +Qed. + + +Global Instance eventually2_refl (R : t A -> t A -> Prop) `{Reflexive _ R} : Reflexive (eventually2 R). +Proof using . intro x. apply Now2. reflexivity. Qed. + +Global Instance eventually2_sym (R : t A -> t A -> Prop) `{Symmetric _ R} : Symmetric (eventually2 R). +Proof using . +intros x y HR. induction HR. +- apply Now2. now symmetry. +- now apply Later2. +Qed. + +(* It does not apprear to be transitive for lack of synchronisation of the streams. *) +Instance eventually2_trans (R : t A -> t A -> Prop) `{Transitive _ R} : Transitive (eventually2 R). +Proof. Abort. +End RelationOperators. + +(** Extensional equality on streams, up to a setoid equality on the stream elements. *) +Global Instance stream_Setoid {A : Type} `{Setoid A} : Setoid (t A). +Proof. +exists (forever2 (instant2 equiv)). split. ++ coinduction Heq. simpl. reflexivity. ++ coinduction Heq; simpl. + - symmetry. match goal with H : forever2 _ _ _ |- _ => now inv H end. + - match goal with H : forever2 _ _ _ |- _ => now inv H end. ++ coinduction Heq; simpl. + - transitivity (hd y); match goal with H : forever2 _ _ _ |- _ => now inv H end. + - transitivity (tl y); match goal with H : forever2 _ _ _ |- _ => now inv H end. +Defined. (** Compatibility lemmas. *) +Section EquivalenceProperties. +Context {A : Type}. +Context {SA : Setoid A}. Global Instance hd_compat : Proper (equiv ==> equiv) hd. Proof using . intros s s' Hs. now inv Hs. Qed. @@ -121,13 +223,13 @@ Proof using . intros s s' Hs. now inv Hs. Qed. Global Instance tl_compat : Proper (equiv ==> equiv) tl. Proof using . intros s s' Hs. now inv Hs. Qed. -Global Instance constant_compat : Proper (equiv ==> equiv) constant. +Global Instance constant_compat : Proper (@equiv A _ ==> equiv) constant. Proof using . unfold constant. now coinduction Heq. Qed. Global Instance aternate_compat : Proper (@equiv A _ ==> equiv ==> equiv) alternate. Proof using . cofix Heq. do 2 (constructor; trivial). cbn. now apply Heq. Qed. -Global Instance instant_compat : Proper ((equiv ==> iff) ==> equiv ==> iff) instant. +Global Instance instant_compat : Proper ((@equiv A _ ==> iff) ==> equiv ==> iff) instant. Proof using . intros P Q HPQ s s' Hs. unfold instant. apply HPQ, Hs. Qed. Global Instance forever_compat : Proper ((equiv ==> iff) ==> equiv ==> iff) forever. @@ -152,11 +254,18 @@ intros P Q HPQ s s' Hs. split; intro eventually. - apply Later. apply IHeventually. now inv Hs. Qed. +(** Alternative caracterisation of [alternate]. *) +Lemma alternate_tl : forall c1 c2 : A, tl (alternate c1 c2) == alternate c2 c1. +Proof using . cofix alt. constructor; [| constructor]; simpl; try reflexivity; []. apply alt. Qed. -Global Instance instant2_compat : Proper ((equiv ==> equiv ==> iff) ==> equiv ==> equiv ==> iff) instant2. +Context {B : Type}. +Context {SB : Setoid B}. + +Global Instance instant2_compat : Proper ((@equiv A _ ==> @equiv B _ ==> iff) ==> equiv ==> equiv ==> iff) instant2. Proof using . intros P Q HPQ ? ? Hs1 ? ? Hs2. unfold instant2. apply HPQ; apply Hs1 || apply Hs2. Qed. -Global Instance forever2_compat : Proper ((equiv ==> equiv ==> iff) ==> equiv ==> equiv ==> iff) forever2. +Global Instance forever2_compat : + Proper ((@equiv (t A) _ ==> @equiv (t B) _ ==> iff) ==> equiv ==> equiv ==> iff) forever2. Proof using . intros P Q HPQ s1 s1' Hs1 s2 s2' Hs2. split. + revert s1 s1' Hs1 s2 s2' Hs2. coinduction Hrec. @@ -167,7 +276,7 @@ intros P Q HPQ s1 s1' Hs1 s2 s2' Hs2. split. - inv Hs1. inv Hs2. match goal with H : forever2 Q _ _ |- _ => inv H end. eapply Hrec; eassumption. Qed. -Global Instance eventually2_compat : Proper ((equiv ==> equiv ==> iff) ==> equiv ==> equiv ==> iff) eventually2. +Global Instance eventually2_compat : Proper ((@equiv (t A) _ ==> @equiv (t B) _ ==> iff) ==> equiv ==> equiv ==> iff) eventually2. Proof using . intros P Q HPQ s1 s1' Hs1 s2 s2' Hs2. split; intro eventually2. + revert s1' s2' Hs1 Hs2. induction eventually2; intros s1' s2' Hs1 Hs2. @@ -178,32 +287,11 @@ intros P Q HPQ s1 s1' Hs1 s2 s2' Hs2. split; intro eventually2. - apply Later2. inv Hs1. inv Hs2. now apply IHeventually2. Qed. - -Lemma instant_impl_compat : forall P Q : A -> Prop, - (forall s, P s -> Q s) -> forall s, instant P s -> instant Q s. -Proof using . unfold instant. auto. Qed. - -Lemma forever_impl_compat : forall P Q : t A -> Prop, - (forall s, P s -> Q s) -> forall s, forever P s -> forever Q s. -Proof using . -cofix Hrec. intros P Q HPQ s HP. constructor. -- inv HP. auto. -- inv HP. eapply Hrec; eauto. -Qed. - -Lemma eventually_impl_compat : forall P Q : t A -> Prop, - (forall s, P s -> Q s) -> forall s, eventually P s -> eventually Q s. -Proof using . -intros P Q HPQ s HP. induction HP as [e HP | HP]. -- apply Now. auto. -- now apply Later. -Qed. - -Lemma instant2_impl_compat : forall P Q : A -> A -> Prop, +Lemma instant2_impl_compat : forall P Q : A -> B -> Prop, (forall s1 s2, P s1 s2 -> Q s1 s2) -> forall s1 s2, instant2 P s1 s2 -> instant2 Q s1 s2. Proof using . unfold instant2. auto. Qed. -Lemma forever2_impl_compat : forall P Q : t A -> t A -> Prop, +Lemma forever2_impl_compat : forall P Q : t A -> t B -> Prop, (forall s1 s2, P s1 s2 -> Q s1 s2) -> forall s1 s2, forever2 P s1 s2 -> forever2 Q s1 s2. Proof using . cofix Hrec. intros P Q HPQ s1 s2 HP. constructor. @@ -211,65 +299,14 @@ cofix Hrec. intros P Q HPQ s1 s2 HP. constructor. - inv HP. eapply Hrec; eauto. Qed. -Lemma eventually2_impl_compat : forall P Q : t A -> t A -> Prop, +Lemma eventually2_impl_compat : forall P Q : t A -> t B -> Prop, (forall s1 s2, P s1 s2 -> Q s1 s2) -> forall s1 s2, eventually2 P s1 s2 -> eventually2 Q s1 s2. Proof using . intros P Q HPQ s1 s2 HP. induction HP as [e1 e2 HP | HP]. - apply Now2. auto. - now apply Later2. Qed. - -(** Relation properties lifted to steams. *) - -Global Instance instant2_refl R `{Reflexive _ R} : Reflexive (instant2 R). -Proof using . intro. simpl. reflexivity. Qed. - -Global Instance instant2_sym R `{Symmetric _ R} : Symmetric (instant2 R). -Proof using . intros x y HR. simpl in *. now symmetry. Qed. - -Global Instance instant2_trans R `{Transitive _ R} : Transitive (instant2 R). -Proof using . intros x y z ? ?. simpl in *. now transitivity (hd y). Qed. - - -Global Instance forever2_refl R `{Reflexive _ R} : Reflexive (forever2 R). -Proof using . coinduction Hrec. reflexivity. Qed. - -Global Instance forever2_sym R `{Symmetric _ R} : Symmetric (forever2 R). -Proof using . -coinduction Hrec; match goal with H : forever2 _ _ _ |- _ => inv H end. -- now symmetry. -- assumption. -Qed. - -Global Instance forever2_trans R `{Transitive _ R} : Transitive (forever2 R). -Proof using . -coinduction Hrec; match goal with H : forever2 _ _ _, H' : forever2 _ _ _ |- _ => inv H; inv H' end. -- now transitivity y. -- now transitivity (tl y). -Qed. - - -Global Instance eventually2_refl R `{Reflexive _ R} : Reflexive (eventually2 R). -Proof using . intro x. apply Now2. reflexivity. Qed. - -Global Instance eventually2_sym R `{Symmetric _ R} : Symmetric (eventually2 R). -Proof using . -intros x y HR. induction HR. -- apply Now2. now symmetry. -- now apply Later2. -Qed. - -(* It does not apprear to be transitive for lack of synchronisation of the streams. *) -Instance eventually2_trans R `{Transitive _ R} : Transitive (eventually2 R). -Proof. Abort. - - -Inductive until (P Q : t A -> Prop) (s : t A) : Prop := - | NotYet : P s -> until P Q (tl s) -> until P Q s - | YesNow : Q s -> until P Q s. - -Definition weak_until P Q s := forever P s \/ until P Q s. -End Streams. +End EquivalenceProperties. (** Map on streams *) Section Map. diff --git a/_CoqProject b/_CoqProject index 9a763c0e1cbf23581aa8aee0b82b7c2b3d504a01..1caec51725b2cc0b53e0d0bf204cc5c2f352bca8 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,4 +1,4 @@ --install none +# -install none # currently this is already the default with coq_makefile -f, and generates awarning -Q . Pactole -arg -w -arg -notation-overridden @@ -6,13 +6,17 @@ # This warning appeared in coq 8.13. We can fix it by putting # "#[export]" in front of Hint commands. But it wouldn't compile in # <8.12. We only support >= 8.12 now so this is nont needed anymore. -#-arg -w -#-arg -deprecated-hint-without-locality +-arg -w +-arg -deprecated-hint-without-locality # Spaces in args are not supported, see Makefile.local #-arg -set #-arg "Suggest Proof Using=true" +## All dependencies that are not case studies +## Typically this should be the target of opam package +Pactole_all.v + ## External libraries #Util/FSets/OrderedType.v Util/FSets/FSetInterface.v @@ -34,6 +38,8 @@ Util/NumberComplements.v Util/ListComplements.v Util/Coqlib.v Util/Bijection.v +Util/Fin.v +Util/Enum.v Util/Stream.v Util/Lexprod.v Util/Ratio.v @@ -50,20 +56,24 @@ Models/Rigid.v Models/RigidFlexibleEquivalence.v Models/RigidFlexibleEquivalence_Assumptions.v Models/DiscreteGraph.v +Models/Isometry.v Models/Similarity.v Models/ContinuousGraph.v Models/GraphEquivalence.v +Models/RingSSync.v ## Spaces Spaces/RealVectorSpace.v Spaces/RealMetricSpace.v Spaces/RealNormedSpace.v Spaces/EuclideanSpace.v +Spaces/Isometry.v Spaces/Similarity.v Spaces/R.v Spaces/R2.v Spaces/Graph.v Spaces/Isomorphism.v +Spaces/ThresholdIsomorphism.v Spaces/Ring.v Spaces/Grid.v @@ -100,11 +110,14 @@ CaseStudies/Gathering/InR2/Viglietta.v CaseStudies/Gathering/InR2/Viglietta_Assumptions.v ## Case Study: Ring Exploration -CaseStudies/Exploration/Definitions.v +CaseStudies/Exploration/ExplorationDefs.v CaseStudies/Exploration/ImpossibilityKDividesN.v CaseStudies/Exploration/ImpossibilityKDividesN_Assumptions.v CaseStudies/Exploration/Tower.v CaseStudies/Exploration/Tower_Assumptions.v ## Case Study: Volume -CaseStudies/Volumes/NoCollisionAndPath.v +CaseStudies/LifeLine/Algorithm.v + +## Minipactole +minipactole/minipactole.v diff --git a/codemeta.json b/codemeta.json new file mode 100644 index 0000000000000000000000000000000000000000..de545b51aa98ebf74a62f91069fbf722824202a4 --- /dev/null +++ b/codemeta.json @@ -0,0 +1,69 @@ +{ + "@context": "https://doi.org/10.5063/schema/codemeta-2.0", + "@type": "SoftwareSourceCode", + "name": "Pactole", + "license": "https://spdx.org/licenses/LGPL-3.0", + "codeRepository": "https://gitlab.liris.cnrs.fr/pactole/coq-pactole", + "relatedLink": "TODO? https://???", + "dateCreated": "TODO 2013-??-??", + "datePublished": "TODO 2024-???", + "version": "TODO 2.0???", + "description": "TODO A Coq library formalizing the distributed computing model for robot swarm called Look-Compute-Move (due to Suzuki and Yamashita) and several case studies in this model. ", + "applicationCategory": "FIXME: Formal development", + "releaseNotes": "TODO Second release.", + "runtimePlatform": "Coq-8.20", + "developmentStatus": "active", + "funding": [ + { "@type":"Grant", + "identifier": "Digiteo Project #2009-38HD" } , + { "@type":"Grant", + "identifier": "ANR Project 2019-CE25-0005" } + ], + "keywords": [ + "distributed systems", + "formal proof", + "robot swarm" + ], + "programmingLanguage": [ + "Coq/Rocq" + ], + "runtimePlatform": [ + "Coq/Rocq" + ], + "author": [ + { + "@type": "Person", + "@id": "https://orcid.org/0000-0001-8789-9781", + "givenName": "Pierre", + "familyName": "Courtieu", + "email": "Pierre.Courtieu@lecnam.net", + "affiliation": { + "@type": "Organization", + "name": "Conservatoire National des Arts et Métiers" + } + }, + { + "@type": "Person", + "@id": "https://orcid.org/0000-0001-7442-2538", + "givenName": "Xavier", + "familyName": "Urbain", + "email": "Xavier.Urbain@univ-lyon1.fr", + "affiliation": { + "@type": "Organization", + "name": "Université Claude Bernard Lyon 1" + } + } + { + "@type": "Person", + "givenName": "Lionel", + "familyName": "Rieg", + "email": "lionel.rieg@univ-grenoble-alpes.fr", + "affiliation": { + "@type": "Organization", + "name": "Verimag, Grenoble INP -- UGA" + } + } + ], + "contributor": "TODO: Les stagiaires comme ci-dessus?", + "citation": "TODO un papier décrivant Pactole" +} diff --git a/build_package.sh b/dev/build_package.sh similarity index 100% rename from build_package.sh rename to dev/build_package.sh diff --git a/dev/depgraph.sh b/dev/depgraph.sh new file mode 100755 index 0000000000000000000000000000000000000000..a89e298a68b2b347c8f992c45229b3073375c80d --- /dev/null +++ b/dev/depgraph.sh @@ -0,0 +1,19 @@ +#!/bin/bash + +OPT="-Q . Pactole" +VFILES=$(grep "^[^#].*\.v" _CoqProject) + +( echo "digraph interval_deps {" ; + echo "node [shape=ellipse, style=filled, URL=\"html/Interval.\N.html\", color=black];"; + ( coqdep $OPT $VFILES ) | + sed -n -e 's,/,.,g;s/[.]vo.*: [^ ]*[.]v//p' | + while read src dst; do + color=$(echo "$src" | sed -e 's,Real.*,turquoise,;s,Interval[.].*,plum,;s,Integral.*,lightcoral,;s,Poly.*,yellow,;s,Float.*,cornflowerblue,;s,Eval.*,green,;s,[A-Z].*,white,') + echo "\"$src\" [fillcolor=$color];" + for d in $dst; do + echo "\"$src\" -> \"${d%.vo}\" ;" + done + done; + echo "}" ) | tred > deps.dot +dot -T png deps.dot > deps.png +dot -T cmap deps.dot | sed -e 's,>$,/>,' > deps.map diff --git a/dev/rebuild_pactole_all.sh b/dev/rebuild_pactole_all.sh new file mode 100755 index 0000000000000000000000000000000000000000..eaccf5ac39d66509db26ff8c2400a8eed4ad14df --- /dev/null +++ b/dev/rebuild_pactole_all.sh @@ -0,0 +1,10 @@ + + + + +TOLERATED="Pactole_all.v\\|Util/FSets/OrderedType.v\\|CaseStudies/Volumes/NoCollisionAndPath.v" +VFILES=$(find . -name "*.v" | sed -e 's/\.\///' | grep -v "\.#\\|*~" | grep -v $TOLERATED | grep -v CaseStudies) + +for i in $VFILES ; do + echo $i | sed -e 's/\.v/./' | sed -e 's/\//./g' | sed -e 's/^/Require Import Pactole./' +done diff --git a/sanity-git.sh b/dev/sanity-git.sh similarity index 93% rename from sanity-git.sh rename to dev/sanity-git.sh index c6d79027eccc8855b01e20e882f908d911cfb7a2..75f5908b5329632714018f3c1ee2389774e59b53 100755 --- a/sanity-git.sh +++ b/dev/sanity-git.sh @@ -23,7 +23,7 @@ AUXFILES=".allvfiles .trackedfiles .untrackedvfiles .unfoundvfiles" # again? # this is a regexp, typically: # TOLERATED="Foo/Bar/toto.v\\|Baz/titi.v" -TOLERATED="Util/FSets/OrderedType.v" +TOLERATED="Util/FSets/OrderedType.v\\|CaseStudies/Volumes/NoCollisionAndPath.v" find . -name "*.v" | sed -e 's/\.\///' | grep -v "\.#\\|*~" | grep -v $TOLERATED > .allvfiles @@ -71,7 +71,7 @@ then else echo "${RED}Commit was aborted for the reason described above.${NC}" echo - echo "The_coqProject seems to be outdated wrt actual files in your" + echo "The _CoqProject file seems to be outdated wrt actual files in your" echo "working directory." echo echo "Please fix the problem above or use \"git commit --no-verify.\"" diff --git a/minipactole/minipactole.v b/minipactole/minipactole.v new file mode 100644 index 0000000000000000000000000000000000000000..aef316fbf954d6fa0d6932982b2b807184e030fd --- /dev/null +++ b/minipactole/minipactole.v @@ -0,0 +1,227 @@ +Require Import Reals Morphisms Setoid. +Require Import Decidable. +Require Import SetoidDec SetoidList SetoidClass. +Require Pactole.Util.Stream. +Require Import Pactole.Util.SetoidDefs Pactole.Util.NumberComplements. + +Open Scope R_scope. + +Inductive ident:Type := A | B. (* Nous avons seulement 2 robots. *) + +Definition position := (R*R)%type. (*les coordonnées dans le plan euclidien.*) + +Class State_info {info:Type}:Type := mkStI { }. (* Paramètre du modèle *) +Class State_pos {pos:Type}:Type := mkStP { }. (* Paramètre du modèle *) + +(* À partir des paramètres on a un type State. *) +Class State {I:Type} {L:Type} {info:@State_info I} {position:@State_pos L} + : Type := + mkSt { + pos: L; + info: I + }. +(* Point technique: le dernier argument de pos et info n'est pas implicite. *) +Arguments pos {I}%_type {L}%_type {info} {position} State. +Arguments info {I}%_type {L}%_type {info} {position} State. + +Module Minipactole. (* comme si c'était dans un autre fichier. *) + + Section minipactole. + (* On laisse Info comme paramètre *) + Context {info} {Info: @State_info info}. (* À définir ou bien faire varier selon + les algorithmes considérés: + couleur etc *) + (* Mais on fixe pos *) + Context {pos} {Pos: @State_pos pos}. + (* Local Instance Pos: State_pos := (mkStP position). *) + + (* config = fonction retournant l'état de n'importe quel robot. *) + Definition configuration := ident -> State. + + (* Comme une configuration est une fonction, l'égalité "=" de coq + ne convient pas: deux fonctions équivalentes mais définie + différemment ne seraient pas considérées comme "=". On définit + donc une autre égalité: "equiv". Deux configurations sont + "equiv" si elles ont le même comportement. Ceci devrait + permettre d'utiliser "rewrite" de façon plus agréable. *) + Local Instance configuration_Setoid : Setoid configuration := fun_Setoid ident _. + + (* Le robogram est le programme qui tourne localement sur le + robot: il prend la configuration (observation parfaite de la + configuration). et retourne le nouvel état (position + infos + couleur etc). Il est compatible avec le "equiv" de + configuration et de state. *) + Record robogram := { + pgm :> configuration -> ident -> State; + pgm_compat :> Proper (equiv ==> eq ==> eq) pgm + }. + + (* Une demonic_action détermine qui est activé au round courant. Le + démon sera une suite infinie de demonic_action. *) + Record demonic_action := { + activate : ident -> bool; (** Select which robots are activated *) + }. + + (** A [demon] is just a stream of [demonic_action]s. *) + Definition demon := Stream.t demonic_action. + + (* Un round consiste à fabriquer la configuration suivante à + partir de la configuration donnée (config) en appliquant le + robogram sur les robots activés. *) + Definition round (r:robogram) da config := + fun id => + if da.(activate) id then r config id (* on applique le robogram *) + else config id (* on garde le même état que dans config *) + . + + + Global Instance round_compat : Proper (equiv ==> equiv ==> equiv ==> equiv) round. + Proof using Type. + intros r1 r2 Hr da1 da2 Hda config1 config2 Hconfig id. + unfold round. rewrite Hda. + destruct (activate da2 id). + - (* active robot *) + destruct id. + + rewrite Hr. + apply pgm_compat;auto. + + rewrite Hr. + apply pgm_compat;auto. + - apply Hconfig. + Qed. + + (** Executions are simply streams of configurations. *) + Definition execution := Stream.t configuration. + + (** [execute r d config] returns an (infinite) execution from an + initial global configuration [config], a demon [d] and a + robogram [r] running on each good robot. *) + Definition execute (r : robogram) : demon -> configuration -> execution := + cofix execute d config := + Stream.cons config (execute (Stream.tl d) (round r (Stream.hd d) config)). + + End minipactole. +End Minipactole. + +(* On développe en parallèle l'exemple le plus simple *) +Section Vig2Cols. + + Inductive Couleur:Type := Black | White. + + Local Instance Pos: State_pos := (mkStP position). + Local Instance Info: State_info := (mkStI Couleur). + + Definition theother id := + match id with + A => B + | B => A + end. + + Lemma RRdec : forall x y : position, {x = y} + {x <> y}. + Proof. + intros. + case x;case y. + intros r r0 r1 r2. + case (Req_dec_T r1 r);intros; subst. + - case (Req_dec_T r2 r0);intros;cbn [equiv] in *;subst. + + left; reflexivity. + + right. + intro abs. + injection abs. + intros H. + contradiction. + - right. + intro abs. + injection abs. + intros H H0. + contradiction. + Qed. + + Definition middle (p1 p2: position) := ((fst p1 + fst p2)/2, (snd p1 + snd p2)/2). + + (* (BLACK, BLACK) -> WHITE, STAY + (BLACK, WHITE) -> skip + (WHITE, BLACK) -> –, M2O + (WHITE, WHITE) -> BLACK, M2H *) + Definition rVig2Cols (conf:Minipactole.configuration) id := + let me := conf id in + let other := conf (theother id) in + match me.(info),other.(info) with + Black,Black => {| pos := me.(pos) ; info := White |} (* (BLACK, BLACK) -> WHITE, STAY *) + | Black,White => me (* skip *) + | White,Black => {| pos := other.(pos) ; info := White |}(*white?*)(* (WHITE, BLACK) -> –, M2O *) + | White,White => {| pos := middle me.(pos) other.(pos) ; info := Black |} (* BLACK, M2H *) + end. + + Definition rbmVig2Cols : Minipactole.robogram. + Proof. + apply (Minipactole.Build_robogram rVig2Cols). + repeat red;simpl. + intros x y H x0 y0 H0. + rewrite H0. + unfold rVig2Cols. + repeat rewrite H. + subst. + reflexivity. + Defined. (* Important *) + + (* Exemple de théorème prouvable sur l'algo concret. *) + (* une fois le RDV atteint et A=White et B=black, personne ne bouge plus. *) + Lemma once_RDV1 : forall c da, + (c A).(info) = Black -> (c B).(info) = White -> + (c A).(pos) = (c B).(pos) -> + Minipactole.round rbmVig2Cols da c == c. + Proof. + intros c da HcA HcB HcPos. + simpl. + intros x. + unfold rbmVig2Cols , Minipactole.round;simpl. + destruct (Minipactole.activate da x);auto. + unfold rVig2Cols. + destruct x;simpl. + - rewrite HcA. + rewrite HcB. + reflexivity. + - rewrite HcA. + rewrite HcPos. + rewrite HcB at 1. + rewrite <- HcB. + destruct (c B);auto. + Qed. + + (* une fois le RDV atteint et A=White et B=black, personne ne bouge plus. *) + Lemma once_RDV2 : forall c da, + (c A).(info) = White -> (c B).(info) = Black -> + (c A).(pos) = (c B).(pos) -> + Minipactole.round rbmVig2Cols da c == c. + Proof. + intros c da HcA HcB HcPos. + simpl. + intros x. + unfold rbmVig2Cols , Minipactole.round;simpl. + destruct (Minipactole.activate da x);auto. + unfold rVig2Cols. + destruct x;simpl. + - rewrite HcA. + rewrite HcB. + rewrite <- HcA. + rewrite <- HcPos. + destruct (c A);auto. + - rewrite HcA. + rewrite HcB at 1. + reflexivity. + Qed. + + Lemma once_RDV: forall c da, + (c A).(info) <> (c B).(info) -> + (c A).(pos) = (c B).(pos) -> + Minipactole.round rbmVig2Cols da c == c. + Proof. + intros c da. + destruct ((c A).(info)) eqn:hA; + destruct ((c B).(info)) eqn:hB; intros; try contradiction. + - apply once_RDV1;auto. + - apply once_RDV2;auto. + Qed. + + +End Vig2Cols.