diff --git a/CaseStudies/Gathering/Definitions.v b/CaseStudies/Gathering/Definitions.v
index c144450b85e8f43e5dcaa5421c887b3c94df78f0..c36036550495c4449d146da111ec430c85a1d5ed 100644
--- a/CaseStudies/Gathering/Definitions.v
+++ b/CaseStudies/Gathering/Definitions.v
@@ -43,7 +43,7 @@ Context `{inactive_choice}.
 Context {UpdFun : update_function _ _ _}.
 Context {InaFun : inactive_function _}.
 
-Notation "!!" := (fun config => obs_from_config config origin).
+(* 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). *)
diff --git a/CaseStudies/Gathering/Impossibility.v b/CaseStudies/Gathering/Impossibility.v
index ac16573c37feb01dc71961afa10365f782bd526d..9d9be3812d417ada3d7822895f9e59a513e200ed 100644
--- a/CaseStudies/Gathering/Impossibility.v
+++ b/CaseStudies/Gathering/Impossibility.v
@@ -91,8 +91,8 @@ Instance InaFun : inactive_function unit := {
 
 (* 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).
-Notation "!! config" := (obs_from_config config origin) (at level 10).
+(* Notation obs_from_confi := (@obs_from_config _ _ _ _ multiset_observation). *)
+Notation "!! config" := (@obs_from_config _ _ _ _ multiset_observation config origin) (at level 10).
 
 (* FIXME: why is this instance not taken from Similarity from Definition *)
 Instance Frame : frame_choice (similarity location) := FrameChoiceSimilarity.
@@ -121,7 +121,6 @@ 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 (pt : location),
   (!! config)[pt] = countA_occ _ equiv_dec pt (List.map get_location (config_list config))
   := WithMultiplicity.obs_from_config_spec.
diff --git a/CaseStudies/Gathering/InR2/Weber/Align_flex_async.v b/CaseStudies/Gathering/InR2/Weber/Align_flex_async.v
new file mode 100644
index 0000000000000000000000000000000000000000..189b13e6a0efe56baa2ab9ff1676a3f7a9c0707c
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Align_flex_async.v
@@ -0,0 +1,1320 @@
+
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file implements an algorithm to ALIGN all robots on an arbitrary 
+ * axis, in the plane (R²). The algorithm assumes there are no byzantine robots,
+ * and works in a FLEXIBLE and ASYNCHRONOUS setting. 
+
+ * The algorithm is as follows : all robots go towards the 'weber point' of 
+ * the configuration. The weber point, also called geometric median, is unique 
+ * if the robots are not aligned, and has the property that moving any robot
+ * towards the weber point in a straight line doesn't change the weber point. 
+ * It thus remains at the same place throughout the whole execution.  *)
+(**************************************************************************)
+
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+(* Pactole basic definitions *)
+Require Export Pactole.Setting.
+(* Specific to R^2 topology *)
+Require Import Pactole.Spaces.RealMetricSpace.
+Require Import Pactole.Spaces.R2.
+(* Specific to gathering *)
+Require Pactole.CaseStudies.Gathering.WithMultiplicity.
+Require Import Pactole.CaseStudies.Gathering.Definitions.
+(* Specific to multiplicity *)
+Require Import Pactole.Observations.MultisetObservation.
+(* Specific to flexibility *)
+Require Import Pactole.Models.Flexible.
+(* Specific to settings with no Byzantine robots *)
+Require Import Pactole.Models.NoByzantine.
+(* Specific to definition and properties of the weber point. *)
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Weber_point.
+
+
+(* User defined *)
+Import Permutation.
+Import Datatypes.
+Import ListNotations.
+
+
+Set Implicit Arguments.
+Close Scope R_scope.
+Close Scope VectorSpace_scope.
+
+
+Section Alignment.
+Local Existing Instances dist_sum_compat R2_VS R2_ES.
+
+(* We assume the existence of a function that calculates a weber point of a collection
+ * (even when the weber point is not unique).
+ * This is a very strong assumption : such a function may not exist in closed form, 
+ * and the Weber point can only be approximated. *)
+Axiom weber_calc : list R2 -> R2.
+Axiom weber_calc_correct : forall ps, Weber ps (weber_calc ps).
+(* We also suppose this function doesn't depend on the order of the points. 
+* This is probably not necessary (we can show that it holds when the points aren't colinear) 
+* but simplifies the proof a bit. *)
+Axiom weber_calc_compat : Proper (PermutationA equiv ==> equiv) weber_calc.
+Local Existing Instance weber_calc_compat.
+  
+(* The number of robots *)
+Variables n : nat.
+Hypothesis lt_0n : 0 < n.
+
+(* There are no byzantine robots. *)
+Local Instance N : Names := Robots n 0.
+Local Instance NoByz : NoByzantine.
+Proof using . now split. Qed.
+
+Lemma list_in_length_n0 {A : Type} x (l : list A) : List.In x l -> length l <> 0.
+Proof using . intros Hin. induction l as [|y l IH] ; cbn ; auto. Qed.
+
+Lemma byz_impl_false : B -> False.
+Proof using . 
+intros b. assert (Hbyz := In_Bnames b). 
+apply list_in_length_n0 in Hbyz. 
+rewrite Bnames_length in Hbyz.
+cbn in Hbyz. intuition.
+Qed.
+
+(* Use this tactic to solve any goal
+ * provided there is a byzantine robot as a hypothesis. *)
+Ltac byz_exfalso :=
+  match goal with 
+  | b : ?B |- _ => exfalso ; apply (byz_impl_false b)
+  end.
+
+(* Since all robots are good robots, we can define a function
+ * from identifiers to good identifiers. *)
+Definition unpack_good (id : ident) : G :=
+  match id with 
+  | Good g => g 
+  | Byz _ => ltac:(byz_exfalso)
+  end.
+
+Lemma good_unpack_good id : Good (unpack_good id) == id.
+Proof using . unfold unpack_good. destruct_match ; [auto | byz_exfalso]. Qed.
+
+Lemma unpack_good_good g : unpack_good (Good g) = g.
+Proof using . reflexivity. Qed.  
+
+(* The robots are in the plane (R^2). *)
+Local Instance Loc : Location := make_Location R2.
+Local Instance LocVS : RealVectorSpace location := R2_VS.
+Local Instance LocES : EuclideanSpace location := R2_ES.
+
+(* This is what represents a robot's state.
+ * The first location is the robot's start position (where it performed its last 'compute').
+ * The second location is the robot's destination (what the robogram computed).
+ * The ratio indicates how far the robot has moved along the straight path
+ * from start to destination. *)
+(* The robogram doesn't have access to all of this information : 
+ * when we create an observation, this state gets reduced to 
+ * only the current position of the robot. *)
+(* I would have prefered to use a path instead of a (start, destination) pair,
+ * but we need an EqDec instance on [info]. *)
+Definition info := ((location * location) * ratio)%type.
+
+Local Instance info_Setoid : Setoid info := 
+  prod_Setoid (prod_Setoid location_Setoid location_Setoid) ratio_Setoid.
+Local Instance info_EqDec : EqDec info_Setoid := 
+  prod_EqDec (prod_EqDec location_EqDec location_EqDec) ratio_EqDec.
+
+Local Instance St : State info.
+simple refine {|
+  get_location := fun '(start, dest, r) => straight_path start dest r ; 
+  state_Setoid := info_Setoid ;
+  state_EqDec := info_EqDec ;
+  precondition f := sigT (fun sim : similarity location => Bijection.section sim == f) ; 
+  lift f := fun '(start, dest, r) => ((projT1 f) start, (projT1 f) dest, r) 
+|} ; autoclass.
+Proof using .
++ abstract (intros [f [sim Hsim]] [[start dest] r] ; cbn -[straight_path] ;
+            now rewrite <-Hsim, straight_path_similarity).
++ abstract (intros [[s d] r] [[s' d'] r'] [[Hs Hd] Hr] ; 
+            cbn -[equiv location] in * |- ; now rewrite Hs, Hd, Hr).
++ abstract (intros [f Hf] [g Hg] ; cbn -[equiv] ; intros Hfg [[s d] r] [[s' d'] r'] [[Hs Hd] Hr] ;
+            cbn -[equiv location] in * |- ; repeat split ; cbn -[equiv] ; auto).
+Defined.
+
+Definition get_start (i : info) := let '(s, _, _) := i in s.
+
+Local Instance get_start_compat : Proper (equiv ==> equiv) get_start.
+Proof using . intros [[? ?] ? ] [[? ?] ?] [[H _] _]. cbn -[equiv] in *. now rewrite H. Qed.
+
+Definition get_destination (i : info) := let '(_, d, _) := i in d.
+
+Local Instance get_destination_compat : Proper (equiv ==> equiv) get_destination.
+Proof using . intros [[? ?] ? ] [[? ?] ?] [[_ H] _]. cbn -[equiv] in *. now rewrite H. Qed.
+
+(* Refolding typeclass instances *)
+Ltac foldR2 :=
+  change R2 with location in * ;
+  change R2_Setoid with location_Setoid in * ;
+  change R2_EqDec with location_EqDec in * ;
+  change R2_VS with LocVS in * ;
+  change R2_ES with LocES in * ;
+  change info_Setoid with state_Setoid in * ;
+  change info_EqDec with state_EqDec in *.
+
+(* Robots choose their destination.
+ * They will move to this destination along a straight path. *)
+Local Instance RobotC : robot_choice location := 
+  { robot_choice_Setoid := location_Setoid }.
+
+(* Robots view the other robots' positions up to a similarity. *)
+Local Instance FrameC : frame_choice (similarity location) := FrameChoiceSimilarity.
+(* The demon doesn't perform any other choice for activated robots. *)
+Local Instance UpdateC : update_choice unit := NoChoice.
+(* The demon chooses how far to move inactive robots towards their destination. 
+ * The ratio chosen by the demon is ADDED to the ratio stored in the robot state
+ * (the result is clamped at 1 of course). *)
+Local Instance InactiveC : inactive_choice ratio := { inactive_choice_Setoid := ratio_Setoid }.
+
+(* In a flexible setting, delta is the minimum distance that robots
+ * are allowed to move before being reactivated. *)
+Variables delta : R.
+Hypothesis delta_g0 : (0 < delta)%R.
+
+(* This is the property that must be verified in a flexible setting. *)
+Definition flex_da_prop da := 
+  forall id (config : configuration), activate da id = true -> 
+    get_location (config id) == get_destination (config id) \/ 
+    (delta <= dist (get_start (config id)) (get_location (config id)))%R.
+
+(* We are in a flexible and semi-synchronous setting. *)
+Local Instance UpdateF : update_function location (similarity location) unit.
+simple refine {| 
+  update config g _ target _ := (get_location (config (Good g)), target, ratio_0)
+|} ; autoclass.
+Proof using .
+intros c c' Hc g g' Hg _ _ _ t t' Ht _ _ _.
+assert (H : c (Good g) == c' (Good g')) by now rewrite Hg, Hc.
+destruct (c (Good g)) as [[start dest] r].
+destruct (c' (Good g')) as [[start' dest'] r'].
+destruct H as [[Hstart Hdest] Hr]. cbn -[equiv] in Hstart, Hdest, Hr.
+repeat split ; cbn -[equiv get_location] ; auto.
+foldR2. apply get_location_compat. now repeat split ; cbn -[equiv].
+Defined.
+
+
+Local Instance InactiveF : inactive_function ratio.
+simple refine {| inactive config id r_demon := 
+  let '(start, dest, r) := config id in (start, dest, add_ratio r r_demon) 
+|} ; autoclass.
+Proof using . 
+intros c c' Hc i i' Hi rd rd' Hrd.
+assert (H : c i == c' i') by now rewrite Hi, Hc.
+destruct (c i) as [[start dest] r].
+destruct (c' i') as [[start' dest'] r'].
+destruct H as [[Hstart Hdest] Hr]. cbn -[equiv] in Hstart, Hdest, Hr.
+repeat split ; cbn -[equiv] ; auto.
+f_equiv ; auto.
+Defined.
+
+(* This is a shorthand for the list of positions of robots in a configuration. *)
+Definition pos_list (config : configuration) : list location := 
+  List.map get_location (config_list config).
+
+(* The support of a multiset, but elements are repeated 
+ * a number of times equal to their multiplicity. 
+ * This is needed to convert an observation from multiset to list format, 
+ * so that we can use functions [colinear_dec] and [weber_calc]. *)
+Definition multi_support {A} `{EqDec A} (s : multiset A) :=
+  List.flat_map (fun '(x, mx) => alls x mx) (elements s).
+
+Local Instance multi_support_compat {A} `{EqDec A} : Proper (equiv ==> PermutationA equiv) (@multi_support A _ _).
+Proof using . 
+intros s s' Hss'. unfold multi_support. apply flat_map_compat_perm with eq_pair ; autoclass ; [|]. 
++ intros [x mx] [y my] [H1 H2]. simpl in H1, H2. now rewrite H1, H2.
++ now apply elements_compat.
+Qed.
+
+(* The main algorithm : just move towards the weber point
+ * (in a straight line) until all robots are aligned. *)
+Definition gatherW_pgm obs : location := 
+  if aligned_dec (multi_support obs) 
+  (* Don't move (the robot's local frame is always centered on itself, i.e. its position is at the origin). *)
+  then origin 
+  (* Go towards the weber point. *)
+  else weber_calc (multi_support obs).
+
+Local Instance gatherW_pgm_compat : Proper (equiv ==> equiv) gatherW_pgm.
+Proof using .
+intros s1 s2 Hs. unfold gatherW_pgm.
+repeat destruct_match.
++ reflexivity.
++ rewrite Hs in H. now intuition.
++ rewrite Hs in H. now intuition.
++ apply weber_Naligned_unique with (multi_support s1) ; auto.
+  - rewrite Hs. now apply weber_calc_correct.
+  - now apply weber_calc_correct.
+Qed.
+
+Definition gatherW : robogram := {| pgm := gatherW_pgm |}.
+
+Lemma multi_support_add {A : Type} `{EqDec A} s x k : ~ In x s -> k > 0 ->
+  PermutationA equiv (multi_support (add x k s)) (alls x k ++ multi_support s).
+Proof using . 
+intros Hin Hk. unfold multi_support. 
+transitivity (flat_map (fun '(x0, mx) => alls x0 mx) ((x, k) :: elements s)).
++ apply flat_map_compat_perm with eq_pair ; autoclass ; [|].
+  - intros [a ka] [b kb] [H0 H1]. cbn in H0, H1. now rewrite H0, H1.
+  - apply elements_add_out ; auto.
++ now cbn -[elements].
+Qed.
+
+Lemma multi_support_countA {A : Type} `{eq_dec : EqDec A} s x :
+  countA_occ equiv eq_dec x (multi_support s) == s[x]. 
+Proof using .
+pattern s. apply MMultisetFacts.ind.
++ intros m m' Hm. f_equiv. 
+  - apply countA_occ_compat ; autoclass. now rewrite Hm.
+  - now rewrite Hm.
++ intros m x' n' Hin Hn IH. rewrite add_spec, multi_support_add, countA_occ_app by auto.
+  destruct_match.
+  - now rewrite <-H, countA_occ_alls_in, Nat.add_comm, IH ; autoclass.
+  - now rewrite countA_occ_alls_out, IH, Nat.add_0_l ; auto.  
++ now reflexivity.
+Qed.
+
+(* This is the main result about multi_support. *)
+(* RMK : typeclass instance inference seems to be EXTREMELY slow in this proof.
+ * Thankfully I found that the 'change' tactic is fast here. *)
+Lemma multi_support_config (config : configuration) (id : ident) : 
+  @PermutationA location equiv 
+    (@multi_support location _ _ (obs_from_config config (config id))) 
+    (pos_list config).
+Proof using . 
+pose (l := pos_list config). fold l.
+change (obs_from_config config (config id)) with (make_multiset l).
+apply PermutationA_countA_occ with R2_EqDec ; autoclass ; []. 
+intros x. rewrite multi_support_countA. now apply make_multiset_spec.
+Qed. 
+
+Corollary multi_support_map f config id : 
+  Proper (equiv ==> equiv) (projT1 f) ->
+  PermutationA (@equiv location _) 
+    (@multi_support location _ _ (obs_from_config (map_config (lift f) config) (lift f (config id))))
+    (List.map (fun x => (projT1 f) (get_location x)) (config_list config)).
+Proof using .  
+intros H. destruct f as [f Pf].
+change (lift (existT precondition f Pf) (config id)) with (map_config (lift (existT precondition f Pf)) config id).
+rewrite multi_support_config. unfold pos_list. rewrite config_list_map, map_map.
+apply eqlistA_PermutationA. (* f_equiv. *)
+apply (@map_extensionalityA_compat _ _ equiv equiv _).
+- intros [[s d] r] [[s' d'] r'] Hsdr.
+  inversion Hsdr.
+  cbn in H0,H1.
+  destruct H0.
+  rewrite <- Subset.subset_eq in H1.
+  subst.
+  cbn -[equiv straight_path]. destruct Pf as [sim Hsim]. rewrite <-Hsim.
+  apply  straight_path_similarity.
+- reflexivity.
+Qed.
+
+Lemma lift_update_swap da config1 config2 g target :
+  @equiv info _
+    (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                               (precondition_satisfied_inv da config1 g))
+          (update config2
+           g (change_frame da config1 g) target (choose_update da config2 g target)))
+    (update (map_config (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                                      (precondition_satisfied_inv da config1 g)))
+                        config2)
+            g Similarity.id
+            ((frame_choice_bijection (change_frame da config1 g ⁻¹)) target)
+            (choose_update da config2 g target)).
+Proof using .
+pose (sim := change_frame da config1 g). fold sim.
+cbn -[inverse equiv straight_path]. destruct (config2 (Good g)) as [[start dest] r].
+now rewrite straight_path_similarity.
+Qed.
+
+(* Simplify the [round] function and express it in the global frame of reference. *)
+(* All the proofs below use this simplified version. *)
+Lemma round_simplify da config : similarity_da_prop da ->
+  exists r : ident -> ratio,
+  round gatherW da config == 
+  fun id => if activate da id then
+              let target := 
+                if aligned_dec (pos_list config) 
+                then get_location (config id) 
+                else weber_calc (pos_list config) 
+              in (get_location (config id), target, ratio_0)
+            else inactive config id (r id).
+Proof using . 
+intros Hsim. eexists ?[r]. intros id. unfold round. 
+destruct_match ; [|reflexivity].
+destruct_match ; [|byz_exfalso].
+cbn -[inverse equiv lift precondition frame_choice_bijection config_list origin update get_location].
+rewrite (lift_update_swap da config _ g). 
+pose (f := existT precondition
+  (change_frame da config g)
+  (precondition_satisfied da config g)). 
+pose (f_inv := existT precondition
+  ((change_frame da config g) ⁻¹)
+  (precondition_satisfied_inv da config g)).
+pose (obs := obs_from_config (map_config (lift f) config) (lift f (config (Good g)))).
+change_LHS (update 
+  (map_config (lift f_inv) (map_config (lift f) config)) g Similarity.id
+  (frame_choice_bijection (change_frame da config g ⁻¹)
+    (gatherW_pgm obs))
+  (choose_update da
+    (map_config (lift f) config) g (gatherW_pgm obs))).
+assert (Hcancel : map_config (lift f_inv) (map_config (lift f) config) == config).
+{ intros idt. cbn -[equiv]. destruct (config idt) as [[start dest] r]. now rewrite 2 Bijection.retraction_section. }
+rewrite Hcancel.
+assert (Proper (equiv ==> equiv) (projT1 f)) as f_compat.
+{ unfold f ; cbn -[equiv]. intros x y Hxy ; now rewrite Hxy. }
+assert (Halign_loc_glob : aligned (List.map get_location (config_list config)) <-> aligned (multi_support obs)).
+{ unfold obs. rewrite multi_support_map by auto. unfold f. cbn -[config_list get_location].
+  now rewrite (R2aligned_similarity _ (change_frame da config g)), map_map. }
+destruct_match.
+(* The robots are aligned. *)
++ unfold gatherW_pgm. destruct_match ; [|intuition].
+  change (frame_choice_bijection (change_frame da config g ⁻¹) 0%VS) with (center (change_frame da config g)).
+  rewrite Hsim. cbn -[equiv inverse straight_path lift].
+  destruct (config (Good g)) as [[start dest] r].
+  reflexivity.
+(* The robots aren't aligned. *)
++ unfold gatherW_pgm. destruct_match ; [intuition|].
+  pose (sim := change_frame da config g). foldR2. fold sim.
+  assert (Hweb : weber_calc (multi_support obs) == sim (weber_calc (List.map get_location (config_list config)))).
+  {
+    unfold obs. rewrite multi_support_map by auto. unfold f. cbn -[equiv config_list get_location].
+    foldR2. fold sim. 
+    apply weber_Naligned_unique with (List.map sim (List.map get_location (config_list config))).
+    - now rewrite <-R2aligned_similarity.
+    - apply weber_similarity, weber_calc_correct.
+    - rewrite map_map. apply weber_calc_correct.
+  }
+  rewrite Hweb. cbn -[equiv config_list straight_path].
+  destruct (config (Good g)) as [[start dest] r].
+  rewrite Bijection.retraction_section. reflexivity.
+Qed.
+
+(* This is the goal (for all demons and configs).
+ * Notice that we allow robots to move once they are aligned ; 
+ * the line could even change (as long as the robots stay aligned). *)
+Definition eventually_aligned config (d : demon) (r : robogram) := 
+  Stream.eventually 
+    (Stream.forever (Stream.instant (fun c => aligned (pos_list c)))) 
+    (execute r d config).
+
+(* This is the property : all robots stay where they are. 
+ * This is what should be verified in the initial configuration. *)
+Definition config_stay (config : configuration) : Prop := 
+  forall id, let '(start, dest, _) := config id in dest == start.
+
+Local Instance config_stay_compat : Proper (equiv ==> iff) config_stay.
+Proof using . 
+intros c c' Hc. unfold config_stay. 
+assert (H : forall id, c id == c' id) by (intros id ; now specialize (Hc id)).
+split ; intros H1 id ; specialize (H1 id) ; specialize (H id) ;
+  destruct (c id) as [[s d] r] ; destruct (c' id) as [[s' d'] r'] ;
+  destruct H as [[Hs Hd] _] ; cbn -[equiv] in Hs, Hd.
++ now rewrite <-Hs, <-Hd.
++ now rewrite Hs, Hd.    
+Qed.
+
+(* This is the property : all robots stay where they are OR 
+ * go towards point p. *)
+Definition config_stay_or_go (config : configuration) p : Prop := 
+  forall id, let '(start, dest, _) := config id in dest == start \/ dest == p.
+
+Local Instance config_stay_or_go_compat : Proper (equiv ==> equiv ==> iff) config_stay_or_go.
+Proof using . 
+intros c c' Hc p p' Hp. unfold config_stay_or_go. 
+assert (H : forall id, c id == c' id) by (intros id ; now specialize (Hc id)).
+split ; intros H1 id ; specialize (H1 id) ; specialize (H id) ;
+  destruct (c id) as [[s d] r] ; destruct (c' id) as [[s' d'] r'] ;
+  destruct H as [[Hs Hd] _] ; cbn -[equiv] in Hs, Hd ; case H1 as [Hstay | Hgo].
++ left. now rewrite <-Hs, <-Hd.
++ right. now rewrite <-Hd, <-Hp.
++ left. now rewrite Hs, Hd.
++ right. now rewrite Hd, Hp.      
+Qed.
+
+Lemma config_stay_impl_config_stg config :
+  config_stay config -> forall p, config_stay_or_go config p.
+Proof using .
+unfold config_stay, config_stay_or_go. intros Hstay p i. specialize (Hstay i). 
+destruct (config i) as [[start dest] _]. now left.
+Qed.
+
+(* Typeclasses eauto := (bfs).   *)
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine x x' (c c':configuration) : 
+  InA equiv (x, x') (combine (config_list c) (config_list c')) <-> 
+    exists id, (x == c id /\ x' == c' id)%VS.
+Proof using lt_0n.
+  pose (id0 :=(Fin.Fin lt_0n)).
+  pose (g0 :=(Good id0)).
+  pose (default :=(c g0)).
+  pose (default' :=(c' g0)).
+  split.
+  + intros Hin.
+    specialize @config_list_InA_combine_same_lgth2 with (1:=Hin)(dfA:=default)(dfB:=default') as h.
+    destruct h as [i [Hn [Hx Hx']]].
+    rewrite config_list_spec in Hx,Hx'.
+    unfold default in Hx,Hx'.
+    unfold default' in Hx'.
+    rewrite map_nth in Hx .
+    rewrite map_nth in Hx' .
+    exists (nth i names g0);auto.
+  + intros [[g|b] [Hx Hx']] ; [|byz_exfalso]. 
+    rewrite Hx,Hx'.
+    rewrite config_list_spec.
+    rewrite config_list_spec.
+    rewrite combine_map.
+    apply InA_alt.
+    exists (c (Good g),c' (Good g)).
+    split;auto.
+    apply in_map_iff.
+    exists (Good g,Good g);split;auto.
+    apply in_combine_id;split;auto.
+    unfold names.
+    apply in_app_iff.
+    left.
+    apply in_map_iff.
+    exists g;split;auto.
+    apply In_Gnames.
+Qed.
+
+Lemma config_list_In_combine x x' c c' : 
+  List.In (x, x') (combine (config_list c) (config_list c')) <-> 
+  exists id, x == c id /\ x' == c' id.
+Proof using lt_0n.
+  rewrite <- In_InA_is_leibniz with (eqA:=equiv).
+  - apply config_list_InA_combine.
+  - intros x0 y. 
+    cbn.
+    split;[ intros [[[h0 h1] h2] [[h3 h4] h5]]| intro h].
+    + destruct x0 as [[[u1 u2] t] [[v1 v2] w]]; destruct y as [[[u1' u2'] t'] [[v1' v2'] w']];cbn in *.
+      rewrite <- Subset.subset_eq in h2, h5.
+      subst;auto.
+    + now subst.
+Qed.
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+(*Lemma config_list_InA_combine x x' c c' : 
+  InA equiv (x, x') (combine (config_list c) (config_list c')) <-> 
+  exists id, x == c id /\ x' == c' id.
+Proof using lt_0n.
+assert (g0 : G).
+{ change G with (fin n). apply (exist _ 0). lia. }
+split.
++ intros Hin.
+  apply (@InA_nth (info * info) equiv (c (Good g0), c' (Good g0))) in Hin.
+  destruct Hin as [i [[y y'] [Hi [Hxy Hi']]]]. 
+  rewrite combine_nth in Hi' by now repeat rewrite config_list_length. 
+  inv Hi'. inv Hxy ; cbn -[equiv config_list] in * |-.
+  setoid_rewrite H. setoid_rewrite H0.
+  assert (i < n) as Hin.
+  { 
+    eapply Nat.lt_le_trans ; [exact Hi|]. rewrite combine_length.
+    repeat rewrite config_list_length. rewrite Nat.min_id. cbn. lia. 
+  }
+  pose (g := exist (fun x => x < n) i Hin).
+  change (fin n) with G in *. exists (Good g).
+  split ; rewrite config_list_spec, map_nth ; f_equiv ; unfold names ;
+    rewrite app_nth1, map_nth by (now rewrite map_length, Gnames_length) ;
+    f_equiv ; cbn ; change G with (fin n) ; apply nth_enum.  
++ intros [[g|b] [Hx Hx']] ; [|byz_exfalso]. 
+  assert (H : (x, x') == nth (proj1_sig g) (combine (config_list c) (config_list c')) (c (Good g0), c' (Good g0))).
+  { 
+    rewrite combine_nth by now repeat rewrite config_list_length.
+    destruct g as [g Hg].
+    repeat rewrite config_list_spec, map_nth. unfold names.
+    repeat rewrite app_nth1, map_nth by now rewrite map_length, Gnames_length.
+    split ; cbn -[equiv].
+    * rewrite Hx. repeat f_equiv. change G with (fin n). erewrite nth_enum. reflexivity.
+    * rewrite Hx'. repeat f_equiv. change G with (fin n). erewrite nth_enum. reflexivity.
+  }
+  rewrite H. apply nth_InA ; autoclass. rewrite combine_length. repeat rewrite config_list_length.
+  rewrite Nat.min_id. cbn. destruct g. cbn. lia.
+Qed.
+*)
+Lemma pos_list_InA_combine x x' c c' : 
+  InA equiv (x, x') (combine (pos_list c) (pos_list c')) <-> 
+  exists id, x == get_location (c id) /\ x' == get_location (c' id).
+Proof using lt_0n.
+unfold pos_list. rewrite combine_map. rewrite (@InA_map_iff _ _ equiv equiv) ; autoclass.
++ split.
+  - intros [[y y'] [[Hx Hx'] Hin]]. cbn -[equiv get_location] in Hx, Hx'.
+    rewrite config_list_InA_combine in Hin. destruct Hin as [id [Hy Hy']].
+    exists id. rewrite <-Hy, <-Hy', Hx, Hx'. auto.
+  - intros [id [Hx Hx']].
+    exists (c id, c' id). rewrite <-Hx, Hx'. split ; [auto|].
+    rewrite config_list_InA_combine. exists id. auto.
++ intros [? ?] [? ?] [H1 H2]. cbn -[equiv] in H1, H2. split ; cbn -[equiv get_location].
+  - now rewrite H1.
+  - now rewrite H2.
+Qed. 
+
+Lemma pos_list_InA x c : 
+  InA equiv x (pos_list c) <-> exists id, x == get_location (c id).
+Proof using . 
+unfold pos_list. rewrite (@InA_map_iff _ _ equiv equiv) ; autoclass.
++ split.
+  - intros [y [Hx Hin]]. foldR2. rewrite config_list_InA in Hin.
+    destruct Hin as [id Hy]. exists id. now rewrite <-Hx, <-Hy.   
+  - intros [id Hx]. exists (c id). rewrite <-Hx. split ; [auto|].
+    foldR2. rewrite config_list_InA. exists id. auto. 
++ foldR2. apply get_location_compat.
+Qed. 
+
+Typeclasses eauto := (bfs).  
+(* A technical lemma used to prove the fact that the configuration always
+ * contracts towards the weber point. *)
+Lemma segment_progress a b r1 r2 : 
+  segment b (straight_path a b r1) (straight_path a b (add_ratio r1 r2)).
+Proof using .
+Typeclasses eauto := (dfs).  
+assert (Hr1 := ratio_bounds r1).
+assert (Hr2 := ratio_bounds r2).  
+unfold add_ratio. case (Rle_dec R1 (r1 + r2)) as [Hle | HNle].
++ rewrite straight_path_1. apply segment_start.
++ change R1 with 1%R in HNle. cbn -[mul opp RealVectorSpace.add].
+  unfold segment. exists (r2 / (1 - r1))%R. split.
+  - split ; [apply Rdiv_le_0_compat | apply Rdiv_le_1] ; lra.
+  - apply mul_reg_l with (1 - r1)%R ; [lra|].
+    repeat rewrite (mul_distr_add (1 - r1)). repeat rewrite mul_morph.
+    unfold Rdiv. rewrite <-Rmult_assoc, Rinv_r_simpl_m by lra.
+    rewrite Rmult_minus_distr_l, Rmult_1_r, <-Rmult_assoc, Rinv_r_simpl_m by lra.
+    assert (H : (a + r1 * (b - a) == r1 * b + (1 - r1) * a)%VS).
+    { unfold Rminus. rewrite mul_distr_add, <-add_morph, mul_1, 2 RealVectorSpace.add_assoc.
+      rewrite minus_morph, mul_opp. f_equiv. rewrite RealVectorSpace.add_comm. reflexivity. }
+    rewrite H, (RealVectorSpace.add_comm ((1 - r1) * a)), 2 mul_distr_add.
+    rewrite <-RealVectorSpace.add_assoc, (RealVectorSpace.add_assoc (r2 * b)). f_equiv.
+    * rewrite mul_morph, add_morph. f_equiv. lra.
+    * rewrite mul_opp, <-minus_morph, add_morph, mul_morph. f_equiv. lra. 
+Qed.
+
+(* This is the main invariant : the robots are alway headed towards a weber point. *)
+Definition invariant w config : Prop := 
+  config_stay_or_go config w /\ Weber (pos_list config) w. 
+
+Local Instance invariant_compat : Proper (equiv ==> equiv ==> iff) invariant.
+Proof using . intros w w' Hw c c' Hc. unfold invariant. now rewrite Hc, Hw. Qed. 
+
+
+(* What is remarquable is that the invariant is preserved regardless of 
+ * whether the robots are aligned or not. *)
+Lemma round_preserves_invariant config da w : similarity_da_prop da -> 
+  invariant w config -> invariant w (round gatherW da config).
+Proof using lt_0n.
+unfold invariant. intros Hsim [Hstg Hweb]. destruct (round_simplify config Hsim) as [r Hround].
+split.
++ case (aligned_dec (pos_list config)) as [Halign | HNalign] ; cbn zeta in Hround.
+  (* The robots are aligned. *)  
+  - rewrite Hround. intros i. destruct_match ; [now left |].
+    specialize (Hstg i). cbn -[equiv]. now destruct (config i) as [[s d] _].
+  (* The robots aren't aligned. *)
+  - rewrite Hround. intros i. destruct_match.
+    * right. apply weber_Naligned_unique with (pos_list config) ; auto.
+      apply weber_calc_correct.
+    * specialize (Hstg i). cbn -[equiv]. now destruct (config i) as [[s d] _].   
++ revert Hweb. apply weber_contract. unfold contract. 
+  rewrite Forall2_Forall, Forall_forall by (now unfold pos_list ; repeat rewrite map_length, config_list_length).
+  intros [x x'] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+  rewrite pos_list_InA_combine in Hin. destruct Hin as [id [Hx Hx']].
+  rewrite Hx, Hx', (Hround id).
+  destruct_match.
+  (* Activated robots don't move. *)
+  * cbn zeta. cbn -[config_list straight_path]. rewrite straight_path_0. 
+    apply segment_end.
+  (* Inactive robots move along a straight line towards w. *)
+  * cbn -[straight_path]. 
+    specialize (Hstg id). destruct (config id) as [[s d] ri].
+    case Hstg as [Hstay | Hgo].
+    --rewrite Hstay, 2 straight_path_same. apply segment_end.
+    --rewrite Hgo. apply segment_progress.
+Qed.
+
+(* If the robots aren't aligned, then the point refered to in the invariant
+ * is necessarily the unique weber point. *)
+Lemma invariant_weber_calc config w :
+  invariant w config -> ~aligned (pos_list config) -> 
+  w == weber_calc (pos_list config).
+Proof using .
+intros [Hweb Hstg] HNalign. apply weber_Naligned_unique with (pos_list config) ; auto.
+apply weber_calc_correct.
+Qed.
+
+(* This part is rather complicated in ASYNC, since robots continue moving
+ * even after being aligned. However we show that they stay aligned : 
+ * this mainly comes from the fact that if robots are aligned,
+ * then any weber point is also on the same line. *)
+Lemma round_preserves_aligned config da w : similarity_da_prop da -> 
+  invariant w config -> 
+  aligned (pos_list config) -> 
+  aligned (pos_list (round gatherW da config)).
+Proof using lt_0n. 
+intros Hsim [Hstg Hweb] [L Halign]. 
+assert (H : aligned (w :: pos_list config)).
+{ exists L. rewrite aligned_on_cons_iff. split ; [apply weber_aligned with (pos_list config) | assumption].
+  + unfold pos_list. rewrite <-length_zero_iff_nil, map_length, config_list_length. cbn. lia.
+  + assumption.
+  + assumption. }
+rewrite aligned_spec in H. destruct H as [v H].
+apply aligned_tail with w. rewrite aligned_spec. 
+exists v. intros p. unfold pos_list. 
+rewrite (@InA_map_iff _ _ equiv equiv) ; autoclass ; 
+  [|foldR2 ; apply get_location_compat].
+intros [x [Hx Hin]]. foldR2. rewrite config_list_InA in Hin.
+destruct Hin as [id Hid]. revert Hid.
+specialize (H (get_location (config id))). feed H.
+{ rewrite pos_list_InA. now exists id. }
+destruct H as [t H].
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id).
+destruct_match.
+(* The robot is activated. *)
++ cbn zeta. case ifP_sumbool ; [intros _ | intros HNalign].
+  - intros Hid. exists t. rewrite <-Hx, <-H, Hid.
+    cbn -[equiv straight_path]. now rewrite straight_path_0.
+  - exfalso. apply HNalign. now exists L.
+(* The robot isn't activated. *)
++ cbn -[equiv straight_path mul RealVectorSpace.add].
+  specialize (Hstg id). destruct (config id) as [[s d] rx].
+  rewrite <-Hx. clear Hx Hround.
+  case Hstg as [Hstay | Hgo].
+  - rewrite Hstay in *. intros Hid. exists t. rewrite Hid, <-H.
+    cbn -[equiv straight_path]. now repeat rewrite straight_path_same.
+  - rewrite Hgo in *. intros Hid. setoid_rewrite Hid.
+    unfold add_ratio. case (Rle_dec R1 (rx + r id)) as [Rle | RNle].
+    * cbn -[equiv straight_path mul RealVectorSpace.add]. 
+      setoid_rewrite straight_path_1. exists 0%R. now rewrite mul_0, add_origin_r.
+    * cbn -[equiv mul opp RealVectorSpace.add] in H |- *. change R1 with 1%R in *.
+      assert (Halg : exists t', (w - s == t' * v)%VS).
+      {
+        exists (t / (rx - 1))%R. unfold Rdiv. 
+        rewrite Rmult_comm, <-mul_morph. 
+        apply R2div_reg_l ; [generalize (ratio_bounds (r id)) ; lra|].
+        unfold Rminus. rewrite <-add_morph, minus_morph, mul_1.
+        rewrite opp_distr_add, opp_opp, (RealVectorSpace.add_comm _ s).
+        rewrite RealVectorSpace.add_assoc, (RealVectorSpace.add_comm _ s).
+        rewrite H, (RealVectorSpace.add_comm w), <-RealVectorSpace.add_assoc, add_opp.
+        now rewrite add_origin_r. 
+      }
+      destruct Halg as [t' Halg]. 
+      exists (-t' + (rx + r id) * t')%R.
+      rewrite Halg, <-(add_morph (-t')%R), RealVectorSpace.add_assoc, mul_morph.
+      f_equiv. rewrite minus_morph, <-Halg, opp_distr_add, opp_opp.
+      rewrite RealVectorSpace.add_assoc, add_opp, add_origin_l.
+      reflexivity.
+Qed.
+
+(* If the robots are aligned at any point, they stay aligned forever. *)
+Corollary aligned_over config (d : demon) w :
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  invariant w config -> 
+  aligned (pos_list config) -> 
+  Stream.forever (Stream.instant (fun c => aligned (pos_list c))) (execute gatherW d config).
+Proof using lt_0n.
+revert config d.
+cofix Hind. intros config d Hsim Hinv Halign. constructor.
++ cbn -[pos_list]. exact Halign.
++ cbn -[pos_list]. apply Hind.
+  - apply Hsim.
+  - apply round_preserves_invariant ; [apply Hsim | exact Hinv].
+  - apply round_preserves_aligned with w ; [apply Hsim | exact Hinv | exact Halign].
+Qed.
+
+(* We say that a robot is looping when its start and destination points are equal. *)
+Definition is_looping (robot : info) : bool := 
+  if get_start robot =?= get_destination robot then true else false.
+
+Local Instance is_looping_compat : Proper (equiv ==> eq) is_looping. 
+Proof using . intros [[? ?] ?] [[? ?] ?] [[H1 H2] _]. cbn -[equiv] in *.
+              cbn in H1,H2.
+              now rewrite H1, H2. Qed.     
+
+Lemma is_looping_ratio start dest r1 r2 : 
+  is_looping (start, dest, r1) = is_looping (start, dest, r2).
+Proof using . now unfold is_looping. Qed. 
+
+(* Boolean function to test whether a robot is on a point. *)
+Definition is_on x (robot : info) : bool :=
+  if get_location robot =?= x then true else false.
+
+Local Instance is_on_compat : Proper (equiv ==> equiv ==> eq) is_on.
+Proof using . 
+intros ? ? H1 ? ? H2. unfold is_on. 
+repeat destruct_match ; rewrite H1, H2 in * ; intuition. 
+Qed.
+
+Definition BtoR : bool -> R := fun b => if b then 1%R else 0%R.
+
+(* This measure counts how many robots are [not on the weber point] and [looping]. *)
+Definition measure_loop_nonweb config : R :=
+  let w := weber_calc (pos_list config) in
+  list_sum (List.map 
+    (fun r => BtoR (is_looping r && negb (is_on w r))) 
+    (config_list config)).
+
+(* This measure counts how many robots are not [looping on the weber point]. *)
+Definition measure_loop_web config : R :=
+  let w := weber_calc (pos_list config) in
+  list_sum (List.map 
+    (fun r => BtoR (negb (is_looping r && is_on w r)))
+    (config_list config)).
+
+(* This measure counts the total distance from the weber point to 
+ * the last update position of each robot. 
+ * RMK : this is NOT the distance from the weber point to the 
+ * current position of each robot. *)
+Definition measure_dist config : R := 
+  dist_sum 
+    (List.map get_start (config_list config)) 
+    (weber_calc (pos_list config)).
+
+(* The resulting measure is well-founded, and decreases whenever a robot is activated. *)
+Definition measure config : R := 
+  measure_loop_nonweb config + measure_loop_web config + measure_dist config.
+
+Local Instance measure_compat : Proper (equiv ==> equiv) measure.
+Proof using . 
+intros c c' Hc. unfold measure.
+f_equiv ; [f_equiv|].
++ unfold measure_loop_nonweb. apply list_sum_compat, eqlistA_PermutationA. 
+  f_equiv ; [| now rewrite Hc].
+  intros i i' Hi. now rewrite Hi, Hc.
++ unfold measure_loop_web. apply list_sum_compat, eqlistA_PermutationA. 
+  f_equiv ; [| now rewrite Hc].
+  intros i i' Hi. now rewrite Hi, Hc.
++ unfold measure_dist. f_equiv ; [|now rewrite Hc]. apply eqlistA_PermutationA. 
+  apply map_eqlistA_compat with equiv ; autoclass.
+  now rewrite Hc.
+Qed.
+
+(* The measure is trivially non-negative. *)
+Lemma measure_nonneg config : (0 <= measure config)%R.
+Proof using .
+unfold measure. repeat apply Rplus_le_le_0_compat.
++ apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros x _. unfold BtoR. destruct_match ; lra.
++ apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros x _. unfold BtoR. destruct_match ; lra.
++ unfold measure_dist. apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros x _. apply dist_nonneg.   
+Qed.
+
+Section MeasureDecreaseLemmas.
+Variables (config : configuration) (da : demonic_action) (w : location).
+Hypothesis (Hsim : similarity_da_prop da).
+Hypothesis (Hinv : invariant w config).
+Hypothesis (HRNalign : ~aligned (pos_list (round gatherW da config))).
+
+Lemma HNalign : ~aligned (pos_list config).
+Proof using lt_0n HRNalign Hinv Hsim. 
+intro Halign. now apply HRNalign, round_preserves_aligned with w. 
+Qed.
+
+Lemma HRw : w == weber_calc (pos_list (round gatherW da config)).
+Proof using HRNalign Hinv Hsim lt_0n. 
+now apply invariant_weber_calc ; [apply round_preserves_invariant|]. 
+Qed.
+
+Lemma Hw : w == weber_calc (pos_list config).
+Proof using lt_0n HRNalign Hinv Hsim. 
+apply invariant_weber_calc ; auto. apply HNalign. 
+Qed. 
+
+Lemma BtoR_le b1 b2 : (b1 = true -> b2 = true) <-> (BtoR b1 <= BtoR b2)%R.
+Proof using HRNalign. unfold BtoR. repeat destruct_match ; solve [lra | intuition]. Qed.
+
+Lemma BtoR_le_1 b1 b2 : (b1 = false /\ b2 = true) <-> (BtoR b1 <= BtoR b2 - 1)%R.
+Proof using HRNalign. unfold BtoR. repeat destruct_match ; intuition ; lra. Qed.
+
+Lemma weber_dist_decreases id :
+  (dist w (get_start (round gatherW da config id)) <= dist w (get_start (config id)))%R.
+Proof using lt_0n HRNalign Hinv Hsim. 
+destruct Hinv as [Hstg Hweb].
+destruct (round_simplify config Hsim) as [r Hround].
+rewrite (Hround id). destruct_match. 
++ destruct_match ; [assert (HNalign := HNalign) ; intuition |].
+  cbn -[dist straight_path]. specialize (Hstg id).
+  destruct (config id) as [[s d] ri]. cbn [get_start].
+  case Hstg as [Hstay | Hgo].
+  - rewrite Hstay, straight_path_same. reflexivity.
+  - rewrite Hgo. rewrite straight_path_dist_end, dist_sym. 
+    rewrite <-Rmult_1_l. apply Rmult_le_compat_r ; [apply dist_nonneg|].
+    generalize (ratio_bounds ri). lra.
++ cbn -[dist straight_path]. destruct (config id) as [[s d] ri]. cbn [get_start]. reflexivity.
+Qed.
+
+Lemma weber_dist_decreases_strong id :
+  activate da id = true -> 
+  get_destination (config id) == w ->
+  get_location (config id) =/= w ->
+  flex_da_prop da ->
+  (dist w (get_start (round gatherW da config id)) <= dist w (get_start (config id)) - delta)%R.
+Proof using lt_0n HRNalign Hinv Hsim.
+intros Hact Hdest Hloc Hflex. specialize (Hflex id config Hact).
+destruct Hinv as [Hstg Hweb].
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id). 
+foldR2. change FrameChoiceSimilarity with FrameC in *. destruct_match_eq H ; [|intuition]. 
+destruct_match ; [assert (HNalign := HNalign) ; intuition |].
+cbn -[dist straight_path]. specialize (Hstg id).
+destruct (config id) as [[s d] ri]. cbn -[dist equiv straight_path] in Hdest, Hloc, Hflex |- *.
+case Hstg as [Hstay | Hgo].
++ exfalso. apply Hloc. now rewrite <-Hdest, Hstay, straight_path_same.
++ rewrite Hgo in *. case Hflex as [Hreached | Hdelta] ; [intuition|].
+  transitivity (dist w s - dist s (straight_path s w ri))%R.
+  - rewrite straight_path_dist_end, straight_path_dist_start, Rmult_minus_distr_r, Rmult_1_l. 
+    repeat rewrite (dist_sym w s). reflexivity.
+  - unfold Rminus. apply Rplus_le_compat_l, Ropp_le_contravar. exact Hdelta.
+Qed.
+  
+Lemma measure_dist_decreases : 
+  (measure_dist (round gatherW da config) <= measure_dist config)%R.
+Proof using lt_0n HRNalign Hinv Hsim lt_0n. 
+apply list_sum_le. rewrite Forall2_Forall by (now repeat rewrite map_length ; repeat rewrite config_list_length).
+rewrite combine_map, Forall_map, Forall_forall.
+intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+rewrite combine_map, (@InA_map_iff _ _ equiv equiv) in Hin ; autoclass.
++ destruct Hin as [[y' y] [[Hy' Hy] Hin]]. cbn -[equiv] in Hy, Hy'.
+  rewrite config_list_InA_combine in Hin. destruct Hin as [id [Hx' Hx]].
+  rewrite <-Hw, <-HRw, <-Hy, <-Hy', Hx, Hx'. apply weber_dist_decreases ; auto.
++ intros [? ?] [? ?] [H1 H2]. now rewrite H1, H2.
+Qed.
+
+Lemma contra_bool b1 b2 : (b2 = true -> b1 = true) -> (negb b1 = true -> negb b2 = true).
+Proof using . case b1 ; case b2 ; intuition. Qed.
+
+Lemma loop_web_decreases id : 
+  (BtoR (negb (is_looping (round gatherW da config id) && is_on w (round gatherW da config id))) <=  
+  BtoR (negb (is_looping (config id) && is_on w (config id))))%R.
+Proof using lt_0n HRNalign Hinv Hsim. 
+destruct Hinv as [Hstg Hweb]. specialize (Hstg id).
+apply BtoR_le. apply contra_bool. rewrite 2 andb_true_iff. intros [Hloop Hon]. revert Hloop Hon.
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id).
+destruct_match.
++ destruct_match ; [assert (HNalign := HNalign) ; intuition |].
+  rewrite <-Hw. unfold is_looping, is_on. cbn -[straight_path equiv_dec].
+  destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+  repeat (case ifP_sumbool ; try now intuition).
+  rewrite straight_path_0. intuition.
++ unfold is_looping, is_on. cbn -[straight_path equiv_dec]. 
+  destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+  repeat (case ifP_sumbool ; try now intuition).
+  intros H1 H2 Hsd. exfalso. revert H1 H2. rewrite Hsd, 2 straight_path_same.
+  now intuition.
+Qed.
+
+Lemma loop_web_decreases_strong id :
+  activate da id = true -> 
+  get_location (config id) == w -> 
+  get_start (config id) =/= w ->
+  (BtoR (negb (is_looping (round gatherW da config id) && is_on w (round gatherW da config id))) <=  
+  BtoR (negb (is_looping (config id) && is_on w (config id))) - 1)%R.
+Proof using lt_0n HRNalign Hinv Hsim.
+intros Hact Hloc Hstart. apply BtoR_le_1. 
+rewrite negb_true_iff, negb_false_iff, andb_true_iff, andb_false_iff.
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id).
+destruct Hinv as [Hstg Hweb]. specialize (Hstg id).
+foldR2. change FrameChoiceSimilarity with FrameC.
+revert Hact. case ifP_bool ; [intros Hact _ | discriminate].
+destruct_match ; [assert (HNalign := HNalign) ; intuition |]. rewrite <-Hw. cbn zeta.
+revert Hloc Hstart. unfold is_looping, is_on. cbn -[straight_path equiv equiv_dec].
+destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+case Hstg as [Hstay | Hgo].
++ rewrite Hstay. 
+  repeat (case ifP_sumbool ; try now intuition) ;
+    repeat rewrite straight_path_same ; intuition.
++ rewrite Hgo. 
+  repeat (case ifP_sumbool ; try now intuition) ; rewrite straight_path_0. 
+  intuition.
+Qed.
+
+Lemma measure_loop_web_decreases : 
+  (measure_loop_web (round gatherW da config) <= measure_loop_web config)%R.
+Proof using lt_0n HRNalign Hinv Hsim lt_0n. 
+apply list_sum_le. rewrite Forall2_Forall by (now repeat rewrite map_length ; repeat rewrite config_list_length).
+rewrite combine_map, Forall_map, Forall_forall.
+intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+rewrite config_list_InA_combine in Hin. destruct Hin as [id [Hx' Hx]].
+rewrite <-Hw, <-HRw, Hx, Hx'. apply loop_web_decreases ; auto.
+Qed.
+
+Lemma loop_nonweb_decreases id :
+  (BtoR (is_looping (round gatherW da config id) && negb (is_on w (round gatherW da config id))) <=
+  BtoR (is_looping (config id) && negb (is_on w (config id))))%R.
+Proof using lt_0n HRNalign Hinv Hsim.
+destruct Hinv as [Hstg Hweb]. specialize (Hstg id).
+apply BtoR_le. rewrite 2 andb_true_iff, 2 negb_true_iff.
+intros [Hloop Hon]. revert Hloop Hon.
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id).
+destruct_match.
++ destruct_match ; [assert (HNalign := HNalign) ; intuition |].
+  rewrite <-Hw. unfold is_looping, is_on. cbn -[straight_path equiv_dec].
+  destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+  repeat (case ifP_sumbool ; try now intuition) ; rewrite straight_path_0 ; intuition.
++ unfold is_looping, is_on. cbn -[straight_path equiv_dec].
+  destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+  repeat (case ifP_sumbool ; try now intuition).
+  intros H1 H2 ->. exfalso. rewrite straight_path_same in H1, H2. intuition.
+Qed.
+
+Lemma loop_nonweb_decreases_strong id :
+  activate da id = true -> 
+  get_destination (config id) =/= w -> 
+  (BtoR (is_looping (round gatherW da config id) && negb (is_on w (round gatherW da config id))) <=
+  BtoR (is_looping (config id) && negb (is_on w (config id))) - 1)%R.
+Proof using lt_0n HRNalign Hinv Hsim. 
+intros Hact Hdest. apply BtoR_le_1. 
+rewrite andb_true_iff, andb_false_iff, negb_false_iff, negb_true_iff.
+destruct (round_simplify config Hsim) as [r Hround]. rewrite (Hround id).
+destruct Hinv as [Hstg Hweb]. specialize (Hstg id).
+foldR2. change FrameChoiceSimilarity with FrameC. 
+revert Hact. case ifP_bool ; [|discriminate].
+destruct_match ; [assert (HNalign := HNalign) ; intuition |]. rewrite <-Hw. cbn zeta.
+revert Hdest. unfold is_looping, is_on. cbn -[straight_path equiv equiv_dec].
+destruct (config id) as [[s d] ri]. cbn [get_start get_destination].
+repeat (case ifP_sumbool ; try now intuition) ;
+  intros -> ; rewrite straight_path_same ; intuition.
+Qed.
+
+Lemma measure_loop_nonweb_decreases : 
+  (measure_loop_nonweb (round gatherW da config) <= measure_loop_nonweb config)%R.
+Proof using lt_0n HRNalign Hinv Hsim lt_0n. 
+apply list_sum_le. rewrite Forall2_Forall by (now repeat rewrite map_length ; repeat rewrite config_list_length).
+rewrite combine_map, Forall_map, Forall_forall.
+intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+rewrite config_list_InA_combine in Hin. destruct Hin as [id [Hx' Hx]].
+rewrite <-Hw, <-HRw, Hx, Hx'. apply loop_nonweb_decreases ; auto.
+Qed.
+
+End MeasureDecreaseLemmas.
+    
+Lemma round_decrease_measure config da w :
+  similarity_da_prop da ->
+  invariant w config ->
+  ~aligned (pos_list (round gatherW da config)) ->  
+    (measure (round gatherW da config) <= measure config)%R. 
+Proof using lt_0n.
+intros Hsim Hinv HRNalign.
+unfold measure. repeat apply Rplus_le_compat.
++ apply measure_loop_nonweb_decreases with w ; auto.
++ apply measure_loop_web_decreases with w ; auto.
++ apply measure_dist_decreases with w ; auto.
+Qed. 
+ 
+Lemma Rplus_le_compat3_1 x y z x' y' z' eps : 
+  (x <= x' - eps)%R -> (y <= y')%R -> (z <= z')%R -> (x + y + z <= x' + y' + z' - eps)%R.
+Proof using . lra. Qed.
+
+Lemma Rplus_le_compat3_2 x y z x' y' z' eps : 
+  (x <= x')%R -> (y <= y' - eps)%R -> (z <= z')%R -> (x + y + z <= x' + y' + z' - eps)%R.
+Proof using . lra. Qed.
+
+Lemma Rplus_le_compat3_3 x y z x' y' z' eps : 
+  (x <= x')%R -> (y <= y')%R -> (z <= z' - eps)%R -> (x + y + z <= x' + y' + z' - eps)%R.
+Proof using . lra. Qed. 
+
+Typeclasses eauto := (dfs) 5.
+(* If a robot that is not [looping on the weber point] is activated, 
+ * the measure strictly decreases. *)
+Lemma round_decreases_measure_strong config da w : 
+  similarity_da_prop da -> 
+  flex_da_prop da ->
+  invariant w config -> 
+  ~aligned (pos_list (round gatherW da config)) ->
+  (exists id, activate da id = true /\ is_looping (config id) && is_on w (config id) = false) -> 
+    (measure (round gatherW da config) <= measure config - Rmin delta 1)%R.
+Proof using lt_0n. 
+intros Hsim Hflex Hinv HRNalign [id [Hact Hnlw]].
+assert (H := Hinv). destruct H as [Hstg Hweb]. specialize (Hstg id). 
+assert (HNalign : ~aligned (pos_list config)).
+{ intro Halign. now apply HRNalign, round_preserves_aligned with w. }
+assert (HRw : w == weber_calc (pos_list (round gatherW da config))).
+{ now apply invariant_weber_calc ; [apply round_preserves_invariant|]. }
+assert (Hw : w == weber_calc (pos_list config)).
+{ apply invariant_weber_calc ; auto. }
+rewrite andb_false_iff in Hnlw. 
+case (Sumbool.sumbool_of_bool (is_looping (config id))) as [Hloop | HNloop]. 
++ destruct Hnlw as [HNloop | HNon] ; [now rewrite Hloop in HNloop |].
+  transitivity (measure config - 1)%R ; [|generalize (Rmin_r delta 1%R) ; lra].
+  apply Rplus_le_compat3_1 ; [| eapply measure_loop_web_decreases ; eauto | eapply measure_dist_decreases ; eauto].
+  apply list_sum_le_eps ; rewrite <-Hw, <-HRw.
+  * rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall 
+      by now repeat rewrite map_length, config_list_length.
+    intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+    rewrite config_list_InA_combine in Hin. destruct Hin as [id' [-> ->]].
+    now apply loop_nonweb_decreases.
+  * pose (f := fun r0 => BtoR (is_looping r0 && negb (is_on w r0))). fold f. 
+    rewrite Exists_exists, combine_map.
+    eexists (?[x], ?[x']).
+    rewrite <-(@In_InA_is_leibniz _ equiv), (@InA_map_iff _ _ equiv equiv) ; autoclass.
+    ++split. 
+      --exists (round gatherW da config id, config id). split ; [reflexivity|].
+        rewrite config_list_InA_combine. now exists id.
+      --unfold f. apply loop_nonweb_decreases_strong ; auto.
+        destruct (config id) as [[s d] ri]. revert HNon Hloop. unfold is_on, is_looping.
+        case ifP_sumbool ; case ifP_sumbool ; try discriminate.
+        cbn -[equiv straight_path]. intros ->. now rewrite straight_path_same.
+    ++intros [? ?] [? ?] [H1 H2]. unfold f.
+      cbn -[equiv] in H1,H2.
+      foldR2.
+      rewrite H1, H2. reflexivity.
+    ++intros [? ?] [? ?]. split ; auto.
++ unfold is_looping in HNloop. destruct (config id) as [[s d] r] eqn:Econfig. simpl in HNloop.
+  case (s =?= d) as [Hsd | Hsd] ; [exfalso ; revert HNloop ; destruct_match ; intuition |].
+  case Hstg as [Hstay | Hgo] ; [intuition|]. clear HNloop.
+  case (get_location (config id) =?= w) as [HReached | HNreached].
+  - transitivity (measure config - 1)%R ; [|generalize (Rmin_r delta 1%R) ; lra].
+    apply Rplus_le_compat3_2 ; [eapply measure_loop_nonweb_decreases ; eauto | | eapply measure_dist_decreases ; eauto].
+    apply list_sum_le_eps ; rewrite <-Hw, <-HRw.
+    * rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall 
+        by now repeat rewrite map_length, config_list_length.
+      intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+      rewrite config_list_InA_combine in Hin. destruct Hin as [id' [-> ->]].
+      now apply loop_web_decreases.
+    * pose (f := fun r0 => BtoR (negb (is_looping r0 && is_on w r0))). fold f. 
+      rewrite Exists_exists, combine_map.
+      eexists (?[x], ?[x']).
+      rewrite <-(@In_InA_is_leibniz _ equiv), (@InA_map_iff _ _ equiv equiv) ; autoclass.
+      ++split. 
+        --exists (round gatherW da config id, config id). split ; [reflexivity|].
+          rewrite config_list_InA_combine. now exists id.
+        --unfold f. apply loop_web_decreases_strong ; auto. now rewrite Econfig, <-Hgo.
+      ++intros [? ?] [? ?] [H1 H2]. unfold f. rewrite H1, H2. reflexivity.
+      ++intros [? ?] [? ?] ; split ; auto.
+  - transitivity (measure config - delta)%R ; [|generalize (Rmin_l delta 1%R) ; lra].
+    apply Rplus_le_compat3_3 ; [eapply measure_loop_nonweb_decreases ; eauto | eapply measure_loop_web_decreases ; eauto |].
+    apply list_sum_le_eps ; rewrite <-Hw, <-HRw.
+    * rewrite Forall2_Forall, 2 map_map, combine_map, Forall_map, Forall_forall 
+        by now repeat rewrite map_map, map_length, config_list_length.
+      intros [x' x] Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+      rewrite config_list_InA_combine in Hin. destruct Hin as [id' [-> ->]].
+      now apply weber_dist_decreases.
+    * rewrite 2 map_map, Exists_exists, combine_map.
+      eexists (?[x], ?[x']).  
+      rewrite <-(@In_InA_is_leibniz _ equiv), (@InA_map_iff _ _ equiv equiv) ; autoclass.
+      ++split. 
+        --exists (round gatherW da config id, config id). split ; [reflexivity|].
+          rewrite config_list_InA_combine. now exists id.
+        --apply weber_dist_decreases_strong ; auto. now rewrite Econfig, <-Hgo.
+      ++intros [? ?] [? ?] [H1 H2]. rewrite H1, H2. reflexivity.
+      ++intros [? ?] [? ?] ; split ; auto.  
+Qed.
+
+
+(* This inductive proposition counts how many turns are left before
+ * a robot that isn't [looping on w] is activated.
+ * This is analoguous to FirstMove in the SSYNC case. *)
+Inductive FirstActivNLW w : demon -> configuration -> Prop :=
+  | FirstActivNLW_Now : forall d config id, 
+    activate (Stream.hd d) id = true -> is_looping (config id) && is_on w (config id) = false -> 
+      FirstActivNLW w d config
+  | FirstActivNLW_Later : forall d config,
+    FirstActivNLW w (Stream.tl d) (round gatherW (Stream.hd d) config) -> 
+      FirstActivNLW w d config. 
+
+Typeclasses eauto := (bfs).
+
+Lemma gathered_aligned ps x : 
+  Forall (fun y => y == x) ps -> aligned ps.
+Proof using. 
+Typeclasses eauto := (dfs) 7.
+rewrite Forall_forall. intros Hgathered.
+exists (Build_line x 0%VS). intros y Hin. unfold line_contains.
+rewrite line_proj_spec, line_proj_inv_spec. cbn -[equiv RealVectorSpace.add opp mul inner_product].
+unfold Rdiv. rewrite inner_product_origin_r, Rmult_0_l, mul_0, add_origin_r. 
+symmetry. apply Hgathered. rewrite <-In_InA_is_leibniz ; autoclass. now cbn.
+Qed.
+    
+(* If the robots aren't aligned yet, there exists a robot that isn't [looping on w]. *)
+Lemma exists_non_webloop config w : 
+  ~aligned (pos_list config) ->
+  invariant w config -> 
+  exists id, is_looping (config id) && is_on w (config id) = false.
+Proof using lt_0n. 
+intros HNalign Hinv.
+assert (H := Forall_dec (fun id => is_looping (config id) && is_on w (config id) = true)).
+feed H ; [intros id ; case (is_looping (config id) && is_on w (config id)) ; intuition |].
+case (H names) as [HT | HF] ; clear H.
++ exfalso. apply HNalign. apply gathered_aligned with w. 
+  rewrite Forall_forall in HT |- *. intros x.
+  rewrite <-(In_InA_is_leibniz equiv) by now intros [? ?] [? ?].
+  rewrite pos_list_InA. intros [id ->].
+  specialize (HT id). feed HT ; [apply In_names |].
+  rewrite andb_true_iff in HT. destruct HT as [Hloop Hon].
+  revert Hon. unfold is_on. case ifP_sumbool ; try discriminate. auto. 
++ rewrite <-Exists_Forall_neg, Exists_exists in HF.
+  - destruct HF as [id [_ Hnwl]]. apply not_true_is_false in Hnwl. 
+    now exists id.
+  - intros id. tauto.
+Qed.
+
+(* Fairness entails that if the robots aren't aligned yet, 
+ * then a robot that isn't [looping on w] will eventually be activated. *)
+Lemma non_webloop_will_activate config d w :
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  Fair d ->
+  ~aligned (pos_list config) ->    
+  invariant w config -> 
+  FirstActivNLW w d config.
+Proof using lt_0n.
+intros Hsim Hfair HNalign Hinv.
+destruct (exists_non_webloop HNalign Hinv) as [id Hnwl].
+destruct Hfair as [Hlocallyfair Hfair]. specialize (Hlocallyfair id).
+clear HNalign. generalize dependent config.
+induction Hlocallyfair as [d Hact | d Hlater IH].
++ intros config Hinv Hnwl.
+  apply FirstActivNLW_Now with id ; auto.
++ intros config Hinv Hnwl.
+  case (Sumbool.sumbool_of_bool (activate (Stream.hd d) id)) as [Hact | HNact] ;
+    [apply FirstActivNLW_Now with id ; auto |].
+  apply FirstActivNLW_Later, IH.
+  - apply Hsim.
+  - apply Hfair.
+  - apply round_preserves_invariant ; [apply Hsim | apply Hinv].
+  - destruct Hsim as [Hsim_hd Hsim_tl]. 
+    destruct (round_simplify config Hsim_hd) as [r Hround].
+    rewrite (Hround id). destruct_match_eq Hact ; 
+      [foldR2 ; rewrite Hact in HNact ; discriminate |].
+    cbn. destruct (config id) as [[start dest] ri].
+    rewrite andb_false_iff in Hnwl |- *.
+    rewrite (is_looping_ratio _ _ _ ri).
+    case Hnwl as [HNloop | HNon] ; [now left |].
+    case (start =?= dest) as [Hloop | HNloop] ; [right | left].
+    * revert HNon. cbn in Hloop. rewrite Hloop. unfold is_on. 
+      cbn -[straight_path]. repeat (case ifP_sumbool ; try now intuition). 
+      rewrite 2 straight_path_same. intuition.
+    * unfold is_looping. destruct_match ; intuition.
+Qed. 
+
+(* This is the well founded relation we will perform induction on. *)
+Definition lt_config eps c c' := 
+  (0 <= measure c <= measure c' - eps)%R. 
+
+Typeclasses eauto := (dfs) 5.
+Local Instance lt_config_compat : Proper (equiv ==> equiv ==> equiv ==> iff) lt_config.
+Proof using . intros e e' He c1 c1' Hc1 c2 c2' Hc2. unfold lt_config. now rewrite He, Hc1, Hc2. Qed.
+
+(* We proove this using the well-foundedness of lt on nat. *)
+Lemma lt_config_wf eps : (eps > 0)%R -> well_founded (lt_config eps).
+Proof using . 
+intros Heps. unfold well_founded. intros c.
+pose (f := fun x : R => Z.to_nat (up (x / eps))).
+remember (f (measure c)) as k. generalize dependent c. 
+pattern k. apply (well_founded_ind lt_wf). clear k.
+intros k IH c Hk. apply Acc_intro. intros c' Hc'. apply IH with (f (measure c')) ; auto.
+rewrite Hk ; unfold f ; unfold lt_config in Hc'.
+rewrite <-Z2Nat.inj_lt.
++ apply Zup_lt. unfold Rdiv. rewrite <-(Rinv_r eps) by lra. 
+  rewrite <-Rmult_minus_distr_r. apply Rmult_le_compat_r ; intuition.
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition. 
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition.
+  transitivity eps ; intuition. 
+  apply (Rplus_le_reg_r (- eps)). rewrite Rplus_opp_r. etransitivity ; eauto.
+Qed.
+
+
+(* Notice that [invariant w config] is in the assumptions : 
+ * see the next theorem for the final assumptions. 
+ * The proof is essentially a well-founded induction on [measure config].
+ * Fairness ensures that the measure must decrease at some point,
+ * which leads to a second induction on [FirstMoveNLW w d config]. *)
+Lemma weber_correct_aux w config d :
+  Fair d -> 
+  invariant w config ->
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  Stream.forever (Stream.instant flex_da_prop) d ->
+  eventually_aligned config d gatherW.
+Proof using lt_0n delta_g0.
+assert (Hdelta1 : (Rmin delta 1 > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+revert d.
+induction config as [config IH] using (well_founded_ind (lt_config_wf Hdelta1)).
+intros d Hfair Hinv Hsim Hflex.
+case (aligned_dec (pos_list (round gatherW (Stream.hd d) config))) as [HRalign | HRNalign].
+{ apply Stream.Later, Stream.Now, aligned_over with w ; auto.
+  + apply Hsim.
+  + now apply round_preserves_invariant ; [apply Hsim |]. } 
+assert (HNalign : ~aligned (pos_list config)).
+{ revert HRNalign. now apply contra, round_preserves_aligned with w ; [apply Hsim|]. }
+induction (non_webloop_will_activate Hsim Hfair HNalign Hinv) as [d config id Hact Hnwl | d config Hnwl_later IHnwl].
++ apply Stream.Later. apply IH ; auto.
+  - unfold lt_config. split ; [apply measure_nonneg|].
+    apply round_decreases_measure_strong with w ; auto.
+    * apply Hsim.
+    * apply Hflex.
+    * exists id. intuition.
+  - apply Hfair.  
+  - apply round_preserves_invariant ; [apply Hsim | auto].
+  - apply Hsim.
+  - apply Hflex.
++ apply Stream.Later. 
+  case (aligned_dec (pos_list (round gatherW (Stream.hd (Stream.tl d)) (round gatherW (Stream.hd d) config)))) 
+    as [HRRalign | HRRNalign].
+  - apply Stream.Later, Stream.Now, aligned_over with w ; auto.
+    * destruct Hsim as [_ Hsim]. apply Hsim.
+    * apply round_preserves_invariant ; [destruct Hsim as [_ Hsim] ; apply Hsim |].
+      apply round_preserves_invariant ; [apply Hsim | auto].
+  - apply IHnwl.
+    * intros c Hc. apply IH. unfold lt_config in Hc |- *.
+      split ; [apply measure_nonneg|].
+      etransitivity ; [apply Hc|].
+      unfold Rminus. apply Rplus_le_compat_r.
+      apply round_decrease_measure with w ; auto. apply Hsim.
+    * apply Hfair.
+    * apply round_preserves_invariant ; [apply Hsim | auto].
+    * apply Hsim. 
+    * apply Hflex.
+    * apply HRRNalign.
+    * apply HRNalign. 
+Qed.
+
+(* This is the main theorem : we assume that initially 
+ * the robots are all looping on their position. 
+ * Then they will eventually be aligned. *)
+Theorem weber_correct config d :
+  Fair d -> 
+  (* Initially no robot is moving. *)
+  config_stay config ->
+  (* Frame changes (chosen by the demon) are similarities centered on the observing robot. *)
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  (* We are in a flexible setting. *)
+  Stream.forever (Stream.instant flex_da_prop) d ->
+  eventually_aligned config d gatherW.
+Proof using lt_0n delta_g0.
+intros Hfair Hstay Hsim Hflex. 
+apply weber_correct_aux with (weber_calc (pos_list config)) ; auto.
+now split ; [apply config_stay_impl_config_stg | apply weber_calc_correct]. 
+Qed.
+
+End Alignment.
diff --git a/CaseStudies/Gathering/InR2/Weber/Align_flex_ssync.v b/CaseStudies/Gathering/InR2/Weber/Align_flex_ssync.v
new file mode 100644
index 0000000000000000000000000000000000000000..031c50f7194922e039d8bdeea14a3d106f89b192
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Align_flex_ssync.v
@@ -0,0 +1,760 @@
+
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file implements an algorithm to ALIGN all robots on an arbitrary 
+ * axis, in the plane (R²). The algorithm assumes there are no byzantine robots,
+ * and works in a FLEXIBLE and SEMI-SYNCHRONOUS setting. 
+
+ * The algorithm is as follows : all robots go towards the 'weber point' of 
+ * the configuration. The weber point, also called geometric median, is unique 
+ * if the robots are not aligned, and has the property that moving any robot
+ * towards the weber point in a straight line doesn't change the weber point. 
+ * It thus remains at the same place throughout the whole execution.  *)
+(**************************************************************************)
+
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+(* Pactole basic definitions *)
+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.Definitions.
+(* Specific to multiplicity *)
+Require Import Pactole.Observations.MultisetObservation.
+(* Specific to flexibility *)
+Require Import Pactole.Models.Flexible.
+(* Specific to settings with no Byzantine robots *)
+Require Import Pactole.Models.NoByzantine.
+(* Specific to definition and properties of the weber point. *)
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Weber_point.
+
+(* User defined *)
+Import Permutation.
+Import Datatypes.
+Import ListNotations.
+
+
+Set Implicit Arguments.
+Close Scope R_scope.
+Close Scope VectorSpace_scope.
+
+
+Section Alignment.
+Local Existing Instances dist_sum_compat R2_VS R2_ES.
+
+(* We assume the existence of a function that calculates a weber point of a collection
+ * (even when the weber point is not unique).
+ * This is a very strong assumption : such a function may not exist in closed form, 
+ * and the Weber point can only be approximated. *)
+Axiom weber_calc : list R2 -> R2.
+Axiom weber_calc_correct : forall ps, Weber ps (weber_calc ps).
+(* We also suppose this function doesn't depend on the order of the points. 
+* This is probably not necessary (we can show that it holds when the points aren't colinear) 
+* but simplifies the proof a bit. *)
+Axiom weber_calc_compat : Proper (PermutationA equiv ==> equiv) weber_calc.
+Local Existing Instance weber_calc_compat.
+  
+(* The number of robots *)
+Variables n : nat.
+Hypothesis lt_0n : 0 < n.
+
+(* There are no byzantine robots. *)
+Local Instance N : Names := Robots n 0.
+Local Instance NoByz : NoByzantine.
+Proof using . now split. Qed.
+
+Lemma list_in_length_n0 {A : Type} x (l : list A) : List.In x l -> length l <> 0.
+Proof using . intros Hin. induction l as [|y l IH] ; cbn ; auto. Qed.
+
+Lemma byz_impl_false : B -> False.
+Proof using . 
+intros b. assert (Hbyz := In_Bnames b). 
+apply list_in_length_n0 in Hbyz. 
+rewrite Bnames_length in Hbyz.
+cbn in Hbyz. intuition.
+Qed.
+
+(* Use this tactic to solve any goal
+ * provided there is a byzantine robot as a hypothesis. *)
+Ltac byz_exfalso :=
+  match goal with 
+  | b : ?B |- _ => exfalso ; apply (byz_impl_false b)
+  end.
+
+(* Since all robots are good robots, we can define a function
+ * from identifiers to good identifiers. *)
+Definition unpack_good (id : ident) : G :=
+  match id with 
+  | Good g => g 
+  | Byz _ => ltac:(byz_exfalso)
+  end.
+
+Lemma good_unpack_good id : Good (unpack_good id) == id.
+Proof using . unfold unpack_good. destruct_match ; [auto | byz_exfalso]. Qed.
+
+Lemma unpack_good_good g : unpack_good (Good g) = g.
+Proof using . reflexivity. Qed.  
+
+(* The robots are in the plane (R^2). *)
+Local Instance Loc : Location := make_Location R2.
+Local Instance LocVS : RealVectorSpace location := R2_VS.
+Local Instance LocES : EuclideanSpace location := R2_ES.
+
+(* Refolding typeclass instances *)
+Ltac foldR2 :=
+  change R2 with location in *;
+  change R2_Setoid with location_Setoid in *;
+  change state_Setoid with location_Setoid in *;
+  change R2_EqDec with location_EqDec in *;
+  change state_EqDec with location_EqDec in *;
+  change R2_VS with LocVS in *;
+  change R2_ES with LocES in *.
+
+(* Robots don't have a state (and thus no memory) apart from their location. *)
+Local Instance St : State location := OnlyLocation (fun f => True).
+(* Robots choose their destination and also the path they take to this destination. *)
+Local Instance RobotC : robot_choice (path location) := {| robot_choice_Setoid := @path_Setoid _ location_Setoid |}.
+
+(* Robots view the other robots' positions up to a similarity. *)
+Local Instance FrameC : frame_choice (similarity location) := FrameChoiceSimilarity.
+Local Instance UpdateC : update_choice ratio := OnlyFlexible.
+Local Instance InactiveC : inactive_choice unit := NoChoiceIna.
+
+(* In a flexible setting, the minimum distance that robots are allowed to move is delta. *)
+Variables delta : R.
+Hypothesis delta_g0 : (0 < delta)%R.
+
+(* We are in a flexible and semi-synchronous setting (for now). *)
+Local Instance UpdateF : update_function (path location) (similarity location) ratio := 
+  FlexibleUpdate delta. 
+(* We are in a semi-synchronous setting : inactive robots don't move. *)
+Local Instance InactiveF : inactive_function _.
+  refine {| inactive := fun config id _ => config id |}.
+Proof using . repeat intro. now subst. Defined.
+
+(* The support of a multiset, but elements are repeated 
+ * a number of times equal to their multiplicity. 
+ * This is needed to convert an observation from multiset to list format, 
+ * so that we can use functions [colinear_dec] and [weber_calc]. *)
+Definition multi_support {A} `{EqDec A} (s : multiset A) :=
+  List.flat_map (fun '(x, mx) => alls x mx) (elements s).
+
+Local Instance multi_support_compat {A} `{EqDec A} : Proper (equiv ==> PermutationA equiv) (@multi_support A _ _).
+Proof using . 
+intros s s' Hss'. unfold multi_support. apply flat_map_compat_perm with eq_pair ; autoclass ; [|]. 
++ intros [x mx] [y my] [H1 H2]. simpl in H1, H2. now rewrite H1, H2.
++ now apply elements_compat.
+Qed.
+
+
+(* The main algorithm : just move towards the weber point
+ * (in a straight line) until all robots are aligned. *)
+Definition gatherW_pgm obs : path location := 
+  if aligned_dec (multi_support obs) 
+  (* Don't move (the robot's local frame is always centered on itself, i.e. its position is at the origin). *)
+  then local_straight_path origin 
+  (* Go towards the weber point. *)
+  else local_straight_path (weber_calc (multi_support obs)).
+
+Local Instance gatherW_pgm_compat : Proper (equiv ==> equiv) gatherW_pgm.
+Proof using .
+intros s1 s2 Hs. unfold gatherW_pgm.
+case ifP_sumbool ; case ifP_sumbool ; intros Halign2 Halign1.
++ reflexivity.
++ rewrite Hs in Halign1. now intuition.
++ rewrite Hs in Halign1. now intuition.
++ f_equiv. apply weber_Naligned_unique with (multi_support s1) ; auto.
+  - rewrite Hs. now apply weber_calc_correct.
+  - now apply weber_calc_correct.
+Qed.
+
+Definition gatherW : robogram := {| pgm := gatherW_pgm |}.
+
+
+Lemma multi_support_add {A : Type} `{EqDec A} s x k : ~ In x s -> k > 0 ->
+  PermutationA equiv (multi_support (add x k s)) (alls x k ++ multi_support s).
+Proof using . 
+intros Hin Hk. unfold multi_support. 
+transitivity (flat_map (fun '(x0, mx) => alls x0 mx) ((x, k) :: elements s)).
++ apply flat_map_compat_perm with eq_pair ; autoclass ; [|].
+  - intros [a ka] [b kb] [H0 H1]. cbn in H0, H1. now rewrite H0, H1.
+  - apply elements_add_out ; auto.
++ now cbn -[elements].
+Qed.
+
+Lemma multi_support_countA {A : Type} `{eq_dec : EqDec A} s x :
+  countA_occ equiv eq_dec x (multi_support s) == s[x]. 
+Proof using .
+pattern s. apply MMultisetFacts.ind.
++ intros m m' Hm. f_equiv. 
+  - apply countA_occ_compat ; autoclass. now rewrite Hm.
+  - now rewrite Hm.
++ intros m x' n' Hin Hn IH. rewrite add_spec, multi_support_add, countA_occ_app by auto.
+  destruct_match.
+  - now rewrite <-H, countA_occ_alls_in, Nat.add_comm, IH ; autoclass.
+  - now rewrite countA_occ_alls_out, IH, Nat.add_0_l ; auto.  
++ now reflexivity.
+Qed.
+
+(* This is the main result about multi_support. *)
+Typeclasses eauto := (bfs).
+Lemma multi_support_config config id : 
+  PermutationA equiv 
+    (multi_support (@obs_from_config _ _ _ _ multiset_observation config (config id))) 
+    (config_list config).
+Proof using .
+cbv -[multi_support config_list equiv make_multiset List.map]. rewrite List.map_id.
+apply PermutationA_countA_occ with R2_EqDec ; autoclass ; []. 
+intros x. rewrite multi_support_countA. now apply make_multiset_spec.
+Qed. 
+
+Corollary multi_support_map f config id : 
+  Proper (equiv ==> equiv) (projT1 f) ->
+  PermutationA equiv 
+    (multi_support (@obs_from_config _ _ _ _ multiset_observation (map_config (lift f) config) (lift f (config id))))
+    (List.map (projT1 f) (config_list config)).
+Proof using .  
+Typeclasses eauto := (dfs).
+intros H. destruct f as [f Pf]. cbn -[equiv config_list multi_support]. 
+change (f (config id)) with (map_config f config id).
+now rewrite multi_support_config, config_list_map.
+Qed.
+
+
+Lemma lift_update_swap da config1 config2 g target :
+  @equiv location _
+    (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                               (precondition_satisfied_inv da config1 g))
+          (update config2
+           g (change_frame da config1 g) target (choose_update da config2 g target)))
+    (update (map_config (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                                      (precondition_satisfied_inv da config1 g)))
+                        config2)
+            g Similarity.id
+            (lift_path (frame_choice_bijection (change_frame da config1 g ⁻¹)) target)
+            (choose_update da config2 g target)).
+Proof using .
+cbn -[inverse dist equiv]. unfold id.
+rewrite Similarity.dist_prop, Rmult_1_l.
+destruct_match_eq Hle; destruct_match_eq Hle' ; try reflexivity ; [|];
+rewrite Rle_bool_true_iff, Rle_bool_false_iff in *;
+exfalso; revert_one not; intro Hgoal; apply Hgoal.
+- assert (Hzoom := Similarity.zoom_pos (change_frame da config1 g)).
+  eapply Rmult_le_reg_l; eauto; []. simpl.
+  rewrite <- Rmult_assoc, Rinv_r, Rmult_1_l; trivial; [].
+  foldR2. lra.
+- assert (Hzoom := Similarity.zoom_pos (change_frame da config1 g ⁻¹)).
+  eapply Rmult_le_reg_l; eauto; []. simpl.
+  rewrite <- Rmult_assoc, Rinv_l, Rmult_1_l; trivial; [].
+  foldR2. generalize (Similarity.zoom_pos (change_frame da config1 g)). lra.
+Qed.
+
+(* Simplify the [round] function and express it in the global frame of reference. *)
+(* All the proofs below use this simplified version. *)
+Lemma round_simplify da config : similarity_da_prop da -> 
+  exists r : G -> ratio,
+  round gatherW da config == 
+  fun id => if activate da id then 
+              if aligned_dec (config_list config) then config id 
+              else 
+                let trajectory := straight_path (config id) (weber_calc (config_list config)) in
+                update config (unpack_good id) Similarity.id trajectory (r (unpack_good id))
+            else config id.
+Proof using . 
+intros Hsim. eexists ?[r]. intros id. unfold round. 
+destruct_match ; [|reflexivity].
+destruct_match ; [|byz_exfalso].
+cbn -[inverse equiv lift precondition frame_choice_bijection config_list origin update].
+rewrite (lift_update_swap da config _ g). 
+pose (f := existT precondition
+  (change_frame da config g)
+  (precondition_satisfied da config g)). 
+pose (f_inv := existT (fun _ : location -> location => True)
+  ((change_frame da config g) ⁻¹)
+  (precondition_satisfied_inv da config g)).
+pose (obs := obs_from_config (map_config (lift f) config) (lift f (config (Good g)))).
+change_LHS (update 
+  (map_config (lift f_inv) (map_config (lift f) config)) g Similarity.id
+  (lift_path
+    (frame_choice_bijection (change_frame da config g ⁻¹))
+    (gatherW_pgm obs))
+  (choose_update da
+    (map_config (lift f) config) g (gatherW_pgm obs))).
+assert (Proper (equiv ==> equiv) (projT1 f)) as f_compat.
+{ unfold f ; cbn -[equiv]. intros x y Hxy ; now rewrite Hxy. }
+assert (Halign_loc_glob : aligned (config_list config) <-> aligned (multi_support obs)).
+{ unfold obs. rewrite multi_support_map by auto. unfold f. cbn -[config_list]. apply R2aligned_similarity. }
+destruct_match.
+(* The robots are aligned. *)
++ unfold gatherW_pgm. destruct_match ; [|intuition].
+  cbn -[equiv lift dist mul inverse]. unfold Datatypes.id.
+  repeat rewrite mul_origin. destruct_match ; apply Hsim.
+(* The robots aren't aligned. *)
++ unfold gatherW_pgm. destruct_match ; [intuition|].
+  pose (sim := change_frame da config g). foldR2. fold sim.
+  assert (Hweb : weber_calc (multi_support obs) == sim (weber_calc (config_list config))).
+  {
+    unfold obs. rewrite multi_support_map by auto. unfold f. cbn -[equiv config_list].
+    foldR2. fold sim. 
+    apply weber_Naligned_unique with (List.map sim (config_list config)).
+    - now rewrite <-R2aligned_similarity.
+    - apply weber_similarity, weber_calc_correct.
+    - apply weber_calc_correct.
+  }
+  change location_Setoid with state_Setoid. apply update_compat ; auto.
+  - intros r. cbn -[equiv]. rewrite Bijection.retraction_section. reflexivity.
+  - intros r. rewrite Hweb. 
+    pose (w := weber_calc (config_list config)). fold w. 
+    pose (c := config (Good g)). fold c.
+    cbn -[equiv w c opp RealVectorSpace.add mul inverse].
+    rewrite sim_mul.  
+    assert (Hcenter : (sim ⁻¹) 0%VS == c).
+    { foldR2. change_LHS (center sim). apply Hsim. }
+    assert (Hsim_cancel : forall x, (inverse sim) (sim x) == x).
+    { cbn -[equiv]. now setoid_rewrite Bijection.retraction_section. }
+    rewrite Hcenter, Hsim_cancel. simplifyR2. 
+    rewrite 2 RealVectorSpace.add_assoc. f_equiv. now rewrite RealVectorSpace.add_comm.
+  - rewrite Hweb. 
+    instantiate (r := fun g => choose_update da (map_config ?[sim] config) g
+      (local_straight_path (?[sim'] (weber_calc (config_list config))))).
+    cbn -[equiv config_list]. reflexivity.
+Qed.
+  
+(* This is the goal (for all demons and configs). *)
+Definition eventually_aligned config (d : demon) (r : robogram) := 
+  Stream.eventually 
+    (Stream.forever (Stream.instant (fun c => aligned (config_list c)))) 
+    (execute r d config).
+
+(* If the robots are aligned, they stay aligned. *)
+Lemma round_preserves_aligned da config : similarity_da_prop da ->
+  aligned (config_list config) -> aligned (config_list (round gatherW da config)).
+Proof using . 
+intros Hsim Halign. assert (round gatherW da config == config) as H.
+{ intros id. destruct (round_simplify config Hsim) as [r Hround].
+  rewrite Hround. repeat destruct_match ; auto. }
+now rewrite H.
+Qed.
+
+Lemma aligned_over config (d : demon) :
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  aligned (config_list config) -> 
+  Stream.forever (Stream.instant (fun c => aligned (config_list c))) (execute gatherW d config).
+Proof using .
+revert config d. 
+cofix Hind. intros config d Hsim Halign. constructor.
++ cbn -[config_list]. apply Halign.
++ cbn -[config_list]. simple apply Hind ; [apply Hsim |]. 
+  apply round_preserves_aligned ; [apply Hsim | apply Halign].
+Qed.
+
+
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine x x' (c c':configuration) : 
+  InA equiv (x, x') (combine (config_list c) (config_list c')) <-> 
+    exists id, (x == c id /\ x' == c' id)%VS.
+Proof using lt_0n.
+  pose (id0 :=(Fin.Fin lt_0n)).
+  pose (g0 :=(Good id0)).
+  pose (default :=(c g0)).
+  pose (default' :=(c' g0)).
+  split.
+  + intros Hin.
+    specialize @config_list_InA_combine_same_lgth2 with (1:=Hin)(dfA:=default)(dfB:=default') as h.
+    destruct h as [i [Hn [Hx Hx']]].
+    rewrite config_list_spec in Hx,Hx'.
+    unfold default in Hx,Hx'.
+    unfold default' in Hx'.
+    rewrite map_nth in Hx .
+    rewrite map_nth in Hx' .
+    exists (nth i names g0);auto.
+  + intros [[g|b] [Hx Hx']] ; [|byz_exfalso]. 
+    rewrite Hx,Hx'.
+    rewrite config_list_spec.
+    rewrite config_list_spec.
+    rewrite combine_map.
+    apply InA_alt.
+    exists (c (Good g),c' (Good g)).
+    split;auto.
+    apply in_map_iff.
+    exists (Good g,Good g);split;auto.
+    apply in_combine_id;split;auto.
+    unfold names.
+    apply in_app_iff.
+    left.
+    apply in_map_iff.
+    exists g;split;auto.
+    apply In_Gnames.
+Qed.
+
+Lemma config_list_In_combine x x' c c' : 
+  List.In (x, x') (combine (config_list c) (config_list c')) <-> 
+  exists id, x == c id /\ x' == c' id.
+Proof using lt_0n.
+  rewrite <- In_InA_is_leibniz with (eqA:=equiv).
+  - apply config_list_InA_combine.
+  - cbn.
+    split;[ intros [h1 h2]| intro h].
+    + destruct x0; destruct y;cbn in *;subst;auto.
+    + now subst.
+Qed.
+
+(* This measure counts how many robots aren't on the weber point. *)
+Definition measure_count config : R := 
+  let ps := config_list config in 
+  (*INR (n - countA_occ equiv R2_EqDec (weber_calc ps) ps).*)
+  list_sum (List.map (fun x => if x =?= weber_calc ps then 0%R else 1%R) ps).
+
+(* This measure counts the total distance from the robots to the weber point. *)
+Definition measure_dist config : R :=
+  let ps := config_list config in 
+  dist_sum ps (weber_calc ps).
+
+(* This measure is positive, and decreases whenever a robot moves. *)
+Definition measure config := (measure_count config + measure_dist config)%R.
+
+Local Instance measure_compat : Proper (equiv ==> equiv) measure.
+Proof using . 
+intros c c' Hc. unfold measure, measure_count, measure_dist.
+assert (Rplus_compat : Proper (equiv ==> equiv ==> equiv) Rplus).
+{ intros x x' Hx y y' Hy. now rewrite Hx, Hy. } 
+apply Rplus_compat.
++ apply list_sum_compat, eqlistA_PermutationA. f_equiv ; [| now rewrite Hc].
+  intros ? ? H. repeat destruct_match ; rewrite H, Hc in * ; intuition.
++ now rewrite Hc.
+Qed.
+
+Lemma measure_nonneg config : (0 <= measure config)%R.
+Proof using . 
+unfold measure. apply Rplus_le_le_0_compat.
++ unfold measure_count. apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros x _. destruct_match ; lra.
++ unfold measure_dist. apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros x _. apply dist_nonneg.   
+Qed.
+
+(* All the magic is here : when the robots move 
+ * they go towards the weber point so it is preserved. 
+ * This still holds in an asynchronous setting.
+ * The point calculated by weber_calc thus stays the same during an execution,
+ * until the robots are colinear. *)
+Lemma round_preserves_weber config da w :
+  similarity_da_prop da -> Weber (config_list config) w -> 
+    Weber (config_list (round gatherW da config)) w.
+Proof using lt_0n. 
+intros Hsim Hweb. apply weber_contract with (config_list config) ; auto.
+unfold contract. rewrite Forall2_Forall, Forall_forall by now repeat rewrite config_list_length.
+intros [x x']. rewrite config_list_In_combine.
+intros [id [Hx Hx']]. destruct (round_simplify config Hsim) as [r Hround].
+revert Hx'. rewrite Hround. 
+repeat destruct_match ; intros Hx' ; rewrite Hx, Hx' ; try apply segment_end.
+assert (w == weber_calc (config_list config)) as Hw.
+{ apply weber_Naligned_unique with (config_list config) ; auto. apply weber_calc_correct. }
+cbn zeta. rewrite <-Hw. cbn -[dist straight_path]. unfold Datatypes.id.
+pose (c := config id). fold c.
+destruct_match ; rewrite segment_sym ; apply segment_straight_path.
+Qed.
+
+(* If the robots don't end up colinear, then the point calculated by weber_calc doesn't change. *)
+Corollary round_preserves_weber_calc config da :
+  similarity_da_prop da -> ~aligned (config_list (round gatherW da config)) -> 
+  weber_calc (config_list (round gatherW da config)) == weber_calc (config_list config). 
+Proof using lt_0n. 
+intros Hsim HNalign.
+apply weber_Naligned_unique with (config_list (round gatherW da config)) ; auto.
++ apply round_preserves_weber ; [auto | apply weber_calc_correct].
++ apply weber_calc_correct.
+Qed.
+
+Lemma Forall2_le_count_weber config da : 
+  similarity_da_prop da -> ~aligned (config_list (round gatherW da config)) ->
+  Forall2 Rle
+    (List.map
+      (fun x : R2 => if x =?= weber_calc (config_list (round gatherW da config)) then 0%R else 1%R) 
+      (config_list (round gatherW da config)))
+    (List.map
+      (fun x : R2 => if x =?= weber_calc (config_list config) then 0%R else 1%R)
+      (config_list config)).
+Proof using lt_0n. 
+intros Hsim RNcol.  
+assert (Hweb := round_preserves_weber_calc config Hsim RNcol).
+rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length, config_list_length.
+intros [x' x] Hin. apply config_list_In_combine in Hin. destruct Hin as [id [Hx Hx']].
+repeat destruct_match ; try lra. rewrite Hx, Hx', Hweb in *. rename H into Hround_id. rename H0 into Hconfig_id.
+assert (Hround : round gatherW da config id == weber_calc (config_list config)).
+{ 
+  destruct (round_simplify config Hsim) as [r Hround].
+  rewrite Hround. repeat destruct_match ; auto.
+  cbn zeta. rewrite <-Hconfig_id. cbn -[equiv dist mul opp RealVectorSpace.add]. 
+  simplifyR2. now destruct_match.
+}
+rewrite <-Hround in Hround_id. intuition.
+Qed.
+
+Lemma Forall2_le_dist_weber config da : 
+  similarity_da_prop da -> ~aligned (config_list (round gatherW da config)) ->
+  Forall2 Rle
+    (List.map 
+      (dist (weber_calc (config_list (round gatherW da config))))
+      (config_list (round gatherW da config)))
+    (List.map 
+      (dist (weber_calc (config_list config))) 
+      (config_list config)).
+Proof using lt_0n. 
+intros Hsim RNcol.
+assert (Hweb := round_preserves_weber_calc config Hsim RNcol).
+rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length, config_list_length.
+intros [x' x] Hin. apply config_list_In_combine in Hin. destruct Hin as [id [Hx' Hx]].
+rewrite Hx, Hx', Hweb. destruct (round_simplify config Hsim) as [r Hround].
+rewrite Hround. repeat destruct_match ; try lra.
+cbn zeta. pose (w := weber_calc (config_list config)). fold w. rewrite <-Hx.
+pose (ri := r (unpack_good id)). fold ri.
+cbn -[dist RealVectorSpace.add mul opp w ri]. destruct_match.
++ repeat rewrite norm_dist. rewrite R2_opp_dist, RealVectorSpace.add_assoc.
+  rewrite <-(mul_1 (w - x)) at 1. rewrite <-minus_morph, add_morph, norm_mul.
+  rewrite <-Rmult_1_l. apply Rmult_le_compat_r ; try apply norm_nonneg.
+  unfold Rabs. destruct_match ; generalize (ratio_bounds ri) ; lra.
++ rewrite mul_1, (RealVectorSpace.add_comm w), RealVectorSpace.add_assoc.
+  simplifyR2. rewrite dist_same. apply dist_nonneg.
+Qed.
+
+(* If a robot moves, either the measure strictly decreases or the robots become colinear. *)
+Lemma round_decreases_measure config da : 
+  similarity_da_prop da ->
+  moving gatherW da config <> nil -> 
+    aligned (config_list (round gatherW da config)) \/ 
+    (measure (round gatherW da config) <= measure config - Rmin delta 1)%R.
+Proof using lt_0n. 
+intros Hsim Hmove. 
+destruct (aligned_dec (config_list (round gatherW da config))) as [Rcol | RNcol] ; [now left|right].
+assert (Hweb := round_preserves_weber_calc config Hsim RNcol).
+destruct (not_nil_In Hmove) as [i Hi]. apply moving_spec in Hi.
+destruct (round gatherW da config i =?= weber_calc (config_list config)) as [Hreached | HNreached].
+(* The robot that moved reached its destination. *)
++ transitivity (measure config - 1)%R ; [|generalize (Rmin_r delta 1%R) ; lra].
+  unfold measure, Rminus. rewrite Rplus_assoc, (Rplus_comm _ (-1)%R), <-Rplus_assoc.
+  apply Rplus_le_compat.
+  - unfold measure_count. apply list_sum_le_eps ; [now apply Forall2_le_count_weber|].
+    rewrite combine_map, Exists_map. apply Exists_exists. 
+    exists (round gatherW da config i, config i).
+    split ; [apply config_list_In_combine ; exists i ; intuition |].
+    repeat destruct_match ; solve [lra | rewrite Hreached in * ; intuition]. 
+  - unfold measure_dist. apply list_sum_le. now apply Forall2_le_dist_weber.
+(* The robots that moved didn't reach its destination. *)
++ transitivity (measure config - delta)%R ; [|generalize (Rmin_l delta 1%R) ; lra].
+  unfold measure, Rminus. rewrite Rplus_assoc. 
+  apply Rplus_le_compat.
+  - unfold measure_count. apply list_sum_le. now apply Forall2_le_count_weber.
+  - unfold measure_dist. apply list_sum_le_eps ; [now apply Forall2_le_dist_weber|].
+    rewrite combine_map, Exists_map. apply Exists_exists. 
+    exists (round gatherW da config i, config i).
+    split ; [apply config_list_In_combine ; exists i ; intuition |].
+    rewrite Hweb. destruct (round_simplify config Hsim) as [r Hround].
+    rewrite Hround. destruct_match_eq Hact ; [destruct_match_eq Halign|].
+    * exfalso. now apply RNcol, round_preserves_aligned.
+    * cbn zeta. 
+      pose (w := weber_calc (config_list config)). foldR2. fold w. 
+      pose (x := config i). fold x.
+      pose (ri := r (unpack_good i)). fold ri.
+      cbn -[dist RealVectorSpace.add mul opp w ri]. rewrite Rmult_1_l, mul_1.
+      destruct_match_eq Hdelta ; unfold id in * ; rewrite good_unpack_good in * ; fold x in Hdelta.
+      --rewrite Rle_bool_true_iff in Hdelta. repeat rewrite norm_dist in *.
+        rewrite R2_opp_dist, RealVectorSpace.add_assoc in Hdelta |- *. 
+        rewrite add_opp, add_origin_l, norm_opp, norm_mul, Rabs_pos_eq in Hdelta by (generalize (ratio_bounds ri) ; lra).
+        rewrite <-(mul_1 (w - x)) at 1. rewrite <-minus_morph, add_morph, norm_mul.
+        rewrite Rabs_pos_eq ; [|generalize (ratio_bounds ri) ; lra].
+        rewrite Rmult_plus_distr_r, Rmult_1_l.
+        simpl location in Hdelta |- *.
+        apply Rplus_le_compat ; try lra.
+        rewrite <-Ropp_mult_distr_l. now apply Ropp_le_contravar.
+      --exfalso. apply HNreached. rewrite Hround. destruct_match ; [|intuition].
+        cbn -[config_list dist mul opp RealVectorSpace.add].
+        rewrite Rmult_1_l, good_unpack_good ; unfold id ; fold x ; fold ri ; foldR2 ; fold w.
+        destruct_match ; intuition. 
+        rewrite mul_1, (RealVectorSpace.add_comm w), RealVectorSpace.add_assoc. 
+        now simplifyR2.
+    * exfalso. apply Hi. rewrite Hround. destruct_match ; intuition.
+Qed.
+
+Typeclasses eauto := (bfs).
+Lemma gathered_aligned ps x : 
+  (Forall (fun y => y == x) ps) -> aligned ps.
+Proof using . 
+Typeclasses eauto := (dfs).
+rewrite Forall_forall. intros Hgathered.
+exists (Build_line x 0%VS). intros y Hin. unfold line_contains.
+rewrite line_proj_spec, line_proj_inv_spec. cbn -[equiv RealVectorSpace.add opp mul inner_product].
+unfold Rdiv. rewrite inner_product_origin_r, Rmult_0_l, mul_0, add_origin_r. 
+symmetry. apply Hgathered. rewrite <-In_InA_is_leibniz ; autoclass. now cbn.
+Qed.
+
+(* If the robots aren't aligned yet then there exists at least one robot which, 
+ * if activated, will move. 
+ * Any robot that isn't on the weber point will do the trick. *)
+Lemma one_must_move config : ~aligned (config_list config) ->
+  exists i, forall da, similarity_da_prop da -> activate da i = true ->
+                       round gatherW da config i =/= config i.
+Proof using delta_g0.
+intros Nalign.
+pose (w := weber_calc (config_list config)).
+cut (exists i, config i =/= w). 
+{
+  intros [i Hi]. exists i. intros da Hsim Hact. 
+  destruct (round_simplify config Hsim) as [r Hround]. rewrite Hround.
+  repeat destruct_match ; try intuition. clear Hact. rename H into Hactiv. rename H0 into Halign.
+  cbn -[opp mul RealVectorSpace.add dist config_list equiv complement].
+  rewrite Rmult_1_l, mul_1, good_unpack_good ; unfold id ; fold w.
+  destruct_match_eq Hdelta.
+  + intros H. 
+    rewrite <-(add_origin_r (config i)) in H at 3 ; apply add_reg_l in H.
+    rewrite mul_eq0_iff in H. destruct H as [H1|H2].
+    - rewrite H1, mul_0, add_origin_r, dist_same, Rle_bool_true_iff in Hdelta. lra.
+    - rewrite R2sub_origin in H2. intuition.
+  + rewrite add_sub. intuition.
+}
+assert (List.Exists (fun x => x =/= weber_calc (config_list config)) (config_list config)) as HE.
+{ 
+  apply neg_Forall_Exists_neg ; [intros ; apply equiv_dec|].
+  revert Nalign. apply contra. apply gathered_aligned.
+}
+rewrite Exists_exists in HE. destruct HE as [x [Hin Hx]].
+apply (@In_InA R2 equiv _) in Hin. 
+foldR2. change location_Setoid with state_Setoid in *. rewrite config_list_InA in Hin.
+destruct Hin as [r Hr]. exists r. now rewrite <-Hr.
+Qed.
+
+
+Lemma no_moving_same_config : forall r config da,
+  moving r da config = [] -> round r da config == config.
+Proof using Type.
+intros.
+apply no_changing_same_config.
+assumption.
+Qed.
+
+(* Fairness entails progress. *)
+Lemma fair_first_move (d : demon) config : 
+  Fair d -> Stream.forever (Stream.instant similarity_da_prop) d ->
+  ~(aligned (config_list config)) -> FirstMove gatherW d config.
+Proof using delta_g0.
+intros Hfair Hsim Nalign.
+destruct (one_must_move config Nalign) as [id Hmove].
+destruct Hfair as [locallyfair Hfair].
+specialize (locallyfair id).
+revert config Nalign Hmove.
+induction locallyfair as [d Hnow | d] ; intros config Nalign Hmove.
+* apply MoveNow. apply Hmove in Hnow.
+  + assert (get_location (round gatherW (Stream.hd d) config id) =/= get_location (config id)) as Hnow'.
+    { cbn -[complement equiv].
+      assumption. }
+    rewrite <-(moving_spec gatherW (Stream.hd d) config id) in Hnow'.
+    intros Habs. now rewrite Habs in Hnow'.
+  + apply Hsim.
+* destruct (moving gatherW (Stream.hd d) config) as [| id' mov] eqn:Hmoving.
+  + apply MoveLater ; trivial.
+    apply IHlocallyfair.
+    - apply Hfair.
+    - apply Hsim.
+    - apply no_moving_same_config in Hmoving. now rewrite Hmoving.
+    - intros da Hda Hactive. apply no_moving_same_config in Hmoving.
+      rewrite (Hmoving id).
+      apply (round_compat (reflexivity gatherW) (reflexivity da)) in Hmoving. 
+      rewrite (Hmoving id).
+      now apply Hmove.
+  + apply MoveNow. rewrite Hmoving. discriminate.
+Qed.
+
+(* This is the well founded relation we will perform induction on. *)
+Definition lt_config eps c c' := 
+  (0 <= measure c <= measure c' - eps)%R. 
+
+Local Instance lt_config_compat : Proper (equiv ==> equiv ==> equiv ==> iff) lt_config.
+Proof using .
+  intros ? ? H1 ? ? H2 ? ? H3.
+  unfold lt_config.
+  rewrite H1.
+  rewrite H2 at 1.
+  rewrite H3 at 1.
+  rewrite H2.
+  reflexivity.
+Qed.
+
+
+(* We proove this using the well-foundedness of lt on nat. *)
+Lemma lt_config_wf eps : (eps > 0)%R -> well_founded (lt_config eps).
+Proof using . 
+intros Heps. unfold well_founded. intros c.
+pose (f := fun x : R => Z.to_nat (up (x / eps))).
+remember (f (measure c)) as k. generalize dependent c. 
+pattern k. apply (well_founded_ind lt_wf). clear k.
+intros k IH c Hk. apply Acc_intro. intros c' Hc'. apply IH with (f (measure c')) ; auto.
+rewrite Hk ; unfold f ; unfold lt_config in Hc'.
+rewrite <-Z2Nat.inj_lt.
++ apply Zup_lt. unfold Rdiv. rewrite <-(Rinv_r eps) by lra. 
+  rewrite <-Rmult_minus_distr_r. apply Rmult_le_compat_r ; intuition.
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition. 
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition.
+  transitivity eps ; intuition. 
+  apply (Rplus_le_reg_r (- eps)). rewrite Rplus_opp_r. etransitivity ; eauto.
+Qed.
+
+(* The proof is essentially a well-founded induction on [measure config].
+ * Fairness ensures that the measure must decrease at some point. *)
+Theorem weber_correct config : forall d,
+  Fair d -> 
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  eventually_aligned config d gatherW.
+Proof using delta_g0 lt_0n.
+assert (Hdelta1 : (Rmin delta 1 > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+induction config as [config IH] using (well_founded_ind (lt_config_wf Hdelta1)).
+intros d Hfair Hsim.
+destruct (aligned_dec (config_list config)) as [align | Nalign] ;
+  [now apply Stream.Now, aligned_over|].
+induction (fair_first_move config Hfair Hsim Nalign) as [d config Hmove | d config Hmove FM IH_FM] ;
+  destruct Hsim as [Hsim_hd Hsim_tl] ; cbn in Hsim_hd ; apply Stream.Later.
++ destruct (round_decreases_measure config Hsim_hd Hmove) as [Ralign | Rmeasure].
+  - now apply Stream.Now, aligned_over.
+  - apply IH.
+    * unfold lt_config. split ; [apply measure_nonneg | apply Rmeasure].  
+    * apply Hfair. 
+    * apply Hsim_tl. 
++ apply no_moving_same_config in Hmove. cbn -[config_list].
+  apply IH_FM.
+  - intros c' Hc' d' Hfair' Hsim'. rewrite Hmove in Hc'. apply IH ; auto.
+  - apply Hfair.
+  - apply Hsim_tl.
+  - now rewrite Hmove.
+Qed.  
+
+End Alignment.
diff --git a/CaseStudies/Gathering/InR2/Weber/Align_rigid_ssync.v b/CaseStudies/Gathering/InR2/Weber/Align_rigid_ssync.v
new file mode 100644
index 0000000000000000000000000000000000000000..322ed8391cec231a247586faf6183750b9c7384d
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Align_rigid_ssync.v
@@ -0,0 +1,533 @@
+
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file implements an algorithm to ALIGN all robots on an arbitrary 
+ * axis, in the plane (R²). The algorithm assumes there are no byzantine robots,
+ * and works in a RIGID and SEMI-SYNCHRONOUS setting.
+
+ * The algorithm is as follows : all robots go towards the 'weber point' of 
+ * the configuration. The weber point, also called geometric median, is unique 
+ * if the robots are not aligned, and has the property that moving any robot
+ * towards the weber point in a straight line doesn't change the weber point. 
+ * It thus remains at the same place throughout the whole execution.  *)
+(**************************************************************************)
+
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+Require Import Pactole.Util.ListComplements.
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+(* Specific to R^2 topology *)
+Require Import Pactole.Spaces.R2.
+(* Pactole basic definitions *)
+Require Export Pactole.Setting.
+(* Specific to gathering *)
+Require Pactole.CaseStudies.Gathering.WithMultiplicity.
+Require Import Pactole.CaseStudies.Gathering.Definitions.
+(* Specific to multiplicity *)
+Require Import Pactole.Util.Bijection.
+(* Specific to rigidity *)
+Require Import Pactole.Models.Rigid.
+(* Specific to settings with no Byzantine robots *)
+Require Import Pactole.Models.NoByzantine.
+(* Specific to definition and properties of the weber point. *)
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Weber_point.
+Require Import Pactole.Observations.MultisetObservation.
+
+
+(* User defined *)
+Import Permutation.
+Import Datatypes.
+Import ListNotations.
+
+Set Implicit Arguments.
+Close Scope R_scope.
+Close Scope VectorSpace_scope.
+
+
+
+Section Alignment.
+Local Existing Instances dist_sum_compat R2_VS R2_ES.
+
+(* We assume the existence of a function that calculates a weber point of a collection
+ * (even when the weber point is not unique).
+ * This is a very strong assumption : such a function may not exist in closed form, 
+ * and the Weber point can only be approximated. *)
+Axiom weber_calc : list R2 -> R2.
+Axiom weber_calc_correct : forall ps, Weber ps (weber_calc ps).
+(* We also suppose this function doesn't depend on the order of the points. 
+* This is probably not necessary (we can show that it holds when the points aren't colinear) 
+* but simplifies the proof a bit. *)
+Axiom weber_calc_compat : Proper (PermutationA equiv ==> equiv) weber_calc.
+Local Existing Instance weber_calc_compat.
+ 
+(* The number of robots *)
+Variables n : nat.
+Hypothesis lt_0n : 0 < n.
+
+(* There are no byzantine robots. *)
+Local Instance N : Names := Robots n 0.
+Local Instance NoByz : NoByzantine.
+Proof using . now split. Qed.
+
+Lemma list_in_length_n0 {A : Type} x (l : list A) : List.In x l -> length l <> 0.
+Proof using . intros Hin. induction l as [|y l IH] ; cbn ; auto. Qed.
+
+Lemma byz_impl_false : B -> False.
+Proof using . 
+intros b. assert (Hbyz := In_Bnames b). 
+apply list_in_length_n0 in Hbyz. 
+rewrite Bnames_length in Hbyz.
+cbn in Hbyz. intuition.
+Qed.
+
+(* Use this tactic to solve any goal
+ * provided there is a byzantine robot as a hypothesis. *)
+Ltac byz_exfalso :=
+  match goal with 
+  | b : ?B |- _ => exfalso ; apply (byz_impl_false b)
+  end.
+
+(* Since all robots are good robots, we can define a function
+ * from identifiers to good identifiers. *)
+Definition unpack_good (id : ident) : G :=
+  match id with 
+  | Good g => g 
+  | Byz _ => ltac:(byz_exfalso)
+  end.
+
+Lemma good_unpack_good id : Good (unpack_good id) == id.
+Proof using . unfold unpack_good. destruct_match ; [auto | byz_exfalso]. Qed.
+
+Lemma unpack_good_good g : unpack_good (Good g) = g.
+Proof using . reflexivity. Qed.  
+
+(* The robots are in the plane (R^2). *)
+Local Instance Loc : Location := make_Location R2.
+Local Instance LocVS : RealVectorSpace location := R2_VS.
+Local Instance LocES : EuclideanSpace location := R2_ES.
+
+(* Refolding typeclass instances *)
+Ltac foldR2 :=
+  change R2 with location in *;
+  change R2_Setoid with location_Setoid in *;
+  change state_Setoid with location_Setoid in *;
+  change R2_EqDec with location_EqDec in *;
+  change state_EqDec with location_EqDec in *;
+  change R2_VS with LocVS in *;
+  change R2_ES with LocES in *.
+
+
+(* Robots don't have an state (and thus no memory) apart from their location. *)
+Local Instance St : State location := OnlyLocation (fun f => True).
+Local Instance RobotC : robot_choice location := {| robot_choice_Setoid := location_Setoid |}.
+
+(* Robots view the other robots' positions up to a similarity. *)
+Local Instance FrameC : frame_choice (similarity location) := FrameChoiceSimilarity.
+Local Instance UpdateC : update_choice unit := NoChoice.
+Local Instance InactiveC : inactive_choice unit := NoChoiceIna.
+
+(* We are in a rigid and semi-synchronous setting. *)
+Local Instance UpdateF : update_function _ _ _.
+  refine {| update := fun _ _ _ target _ => target |}.
+Proof using . now repeat intro. Defined. 
+Local Instance InactiveF : inactive_function _.
+  refine {| inactive := fun config id _ => config id |}.
+Proof using . repeat intro. now subst. Defined.
+Local Instance Rigid : RigidSetting.
+Proof using . split. reflexivity. Qed.
+
+(* The support of a multiset, but elements are repeated 
+ * a number of times equal to their multiplicity. 
+ * This is needed to convert an observation from multiset to list format, 
+ * so that we can use functions [colinear_dec] and [weber_calc]. *)
+Definition multi_support {A} `{EqDec A} (s : multiset A) :=
+  List.flat_map (fun '(x, mx) => alls x mx) (elements s).
+
+Local Instance multi_support_compat {A} `{EqDec A} : Proper (equiv ==> PermutationA equiv) (@multi_support A _ _).
+Proof using . 
+intros s s' Hss'. unfold multi_support. apply flat_map_compat_perm with eq_pair ; autoclass ; [|]. 
++ intros [x mx] [y my] [H1 H2]. simpl in H1, H2. now rewrite H1, H2.
++ now apply elements_compat.
+Qed.
+
+
+(* The main algorithm : just move towards the weber point until all robots are aligned. *)
+Definition gatherW_pgm obs : location := 
+  if aligned_dec (multi_support obs) 
+  (* Don't move (the robot's local frame is always centered on itself, i.e. its position is at the origin). *)
+  then origin 
+  (* Go towards the weber point. *)
+  else weber_calc (multi_support obs).
+
+Local Instance gatherW_pgm_compat : Proper (equiv ==> equiv) gatherW_pgm.
+Proof using .
+intros s1 s2 Hs. unfold gatherW_pgm.
+repeat destruct_match.
++ reflexivity.
++ rewrite Hs in H. now intuition.
++ rewrite Hs in H. now intuition.
++ apply weber_Naligned_unique with (multi_support s1) ; auto.
+  - rewrite Hs. now apply weber_calc_correct.
+  - now apply weber_calc_correct.
+Qed.
+
+Definition gatherW : robogram := {| pgm := gatherW_pgm |}.
+
+Lemma multi_support_add {A : Type} `{EqDec A} s x k : ~ In x s -> k > 0 ->
+  PermutationA equiv (multi_support (add x k s)) (alls x k ++ multi_support s).
+Proof using . 
+intros Hin Hk. unfold multi_support. 
+transitivity (flat_map (fun '(x0, mx) => alls x0 mx) ((x, k) :: elements s)).
++ apply flat_map_compat_perm with eq_pair ; autoclass.
+  - intros [a ka] [b kb] [H0 H1]. cbn in H0, H1. now rewrite H0, H1.
+  - apply elements_add_out ; auto.
++ now cbn -[elements].
+Qed.
+
+Lemma multi_support_countA {A : Type} `{eq_dec : EqDec A} s x :
+  countA_occ equiv eq_dec x (multi_support s) == s[x]. 
+Proof using .
+pattern s. apply MMultisetFacts.ind.
++ intros m m' Hm. f_equiv. 
+  - apply countA_occ_compat ; autoclass. now rewrite Hm.
+  - now rewrite Hm.
++ intros m x' n' Hin Hn IH. rewrite add_spec, multi_support_add, countA_occ_app by auto.
+  destruct_match.
+  - now rewrite <-H, countA_occ_alls_in, Nat.add_comm, IH ; autoclass.
+  - now rewrite countA_occ_alls_out, IH, Nat.add_0_l ; auto.  
++ now reflexivity.
+Qed.
+
+Typeclasses eauto := (bfs) 3.
+(* This is the main result about multi_support. *)
+Lemma multi_support_config (config:configuration) id : 
+  PermutationA equiv 
+    (multi_support (@obs_from_config _ _ _ _ multiset_observation config (config id))) 
+    (config_list config).
+Proof using .
+cbv -[multi_support config_list equiv make_multiset List.map]. rewrite List.map_id.
+apply PermutationA_countA_occ with R2_EqDec ; autoclass. 
+intros x. rewrite multi_support_countA. now apply make_multiset_spec.
+Qed. 
+
+Typeclasses eauto := (bfs).
+Corollary multi_support_map f config id : 
+  Proper (equiv ==> equiv) (projT1 f) ->
+  PermutationA equiv 
+    (multi_support (@obs_from_config _ _ _ _ multiset_observation (map_config (lift f) config) (lift f (config id))))
+    (List.map (projT1 f) (config_list config)).
+Proof using .  
+intros H. destruct f as [f Pf]. cbn -[equiv config_list multi_support]. 
+change (f (config id)) with (map_config f config id).
+Typeclasses eauto := (dfs).
+now rewrite multi_support_config, config_list_map.
+Qed.
+
+(* Simplify the [round] function and express it in the global frame of reference. *)
+(* All the proofs below use this simplified version. *)
+Lemma round_simplify da config : similarity_da_prop da -> 
+  round gatherW da config == 
+  fun id => 
+    if activate da id then 
+      if aligned_dec (config_list config) then config id 
+      else weber_calc (config_list config)
+    else config id.
+Proof using . 
+intros Hsim. apply no_byz_eq. intros g. unfold round. 
+cbn -[inverse equiv lift location config_list origin].
+destruct_match ; try reflexivity.
+pose (f := existT (fun _ : location -> location => True)
+  (frame_choice_bijection (change_frame da config g))
+  (precondition_satisfied da config g)).
+pose (f_inv := existT (fun _ : location -> location => True)
+  (frame_choice_bijection (change_frame da config g) ⁻¹)
+  (precondition_satisfied_inv da config g)).
+change_LHS (lift f_inv (gatherW_pgm (obs_from_config 
+  (map_config (lift f) config) 
+  ((lift f) (config (Good g)))
+))).
+assert (Proper (equiv ==> equiv) (projT1 f)) as f_compat.
+{ unfold f ; cbn -[equiv]. intros x y Hxy ; now rewrite Hxy. }
+unfold gatherW_pgm. case ifP_sumbool ; intros Halign.
++ rewrite multi_support_map in Halign by auto.
+  cbn -[equiv inverse config_list location] in *. 
+  rewrite <-R2aligned_similarity in Halign. change_LHS (center (change_frame da config g)).
+  rewrite Hsim ; cbn -[equiv config_list] ; unfold id.
+  now destruct_match ; intuition.
++ rewrite multi_support_map in * by auto.
+  cbn -[equiv inverse config_list location multi_support] in *.
+  pose (sim := change_frame da config g) ; fold sim in Halign ; fold sim.
+  rewrite <-R2aligned_similarity in Halign. case ifP_sumbool ; [now intuition | intros _].
+  apply weber_Naligned_unique with (config_list config) ; auto ; [now apply weber_calc_correct|].
+  apply weber_similarity with sim. cbn -[config_list]. rewrite Bijection.section_retraction.
+  now apply weber_calc_correct.
+Qed.
+  
+(* This is the goal (for all demons and configs). *)
+Definition eventually_aligned config (d : demon) (r : robogram) := 
+  Stream.eventually 
+    (Stream.forever (Stream.instant (fun c => aligned (config_list c)))) 
+    (execute r d config).
+
+(* If the robots are aligned, they stay aligned. *)
+Lemma round_preserves_aligned da config : similarity_da_prop da ->
+  aligned (config_list config) -> aligned (config_list (round gatherW da config)).
+Proof using . 
+intros Hsim Halign. assert (round gatherW da config == config) as H.
+{ intros id. rewrite round_simplify by auto. repeat destruct_match ; auto. }
+now rewrite H.
+Qed.
+
+Lemma aligned_over config (d : demon) :
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  aligned (config_list config) -> 
+  Stream.forever (Stream.instant (fun c => aligned (config_list c))) (execute gatherW d config).
+Proof using .
+revert config d. 
+cofix Hind. intros config d Hsim Halign. constructor.
++ cbn -[config_list]. apply Halign.
++ cbn -[config_list]. simple apply Hind ; [apply Hsim |]. 
+  apply round_preserves_aligned ; [apply Hsim | apply Halign].
+Qed.
+
+Lemma sub_lt_sub (i j k : nat) : j < i <= k -> k - i < k - j.
+Proof using lt_0n. lia. Qed.
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine x x' (c c':configuration) : 
+  InA equiv (x, x') (combine (config_list c) (config_list c')) <-> 
+    exists id, (x == c id /\ x' == c' id)%VS.
+Proof using lt_0n.
+  pose (id0 :=(Fin.Fin lt_0n)).
+  pose (g0 :=(Good id0)).
+  pose (default :=(c g0)).
+  pose (default' :=(c' g0)).
+  split.
+  + intros Hin.
+    specialize @config_list_InA_combine_same_lgth2 with (1:=Hin)(dfA:=default)(dfB:=default') as h.
+    destruct h as [i [Hn [Hx Hx']]].
+    rewrite config_list_spec in Hx,Hx'.
+    unfold default in Hx,Hx'.
+    unfold default' in Hx'.
+    rewrite map_nth in Hx .
+    rewrite map_nth in Hx' .
+    exists (nth i names g0);auto.
+  + intros [[g|b] [Hx Hx']] ; [|byz_exfalso]. 
+    rewrite Hx,Hx'.
+    rewrite config_list_spec.
+    rewrite config_list_spec.
+    rewrite combine_map.
+    apply InA_alt.
+    exists (c (Good g),c' (Good g)).
+    split;auto.
+    apply in_map_iff.
+    exists (Good g,Good g);split;auto.
+    apply in_combine_id;split;auto.
+    unfold names.
+    apply in_app_iff.
+    left.
+    apply in_map_iff.
+    exists g;split;auto.
+    apply In_Gnames.
+Qed.
+
+Lemma config_list_In_combine x x' c c' : 
+  List.In (x, x') (combine (config_list c) (config_list c')) <-> 
+  exists id, x == c id /\ x' == c' id.
+Proof using lt_0n.
+  rewrite <- In_InA_is_leibniz with (eqA:=equiv).
+  - apply config_list_InA_combine.
+  - cbn.
+    split;[ intros [h1 h2]| intro h].
+    + destruct x0; destruct y;cbn in *;subst;auto.
+    + now subst.
+Qed.
+
+(* This measure strictly decreases whenever a robot moves. *)
+Definition measure config := 
+  let ps := config_list config in 
+  n - countA_occ equiv R2_EqDec (weber_calc ps) ps.
+
+Local Instance measure_compat : Proper (equiv ==> eq) measure.
+Proof using . intros c c' Hc. unfold measure. now rewrite Hc. Qed.
+
+(* All the magic is here : when the robots move 
+ * they go towards the weber point so it is preserved. 
+ * This still holds in a flexible and/or asynchronous setting.
+ * The point calculated by weber_calc thus stays the same during an execution,
+ * until the robots are colinear. *)
+Lemma round_preserves_weber config da w :
+  similarity_da_prop da -> Weber (config_list config) w -> 
+    Weber (config_list (round gatherW da config)) w.
+Proof using lt_0n. 
+intros Hsim Hweb. apply weber_contract with (config_list config) ; auto.
+unfold contract. rewrite Forall2_Forall, Forall_forall by now repeat rewrite config_list_length.
+intros [x x'].
+ rewrite config_list_In_combine.
+intros [id [Hx Hx']]. revert Hx'. rewrite round_simplify by auto. 
+repeat destruct_match ; intros Hx' ; rewrite Hx, Hx' ; try apply segment_end.
+assert (w == weber_calc (config_list config)) as Hw.
+{ apply weber_Naligned_unique with (config_list config) ; auto. now apply weber_calc_correct. }
+rewrite <-Hw. apply segment_start.
+Qed.
+
+(* If a robot moves, either the measure decreases or the robots become colinear. *)
+Lemma round_decreases_measure config da : 
+  similarity_da_prop da ->
+  moving gatherW da config <> nil -> 
+    aligned (config_list (round gatherW da config)) \/ 
+    measure (round gatherW da config) < measure config.
+Proof using lt_0n. 
+intros Hsim Hmove. 
+destruct (aligned_dec (config_list (round gatherW da config))) as [Rcol | RNcol] ; [now left|right].
+assert (weber_calc (config_list (round gatherW da config)) == weber_calc (config_list config)) as Hweb.
+{ 
+  apply weber_Naligned_unique with (config_list (round gatherW da config)) ; auto.
+  + apply round_preserves_weber ; [auto | apply weber_calc_correct].
+  + apply weber_calc_correct.  
+}
+unfold measure. apply sub_lt_sub. split.
++ destruct (not_nil_In Hmove) as [i Hi]. apply moving_spec in Hi.
+  rewrite Hweb. apply countA_occ_lt.
+  - rewrite Forall2_Forall, Forall_forall. intros [x x'] Hin.
+    apply config_list_In_combine in Hin. destruct Hin as [j [-> ->]].
+    rewrite round_simplify by auto. repeat destruct_match ; intuition.
+    repeat rewrite config_list_length. reflexivity.
+  - apply Exists_exists. exists (config i, round gatherW da config i).
+    split ; [| now auto].
+    apply config_list_In_combine. exists i ; intuition.
++ etransitivity ; [apply countA_occ_length_le|].
+  rewrite config_list_length. cbn ; lia.
+Qed.
+
+Typeclasses eauto := (bfs).
+
+Lemma gathered_aligned ps x : 
+  (Forall (fun y => y == x) ps) -> aligned ps.
+Proof using . 
+Typeclasses eauto := (dfs).
+rewrite Forall_forall. intros Hgathered.
+exists (Build_line x 0%VS). intros y Hin. unfold line_contains.
+rewrite line_proj_spec, line_proj_inv_spec. cbn -[equiv RealVectorSpace.add opp mul inner_product].
+unfold Rdiv. rewrite inner_product_origin_r, Rmult_0_l, mul_0, add_origin_r. 
+symmetry. apply Hgathered. rewrite <-In_InA_is_leibniz ; autoclass. now cbn.
+Qed.
+
+(* If the robots aren't aligned yet then there exists at least one robot which, 
+ * if activated, will move. 
+ * Any robot that isn't on the weber point will do the trick. *)
+Lemma one_must_move config : ~aligned (config_list config) ->
+  exists r, forall da, similarity_da_prop da -> activate da r = true ->
+                       round gatherW da config r =/= config r.
+Proof using .
+intros Nalign.
+cut (exists r, config r =/= weber_calc (config_list config)). 
+{
+  intros [r Hr]. exists r. intros da Hsim Hact. rewrite round_simplify by auto.
+  repeat destruct_match ; intuition.
+}
+assert (List.Exists (fun x => x =/= weber_calc (config_list config)) (config_list config)) as HE.
+{ 
+  apply neg_Forall_Exists_neg ; [intros ; apply equiv_dec|].
+  revert Nalign. apply contra. apply gathered_aligned.
+}
+rewrite Exists_exists in HE. destruct HE as [x [Hin Hx]].
+apply (@In_InA R2 equiv _) in Hin. 
+foldR2. change location_Setoid with state_Setoid in *. rewrite config_list_InA in Hin.
+destruct Hin as [r Hr]. exists r. now rewrite <-Hr.
+Qed.
+
+Lemma no_moving_same_config : forall r config da,
+  moving r da config = [] -> round r da config == config.
+Proof using Type.
+intros.
+apply no_changing_same_config.
+assumption.
+Qed.
+
+
+(* Fairness entails progress. *)
+Lemma fair_first_move (d : demon) config : 
+  Fair d -> Stream.forever (Stream.instant similarity_da_prop) d ->
+  ~(aligned (config_list config)) -> FirstMove gatherW d config.
+Proof using .
+intros Hfair Hsim Nalign.
+destruct (one_must_move config Nalign) as [id Hmove].
+destruct Hfair as [locallyfair Hfair].
+specialize (locallyfair id).
+revert config Nalign Hmove.
+induction locallyfair as [d Hnow | d] ; intros config Nalign Hmove.
+* apply MoveNow. apply Hmove in Hnow.
+  + assert (get_location (round gatherW (Stream.hd d) config id) =/= get_location (config id)) as Hnow'.
+    { cbn -[complement equiv].
+      assumption. }
+    rewrite <-(moving_spec gatherW (Stream.hd d) config id) in Hnow'.
+    intros Habs. now rewrite Habs in Hnow'.   
+  + apply Hsim.
+* destruct (moving gatherW (Stream.hd d) config) as [| id' mov] eqn:Hmoving.
+  + apply MoveLater ; trivial.
+    apply IHlocallyfair.
+    - apply Hfair.
+    - apply Hsim.
+    - apply no_moving_same_config in Hmoving. now rewrite Hmoving.
+    - intros da Hda Hactive. apply no_moving_same_config in Hmoving.
+      rewrite (Hmoving id).
+      apply (round_compat (reflexivity gatherW) (reflexivity da)) in Hmoving. 
+      rewrite (Hmoving id).
+      now apply Hmove.
+  + apply MoveNow. rewrite Hmoving. discriminate.
+Qed.
+
+(* The proof is essentially a well-founded induction on [measure config].
+ * Fairness ensures that the measure must decrease at some point. *)
+Theorem weber_correct (d : demon) config : 
+  Fair d -> 
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  eventually_aligned config d gatherW.
+Proof using lt_0n.
+remember (measure config) as k. 
+generalize dependent d. generalize dependent config.
+pattern k. apply (well_founded_ind lt_wf). clear k.
+intros k IHk config Hk d Hfair Hsim.
+destruct (aligned_dec (config_list config)) as [align | Nalign] ;
+  [now apply Stream.Now, aligned_over|].
+induction (fair_first_move config Hfair Hsim Nalign) as [d config Hmove | d config Hmove FM IH_FM] ;
+  destruct Hsim as [Hsim_hd Hsim_tl] ; cbn in Hsim_hd ; apply Stream.Later.
++ destruct (round_decreases_measure config Hsim_hd Hmove) as [Ralign | Rmeasure].
+  - now apply Stream.Now, aligned_over.
+  - eapply IHk. 
+    * rewrite Hk. exact Rmeasure.  
+    * reflexivity.
+    * destruct Hfair as [_ Hfair_tl]. exact Hfair_tl.
+    * exact Hsim_tl.
++ apply no_moving_same_config in Hmove.
+  apply IH_FM ; (try rewrite Hmove) ; auto.
+  destruct Hfair as [_ Hfair_tl]. exact Hfair_tl.
+Qed.
+
+End Alignment.
diff --git a/CaseStudies/Gathering/InR2/Weber/Gather_flex_async.v b/CaseStudies/Gathering/InR2/Weber/Gather_flex_async.v
new file mode 100644
index 0000000000000000000000000000000000000000..a4a946ba707d32c85d15f484c5d7ec6922d36798
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Gather_flex_async.v
@@ -0,0 +1,5264 @@
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file implements an algorithm to GATHER all robots in the plane (R²). 
+ * The algorithm assumes there are no byzantine robots,
+ * and works in a FLEXIBLE and ASYNCHRONOUS setting. 
+
+ * The algorithm is as follows : all robots go towards the 'weber point' of 
+ * the configuration.
+ * The algorithm works on initial configurations where the weber point is unique.
+ * Thanks to a property of the weber point (see the thesis of Zohir Bouzid, 
+ * corollary 3.1.1), it remains unique and at the same place throughout the whole execution. *)
+(**************************************************************************)
+
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+Require Import Pactole.Util.ListComplements.
+(* Helping typeclass resolution avoid infinite loops. *)
+(* This does not seem to help much *)
+(* Remove Hints FMapFacts.eq_key_elt_Setoid FMapFacts.eq_key_Setoid FMapFacts.eq_key_elt_Equivalence FMapFacts.eq_key_Equivalence FMapFacts.Equal_ST FMapFacts.eq_key_EqDec: typeclass_instances. *)
+(* each of these sometimes helps, but sometimes makes things worse
+   too. *)
+(* Typeclasses eauto := debug (dfs) 7 . *)
+(* Typeclasses eauto := (bfs). *)
+
+(* Pactole basic definitions *)
+Require Export Pactole.Setting.
+(* Specific to R^2 topology *)
+Require Import Pactole.Spaces.RealMetricSpace.
+Require Import Pactole.Spaces.R2.
+Require Export Pactole.Util.SetoidDefs.
+(* Specific to gathering *)
+Require Pactole.CaseStudies.Gathering.WithMultiplicity.
+Require Import Pactole.CaseStudies.Gathering.Definitions.
+(* Specific to multiplicity *)
+Require Import Pactole.Observations.MultisetObservation.
+(* Specific to flexibility *)
+Require Import Pactole.Models.Flexible.
+(* Specific to settings with no Byzantine robots *)
+Require Import Pactole.Models.NoByzantine.
+(* Specific to definition and properties of the weber point. *)
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Weber_point.
+
+(* User defined *)
+Import Permutation.
+Import Datatypes.
+Import ListNotations.
+(*
+Section LibHyps_stuff.
+Local Open Scope autonaming_scope.
+
+(* Define the naming scheme as new tactic pattern matching on a type
+th, and the depth n of the recursive naming analysis. Here we state
+that a type starting with Nat.eqb should start with _Neqb, followed by
+the name of both arguments. #n here means normal decrement of depth.
+(S n) would increase depth by 1 (n-1) would decrease depth. *)
+Ltac rename_hyp_2 n th :=
+  match th with
+  | multiplicity ?x ?y => name(`_mult` ++ x#n ++ y#n)
+  end.
+
+(* Then overwrite the customization hook of the naming tactic *)
+Ltac rename_hyp ::= rename_hyp_2.
+
+End LibHyps_stuff.*)
+
+(* Preventing coq bug https://github.com/coq/coq/issues/19861 from
+biting us at "Print Assumptions". This takes precedence over the
+"abstract lia" hints.
+
+To check hints: 
+Print Hint *.
+
+ *)
+#[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.
+
+
+Set Implicit Arguments.
+Close Scope R_scope.
+Close Scope VectorSpace_scope.
+
+
+Section Gathering.
+Local Existing Instances dist_sum_compat R2_VS R2_ES.
+
+(* We assume the existence of a function that calculates a weber point of a collection
+ * (even when the weber point is not unique).
+ * This is a very strong assumption : such a function may not exist in closed form, 
+ * and the Weber point can only be approximated. *)
+(* Axiom weber_calc : list R2 -> R2. *)
+(* Axiom weber_calc_correct : forall ps, Weber ps (weber_calc ps). *)
+ (* We also suppose this function doesn't depend on the order of the points. 
+ * This is probably not necessary (we can show that it holds when the points aren't colinear) 
+ * but simplifies the proof a bit. *)
+(* Axiom weber_calc_compat : Proper (PermutationA equiv ==> equiv) weber_calc. *)
+(* Local Existing Instance weber_calc_compat. *)
+
+(* We then build a function that calculates the weber segment of a set of points. *)
+Axiom weber_calc_seg : list R2 -> R2 * R2. 
+
+Axiom weber_calc_seg_correct : forall ps,
+    ps <> [] -> forall w,
+    let '(w1, w2) := weber_calc_seg ps in 
+    Weber ps w <-> segment w1 w2 w.
+Lemma weber_calc_seg_correct_1 : forall ps w1 w2,
+    ps <> [] ->
+    weber_calc_seg ps = (w1, w2) -> 
+    forall w, Weber ps w <-> segment w1 w2 w.
+Proof. 
+  intros ps w1 w2 H H0 w. 
+  specialize (weber_calc_seg_correct H ) as h.
+  rewrite H0 in h.
+  apply h.
+Qed.
+
+Axiom weber_calc_seg_compat : Proper (PermutationA equiv ==> perm_pairA equiv) weber_calc_seg.
+
+Local Existing Instance weber_calc_seg_compat.
+
+
+(* THIS IS FALSE BECAUSE OF THE DEGENERATED CASE OF PS = [].
+(* The weber segment isn't reduced to a point iff 
+ * there are two distinct weber points in the original set of points. *)
+Lemma weber_calc_seg_diff ps : 
+  (let (w1, w2) := weber_calc_seg ps in w1 =/= w2) <-> 
+  exists p p', List.In p ps /\ Weber ps p /\ List.In p' ps /\ Weber ps p' /\ p =/= p'.
+Proof. Admitted.
+*)
+
+(* The number of robots *)
+Variables n : nat.
+Hypothesis lt_0n : 0 < n.
+
+(* There are no byzantine robots. *)
+Local Instance N : Names := Robots n 0.
+Local Instance NoByz : NoByzantine.
+Proof using . now split. Qed.
+
+Lemma list_in_length_n0 {A : Type} x (l : list A) : List.In x l -> length l <> 0.
+Proof using . intros Hin. induction l as [|y l IH] ; cbn ; auto. Qed.
+
+Lemma byz_impl_false : B -> False.
+Proof using . 
+intros b. assert (Hbyz := In_Bnames b). 
+apply list_in_length_n0 in Hbyz. 
+rewrite Bnames_length in Hbyz.
+cbn in Hbyz. intuition.
+Qed.
+
+(* Use this tactic to solve any goal
+ * provided there is a byzantine robot as a hypothesis. *)
+Ltac byz_exfalso :=
+  match goal with 
+  | b : ?B |- _ => exfalso ; apply (byz_impl_false b)
+  end.
+
+(* Since all robots are good robots, we can define a function
+ * from identifiers to good identifiers. *)
+Definition unpack_good (id : ident) : G :=
+  match id with 
+  | Good g => g 
+  | Byz _ => ltac:(byz_exfalso)
+  end.
+
+Lemma good_unpack_good id : Good (unpack_good id) == id.
+Proof using . unfold unpack_good. destruct_match ; [auto | byz_exfalso]. Qed.
+
+Lemma unpack_good_good g : unpack_good (Good g) = g.
+Proof using . reflexivity. Qed.  
+
+(* The robots are in the plane (R^2). *)
+Local Instance Loc : Location := make_Location R2.
+Local Instance LocVS : RealVectorSpace location := R2_VS.
+Local Instance LocES : EuclideanSpace location := R2_ES.
+
+(* - This is what represents a robot's state.
+ * The first location is the robot's start position (where it performed its last 'compute').
+ * The second location is the robot's destination (what the robogram computed).
+ * The ratio indicates how far the robot has moved along the straight path
+ * from start to destination.
+ * - The robogram doesn't have access to all of this information : 
+ * when we create an observation, this state gets reduced to 
+ * only the current position of the robot.
+ * - I would have prefered to use a path instead of a (start, destination) pair,
+ * but we need an EqDec instance on [info]. *)
+Definition info := ((location * location) * ratio)%type.
+
+Local Instance info_Setoid : Setoid info := 
+  prod_Setoid (prod_Setoid location_Setoid location_Setoid) ratio_Setoid.
+Local Instance info_EqDec : EqDec info_Setoid := 
+  prod_EqDec (prod_EqDec location_EqDec location_EqDec) ratio_EqDec.
+
+Local Instance St : State info.
+simple refine {|
+  get_location := fun '(start, dest, r) => straight_path start dest r ; 
+  state_Setoid := info_Setoid ;
+  state_EqDec := info_EqDec ;
+  precondition f := sigT (fun sim : similarity location => Bijection.section sim == f) ; 
+  lift f := fun '(start, dest, r) => ((projT1 f) start, (projT1 f) dest, r) 
+|} ; autoclass.
+Proof using .
++ abstract (intros [f [sim Hf]] [[start dest] r] ; cbn -[equiv straight_path] ;
+            rewrite <-Hf ; apply straight_path_similarity).
++ abstract (intros [[s d] r] [[s' d'] r'] [[Hs Hd] Hr] ; 
+            cbn -[equiv location] in * |- ; now rewrite Hs, Hd, Hr).
++ abstract (intros [f Hf] [g Hg] ; cbn -[equiv] ; intros Hfg [[s d] r] [[s' d'] r'] [[Hs Hd] Hr] ;
+            cbn -[equiv location] in * |- ; repeat split ; cbn -[equiv] ; auto).
+Defined.
+
+Definition get_start (i : info) := let '(s, _, _) := i in s.
+
+Local Instance get_start_compat : Proper (equiv ==> equiv) get_start.
+Proof using . intros [[? ?] ? ] [[? ?] ?] [[H _] _]. cbn -[equiv] in *. now rewrite H. Qed.
+
+Definition get_dest (i : info) := let '(_, d, _) := i in d.
+
+Local Instance get_dest_compat : Proper (equiv ==> equiv) get_dest.
+Proof using . intros [[? ?] ? ] [[? ?] ?] [[_ H] _]. cbn -[equiv] in *. now rewrite H. Qed.
+
+Lemma get_location_ratio_0 x y : get_location (x, y, ratio_0) == x.
+Proof using . cbn -[straight_path]. now rewrite straight_path_0. Qed. 
+
+(* Refolding typeclass instances *)
+Ltac foldR2 :=
+  change R2 with location in * ;
+  change R2_Setoid with location_Setoid in * ;
+  change R2_EqDec with location_EqDec in * ;
+  change R2_VS with LocVS in * ;
+  change R2_ES with LocES in * ;
+  change info_Setoid with state_Setoid in * ;
+  change info_EqDec with state_EqDec in *.
+
+(* Robots choose their destination.
+ * They will move to this destination along a straight path. *)
+Local Instance RobotC : robot_choice location := 
+  { robot_choice_Setoid := location_Setoid }.
+
+(* Robots view the other robots' positions up to a similarity. *)
+Local Instance FrameC : frame_choice (similarity location) := FrameChoiceSimilarity.
+(* The demon doesn't perform any other choice for activated robots. *)
+Local Instance UpdateC : update_choice unit := NoChoice.
+(* The demon chooses how far to move inactive robots towards their destination. 
+ * The ratio chosen by the demon is ADDED to the ratio stored in the robot state
+ * (the result is clamped at 1 of course). *)
+Local Instance InactiveC : inactive_choice ratio := { inactive_choice_Setoid := ratio_Setoid }.
+
+(* In a flexible setting, delta is the minimum distance that robots
+ * are allowed to move before being reactivated. *)
+Variables delta : R.
+Hypothesis delta_g0 : (0 < delta)%R.
+
+(* This is the property that must be verified in a flexible setting. *)
+Definition flex_da_prop da := 
+  forall id (config : configuration), activate da id = true -> 
+    get_location (config id) == get_dest (config id) \/ 
+    (delta <= dist (get_start (config id)) (get_location (config id)))%R.
+
+(* When a robot perfoms a look-compute, it is reset to ration 0. *)
+Local Instance UpdateF : update_function location (similarity location) unit.
+simple refine {| 
+  update config g _ target _ := (get_location (config (Good g)), target, ratio_0)
+|} ; autoclass.
+Proof using .
+intros c c' Hc g g' Hg _ _ _ t t' Ht _ _ _.
+assert (H : c (Good g) == c' (Good g')) by now rewrite Hg, Hc.
+destruct (c (Good g)) as [[start dest] r].
+destruct (c' (Good g')) as [[start' dest'] r'].
+destruct H as [[Hstart Hdest] Hr]. cbn -[equiv] in Hstart, Hdest, Hr.
+repeat split ; cbn -[equiv get_location] ; auto.
+foldR2. apply get_location_compat. now repeat split ; cbn -[equiv].
+Defined.
+
+
+Local Instance InactiveF : inactive_function ratio.
+simple refine {| inactive config id r_demon := 
+  let '(start, dest, r) := config id in (start, dest, add_ratio r r_demon) 
+|} ; autoclass.
+Proof using . 
+intros c c' Hc i i' Hi rd rd' Hrd.
+assert (H : c i == c' i') by now rewrite Hi, Hc.
+destruct (c i) as [[start dest] r].
+destruct (c' i') as [[start' dest'] r'].
+destruct H as [[Hstart Hdest] Hr]. cbn -[equiv] in Hstart, Hdest, Hr.
+repeat split ; cbn -[equiv] ; auto.
+f_equiv ; auto.
+Defined.
+
+Notation "!! config" := (@obs_from_config info Loc St N multiset_observation config (origin, origin, ratio_0)) (at level 10, only parsing).
+
+(* This is a shorthand for the list of positions of robots in a configuration. *)
+Definition pos_list (config : configuration) : list location := 
+  List.map get_location (config_list config).
+
+Lemma pos_list_length config : length (pos_list config) = n.
+Proof using . 
+unfold pos_list. rewrite length_map, config_list_length.
+cbn. now rewrite Nat.add_0_r.
+Qed.
+
+Lemma pos_list_non_nil config : (pos_list config) <> [].
+Proof using lt_0n. 
+  intro abs.
+  apply length_zero_iff_nil in abs.
+  rewrite pos_list_length in abs.
+  lia.
+Qed.
+
+
+Lemma weber_calc_seg_correct_2 : forall c w1 w2,
+    weber_calc_seg (pos_list c) = (w1, w2) -> 
+    forall w, Weber (pos_list c) w <-> segment w1 w2 w.
+Proof using lt_0n.
+  intros c w1 w2 H w. 
+  apply weber_calc_seg_correct_1;auto.
+  apply pos_list_non_nil.
+Qed.
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine x x' (c c':configuration) : 
+  InA equiv (x, x') (combine (config_list c) (config_list c')) <-> 
+    exists id, (x == c id /\ x' == c' id)%VS.
+Proof using lt_0n.
+  pose (id0 :=(Fin.Fin lt_0n)).
+  pose (g0 :=(Good id0)).
+  pose (default :=(c g0)).
+  pose (default' :=(c' g0)).
+  split.
+  + intros Hin.
+    specialize @config_list_InA_combine_same_lgth2 with (1:=Hin)(dfA:=default)(dfB:=default') as h.
+    destruct h as [i [Hn [Hx Hx']]].
+    rewrite config_list_spec in Hx,Hx'.
+    unfold default in Hx,Hx'.
+    unfold default' in Hx'.
+    rewrite map_nth in Hx .
+    rewrite map_nth in Hx' .
+    exists (nth i names g0);auto.
+  + intros [[g|b] [Hx Hx']] ; [|byz_exfalso]. 
+    rewrite Hx,Hx'.
+    rewrite config_list_spec.
+    rewrite config_list_spec.
+    rewrite combine_map.
+    apply InA_alt.
+    exists (c (Good g),c' (Good g)).
+    split;auto.
+    apply in_map_iff.
+    exists (Good g,Good g);split;auto.
+    apply in_combine_id;split;auto.
+    unfold names.
+    apply in_app_iff.
+    left.
+    apply in_map_iff.
+    exists g;split;auto.
+    apply In_Gnames.
+Qed.
+
+Lemma pos_list_InA_combine x x' c c' : 
+  InA equiv (x, x') (combine (pos_list c) (pos_list c')) <-> 
+  exists id, x == get_location (c id) /\ x' == get_location (c' id).
+Proof using lt_0n.
+unfold pos_list. rewrite combine_map. rewrite (@InA_map_iff _ _ equiv equiv) ; autoclass.
++ split.
+  - intros [[y y'] [[Hx Hx'] Hin]]. cbn -[equiv get_location] in Hx, Hx'.
+    rewrite config_list_InA_combine in Hin. destruct Hin as [id [Hy Hy']].
+    exists id. rewrite <-Hy, <-Hy', Hx, Hx'. auto.
+  - intros [id [Hx Hx']].
+    exists (c id, c' id). rewrite <-Hx, Hx'. split ; [auto|].
+    rewrite config_list_InA_combine. exists id. auto.
++ intros [? ?] [? ?] [H1 H2]. cbn -[equiv] in H1, H2. split ; cbn -[equiv get_location].
+  - now rewrite H1.
+  - now rewrite H2.
+Qed. 
+
+Lemma pos_list_InA x c : 
+  InA equiv x (pos_list c) <-> exists id, x == get_location (c id).
+Proof using . 
+unfold pos_list. rewrite (@InA_map_iff _ _ equiv equiv) ; autoclass.
++ split.
+  - intros [y [Hx Hin]]. foldR2. rewrite config_list_InA in Hin.
+    destruct Hin as [id Hy]. exists id. now rewrite <-Hx, <-Hy.   
+  - intros [id Hx]. exists (c id). rewrite <-Hx. split ; [auto|].
+    foldR2. rewrite config_list_InA. exists id. auto. 
++ foldR2. apply get_location_compat.
+Qed. 
+
+Lemma n_nGnB: n = nG + nB.
+Proof using Type.
+  rewrite <- names_length.
+  assert (config:configuration).
+  { red.
+    intro.
+    constructor.
+    - exact (0%VS,0%VS).
+    - exists 0%R;lra. }
+  rewrite <- (pos_list_length config).
+  unfold pos_list.
+  rewrite length_map.
+  unfold config_list, names.
+  rewrite app_length.
+  rewrite app_length.
+  rewrite length_map.
+  rewrite length_map.
+  unfold Gpos, Bpos.
+  rewrite length_map.
+  rewrite length_map.
+  reflexivity.
+Qed.
+
+
+Lemma n_nG: n = nG.
+Proof using n.
+  rewrite n_nGnB.
+  rewrite nB_eq_0.
+  apply Nat.add_0_r.
+Qed.
+
+
+(* I can't use the [invalid] definition from the file WithMultiplicity
+ * because it relies on a state that contains only a location.
+ * Notice that when [n] is even, [n/2] is the same as [(n+1)/2].
+ * When [n] is odd, the weber point is unique so we don't care. *)
+Definition invalid (config : configuration) :=
+  exists x y, 
+    x =/= y /\ 
+    (!! config)[x] = Nat.div2 (n+1) /\ 
+    (!! config)[y] = Nat.div2 (n+1).
+
+Global Instance invalid_compat : Proper (equiv ==> iff) invalid.
+Proof using .
+intros c1 c2 Hc. unfold invalid. setoid_rewrite Hc. reflexivity.
+Qed.
+
+
+
+(*
+Require Pactole.CaseStudies.Gathering.Impossibility.
+
+(* Lemma build_similarity : forall [pt1 pt2 pt3 pt4 : R2], pt1 =/= pt2 -> pt3 =/= pt4 -> similarity R2. *)
+
+
+Lemma invalid_equivalent:
+  forall c, invalid c <-> WithMultiplicity.invalid c.
+
+Proof.
+  
+Qed.
+
+Definition gather_impossible:
+  Nat.Even n ->
+  forall (r : robogram) (config : configuration),
+    WithMultiplicity.invalid config ->
+    exists d : demon,
+      Fair d /\ ~ WillGather (execute r d config) :=
+  fun heven => (@Pactole.CaseStudies.Gathering.Impossibility.noGathering_Fair
+                  n heven (not_eq_sym (Nat.lt_neq _ _ lt_0n)) _ _ _
+                  build_similarity build_similarity_compat build_similarity_eq1
+                  build_similarity_eq2 build_similarity_inverse).
+*)
+
+(*
+Ltac rename_depth ::= constr:(2).
+*)
+#[local] Definition info0: info := (0%VS, 0%VS, ratio_0).
+(* I could have included this in the definition of invalid.
+ * I prefered to use [(n+1)/2] instead of [n/2] and thus have
+ * [Even n] as a consequence. *)
+Lemma invalid_even config : 
+  invalid config -> Nat.Even n.
+Proof using Type.
+  clear delta delta_g0.
+  unfold invalid.
+  change (0%VS, 0%VS, ratio_0) with info0.
+  intros h.
+  destruct h as [ x [ y [h_complement_equiv_x_y_ [ h_eq_multiplicity_div2_ h_eq_multiplicity_div2_0_ ]]] ].
+  destruct (Nat.Even_Odd_dec n) as [h_Even_n_  | h_Odd_n_ ].
+  - assumption.
+  - exfalso.
+    (* specialize Nat.Odd_div2 with (1:= o) as hodd. *)
+    rewrite Nat.add_1_r in *.
+    rewrite <- Nat.Odd_div2 in *;eauto.
+    
+    specialize (cardinal_obs_from_config config info0) as hcard.
+    rewrite <- n_nGnB in hcard.
+    remember (obs_from_config config info0) as cf.
+    assert (hneq2:y =/= x) by now symmetry.
+    assert (hmult: FMultisetsOn location
+                     (MMultisetWMap.FMOps_WMap location location_Setoid location_EqDec FMapList.MapList))
+      by typeclasses eauto.
+    specialize @cardinal_lower with (m := (remove_all x cf)) (x:=y) (1:=hmult) as h_cardlower.
+    clear hneq2 hmult.
+
+    rewrite  @remove_all_cardinal in h_cardlower; [|typeclasses eauto].
+    rewrite @remove_all_other in h_cardlower; try typeclasses eauto.
+    (*2:{ now symmetry in h1. }*) 2:{ now symmetry in h_complement_equiv_x_y_. }
+    (* rewrite h2,h3,hcard in h_5. *)
+    rewrite h_eq_multiplicity_div2_ , h_eq_multiplicity_div2_0_ , hcard in h_cardlower.
+    rewrite Nat.Odd_double with (1:=h_Odd_n_) in h_cardlower at 2.
+    unfold Nat.double in h_cardlower.
+    rewrite Nat.sub_succ in h_cardlower.
+    rewrite Nat.add_sub in h_cardlower.
+    eapply Nat.nle_succ_diag_l;eauto.
+Qed.
+
+
+(* The support of a multiset, but elements are repeated 
+ * a number of times equal to their multiplicity. 
+ * This is needed to convert an observation from multiset to list format, 
+ * so that we can use functions such as [weber_calc]. *)
+Definition multi_support {A} `{EqDec A} (s : multiset A) :=
+  List.flat_map (fun '(x, mx) => alls x mx) (elements s).
+
+Local Instance multi_support_compat {A} `{EqDec A} : Proper (equiv ==> PermutationA equiv) (@multi_support A _ _).
+Proof using . 
+intros s s' Hss'. unfold multi_support. apply flat_map_compat_perm with eq_pair ; autoclass ; [|].
++ intros [x mx] [y my] Hxy. inv Hxy. simpl in H0, H1. now rewrite H0, H1.
++ now apply elements_compat.
+Qed.
+(*
+Lemma multi_support_equiv_pos_list {A} `{EqDec A}:
+  forall (c: configuration) (o:@observation _ _ _ _ multiset_observation),
+    obs_from_config c (0%VS, 0%VS, ratio_0) = o
+    -> PermutationA equiv (multi_support o) (pos_list c).
+Proof.
+  clear lt_0n delta_g0.
+  intros/sng.
+  unfold 
+  
+Qed.
+*)
+(* Choose a robot with multiplicity [k] in the list [ps] *)
+Definition robot_with_mult k (ps : list location) : option location :=
+  find (fun x => countA_occ equiv R2_EqDec x ps ==b k) ps.
+
+Local Instance find_compat A {SA:Setoid A} : 
+  Proper ((equiv ==> eq) ==> eqlistA equiv ==> equiv) (@find A).
+Proof using Type.
+  intros a b h l.
+  induction l;intros.
+  - inversion H;subst.
+    reflexivity.
+  - inversion H;subst.
+    cbn.
+    destruct (a a0) eqn: heqaa0.
+    + cbn.
+      compute in h.
+      rewrite <- (h a0 x');auto.
+      rewrite heqaa0;auto.
+    + cbn.
+      compute in h.
+      rewrite <- (h a0 x');auto.
+      rewrite heqaa0;auto.
+      eapply IHl;eauto.
+Qed.
+
+Local Instance robot_with_mult_compat : 
+  Proper (eq ==> eqlistA equiv ==> equiv) robot_with_mult.
+Proof using Type.
+  red.
+  intros a b h l l' h'.
+  subst.
+  unfold robot_with_mult.
+  rewrite h'.
+  apply find_compat.
+  all:swap 1 2.
+  - reflexivity.
+  - unfold respectful.
+    intros x y H.
+    rewrite  h'.
+    rewrite H.
+    reflexivity.
+Qed.
+
+Lemma find_map {A B} (f : B -> bool) (g : A -> B) (l : list A) : 
+  find f (List.map g l) = option_map g (find (fun x => f (g x)) l).
+Proof using Type.
+  induction l.
+  - cbn.
+    reflexivity.
+  - cbn.
+    destruct (f (g a)).
+    + cbn.
+      reflexivity.
+    + assumption.
+Qed.
+
+Lemma robot_with_mult_map k ps f :
+  Proper (equiv ==> equiv) f ->
+  Preliminary.injective equiv equiv f ->
+  option_map f (robot_with_mult k ps) == robot_with_mult k (List.map f ps).
+Proof using lt_0n delta_g0. 
+intros Hcompat Hinj. induction ps as [|x ps IH] ; auto. 
+unfold robot_with_mult in *. cbn -[equiv countA_occ].
+case ifP_bool ; case ifP_bool ; cbn [countA_occ] ;
+  rewrite (countA_occ_map R2_EqDec R2_EqDec Hcompat Hinj) ;
+  case ifP_sumbool ; (try now intuition) ; intros _ ;
+  case ifP_sumbool ; (try now intuition) ; intros _.
++ intros ->. discriminate.
++ intros ->. discriminate.
++ intros _ _. rewrite find_map. 
+  assert (Hcompat2 : Proper (equiv ==> equiv) (option_map f)).
+  { intros [?|] [?|] H ; try intuition. now inv H. }
+  apply Hcompat2. apply find_compat ; try reflexivity. 
+  intros ? ? H. rewrite H. case ifP_sumbool ; case ifP_sumbool ; intuition ;
+    erewrite countA_occ_map ; autoclass.
+Qed.
+
+Lemma list_perm_find_none: 
+  forall [A : Type] {SA : Setoid A} (f : A -> bool) (l l': list A),
+    PermutationA eq l l' -> 
+    Proper (equiv ==> equiv) f -> 
+    find f l = None -> 
+    find f l' = None.
+Proof using Type.
+  intros A SA f l l' H HProper H0. 
+  specialize find_none with (1:=H0) as H0'.
+  destruct (find f l') eqn:heq_find.
+  - exfalso.
+    specialize find_some with (1:=heq_find) as [x heq_find'].
+    rewrite H0' in heq_find';try discriminate.
+    eapply In_perm_compat. 
+    + reflexivity.
+    + apply PermutationA_Leibniz.
+      eassumption.
+    + assumption.
+  - reflexivity.
+Qed.
+
+
+
+Lemma obs_from_config_spec: forall config pt,
+    (!! config)[pt] = countA_occ _ equiv_dec pt (pos_list config).
+Proof using Type.
+  intros config0 pt. 
+  specialize obs_from_config_spec as h.
+  cbn -[get_location config_list] in h.
+  apply  h.
+Qed.
+
+(* When the position isn't invalid, there is at most one position
+ * with multiplicity [(n+1) / 2]. Thus [robot_with_mult ((n+1)/2)] has 
+ * to return the same result no matter the order of points. *)
+Lemma robot_with_mult_unique (config : configuration) ps : 
+  ~invalid config ->
+  PermutationA eq (pos_list config) ps -> 
+    robot_with_mult (Nat.div2 (n+1)) (pos_list config) == robot_with_mult (Nat.div2 (n+1)) ps.
+Proof using Type.
+  intros Hvalid Hperm.
+  unfold robot_with_mult.
+  assert (find (fun x : R2 => countA_occ equiv R2_EqDec x (pos_list config) ==b Nat.div2 (n + 1))
+    (pos_list config) == find (fun x : R2 => countA_occ equiv R2_EqDec x ps ==b Nat.div2 (n + 1))
+    (pos_list config)).
+  { apply find_compat.
+    - intros x y H.
+      rewrite H.
+      foldR2.
+      specialize countA_occ_compat with (eqA:=@eq (@location Loc)) as h.
+      specialize (h _).
+      unfold Proper, respectful in h.
+      foldR2.
+      specialize h with (x:=y)( y:=y) (2:=Hperm).
+      rewrite h;auto.
+    - reflexivity. }
+  rewrite H.
+  clear H.
+  remember (fun x : R2 => countA_occ equiv R2_EqDec x ps ==b Nat.div2 (n + 1)) as f.
+  match goal with
+    |- ?A == ?B => destruct A eqn:HeqA; destruct B eqn:HeqB ;try discriminate
+  end.
+  all:swap 1 4.
+  - reflexivity.
+  - erewrite list_perm_find_none in HeqA;try typeclasses eauto;eauto.
+    + symmetry. assumption.
+    + rewrite Heqf.
+      now intros i j ->.
+  - erewrite list_perm_find_none in HeqB;try typeclasses eauto;eauto.
+    rewrite Heqf.
+    now intros i j ->.
+  - cbn.
+    destruct (R2_EqDec r r0).
+    + auto.
+    + exfalso.
+      apply Hvalid.
+      unfold invalid.
+      exists r, r0.
+      specialize find_some with (1:=HeqA) as [hinr hfr]. 
+      specialize find_some with (1:=HeqB) as [hinr0 hfr0].
+      repeat rewrite obs_from_config_spec.
+      rewrite Heqf in hfr,hfr0.
+      rewrite eqb_true_iff in hfr,hfr0.
+      repeat constructor;auto.
+      * rewrite <- hfr.
+        apply countA_occ_compat_setoid;auto.
+      * rewrite <- hfr0.
+        apply countA_occ_compat_setoid;auto.
+Qed.
+
+
+Definition invalid_obs (o:observation) :=
+  exists x y,
+    x =/= y /\ cardinal o <> 0 /\
+      o[x] = Nat.div2 (cardinal o + 1) /\
+      o[y] = Nat.div2 (cardinal o + 1).
+
+
+Local Instance invalid_obs_compat: Proper (equiv ==> iff) invalid_obs.
+Proof using Type.
+  unfold Proper, respectful.
+  intros x y h. 
+  split;intro h'.
+  - unfold invalid_obs in h'.
+    decompose [ex and] h'. clear h'.
+    red.
+    exists x0, x1.
+    rewrite <- h.
+    intuition;auto.
+  - unfold invalid_obs in h'.
+    decompose [ex and] h'. clear h'.
+    red.
+    exists x0, x1.
+    rewrite h.
+    intuition;auto.
+Qed.
+
+
+Lemma n_cardinal: forall (c:configuration) (i:info),
+   (@cardinal (@location Loc) (@location_Setoid Loc) 
+      (@location_EqDec Loc)
+      (MMultisetWMap.FMOps_WMap (@location Loc) (@location_Setoid Loc)
+         (@location_EqDec Loc)
+         (@FMapList.MapList (@location Loc) (@location_Setoid Loc)
+            (@location_EqDec Loc))) (obs_from_config c i)) = n.
+Proof using Type.
+  intros c i.
+  rewrite n_nGnB.
+  apply cardinal_obs_from_config.
+Qed.
+
+Lemma invalid_config_obs: forall c info, 
+    invalid c <-> invalid_obs ((obs_from_config c info)).
+Proof using lt_0n.
+  intros c i.
+  split;[intros h | intros h].
+  - red in h .
+    decompose [ex and] h. clear h.
+    red.
+    rewrite n_cardinal.
+    exists x , x0.
+    repeat split;intuition.
+  - red.
+    (* specialize (h (0%VS, 0%VS, ratio_0)). *)
+    red in h.
+    decompose [ex and] h. clear h.
+    rewrite n_cardinal in H1, H3.
+    exists x , x0;intuition.
+Qed.
+
+Lemma fold_left_st_snd: forall T `{FMS: FMapInterface.FMapSpecs T} l o n,
+    (forall xn : T * nat, InA eq_pair xn l -> o[fst xn] = snd xn /\ snd xn > 0) -> 
+    List.fold_left (fun (acc : nat) (xn : T * nat) => snd xn + acc) l n =
+      List.fold_left (fun (acc : nat) (xn : T * nat) => o[fst xn] + acc) l n.
+Proof using Type.
+  intros until l. 
+  induction l.
+  + intros. cbn. reflexivity.
+  + intros o n0 h_prem.
+    cbn.
+    specialize h_prem with (xn:=a) as h.
+    assert (hinxx: InA eq_pair a (a :: l)).
+    { constructor.
+      reflexivity. }
+    specialize (h hinxx).
+    destruct h as [h11 h12].
+    rewrite <- h11.
+    apply IHl.
+    intros xn hinA. 
+    apply h_prem.
+    constructor 2.
+    assumption.
+Qed.
+
+
+
+Lemma cardinal_map_support:
+  forall T `{FMS: FMapInterface.FMapSpecs T} o,
+    cardinal o = List.fold_left (fun acc e => o[e] + acc) (support o) 0.
+Proof using.
+  clear lt_0n.
+  intros until o.
+  rewrite cardinal_spec.
+  rewrite fold_spec.
+  erewrite -> fold_left_st_snd with (o:=o).
+  all:swap 1 3.
+  { intros xn h. 
+    apply elements_spec.
+    assumption. }
+  { typeclasses eauto. }
+  transitivity
+    (fold_left (fun (acc : nat) xn => o[xn] + acc) (List.map fst (elements o)) 0).
+  - rewrite Preliminary.fold_left_map.
+    reflexivity.
+  - eapply @fold_left_symmetry_PermutationA with (eqA:=equiv);try typeclasses eauto.
+    + intros x y z. 
+      unfold respectful;intros u v huv.
+      rewrite z.
+      rewrite huv.
+      reflexivity.
+    + intros x y z.
+      lia.
+(*      rewrite Nat.add_assoc.
+      setoid_rewrite Nat.add_comm at 2.
+      rewrite Nat.add_assoc.
+      reflexivity.*)
+    + symmetry.
+      eapply support_map_elements.
+    + reflexivity.
+Qed.
+
+(*
+Definition cardinal_from_support T `{FMS: FMapInterface.FMapSpecs T} o l :=
+  List.fold_left (fun acc e => o[e] + acc) l 0.
+
+Instance cardinal_from_support_compat:
+  Proper (equiv ==> equiv ==> eq) (@cardinal_from_support _ _ _ _ _). 
+Proof.
+  unfold cardinal_from_support.
+  
+  unfold Proper, respectful.
+  intros x y H x0 y0 H0. 
+  intros x y H. 
+  
+  
+  rewrite <- cardinal_map_support;[|typeclasses eauto ].
+  rewrite <- cardinal_map_support;[|typeclasses eauto ].
+  rewrite H.
+  reflexivity.
+Qed.
+*)
+(*
+Lemma cardinal_map_support: forall o,
+    cardinal o = List.fold_left (fun acc e => o[e] + acc) (support o) 0.
+Proof using.
+  clear lt_0n.
+  intros o. 
+  rewrite cardinal_spec.
+  rewrite fold_spec.
+  erewrite -> toto with (o:=o).
+  2:{ intros xn H. 
+      apply elements_spec.
+      assumption. }
+  transitivity
+    (fold_left (fun (acc : nat) (xn : info) => o[xn] + acc) (List.map fst (elements o)) 0).
+  - rewrite Preliminary.fold_left_map.
+    reflexivity.
+  - eapply @fold_left_symmetry_PermutationA with (eqA:=equiv);try typeclasses eauto.
+    + intros x y z. 
+      unfold respectful;intros.
+      rewrite z.
+      rewrite H.
+      reflexivity.
+    + intros x y z.
+      lia.
+    + symmetry.
+      eapply support_map_elements.
+    + reflexivity.
+Qed.
+*)
+Lemma card_div2: forall {T} `{FMS: FMapInterface.FMapSpecs T} o,
+    cardinal o <> 0 -> 0 < Nat.div2 (cardinal o + 1).
+Proof using Type.
+  intros T S0 H F FMS o H0. 
+  apply Exp_prop.div2_not_R0.
+  clear lt_0n. (* otherwise it is used by lia *)
+  lia.
+(*  rewrite Nat.add_1_r.
+  apply -> Nat.succ_lt_mono.
+  now apply Nat.neq_0_lt_0.
+  (*lia. (* otherwise it is used by lia *)*)*)
+Qed.
+
+Lemma elements_subset: forall (o:multiset location) l l',
+    (* NoDupA equiv (l++l') *)
+    is_elements (l++l')
+    -> PermutationA equiv (elements o) (l ++ l')
+    -> Subset (from_elements l) o.
+Proof using Type.
+  intros o l l' hnodup hperm.
+  foldR2.
+  specialize @from_elements_append with (l1:=l) (l2:=l') as h_all_equiv.
+  specialize (h_all_equiv _ _ _ _).
+  specialize (@union_subset_l _ _ _ _ _ (from_elements l) (from_elements l')) as h_subunion.
+  transitivity (union (from_elements l) (from_elements l'));auto.
+  specialize (@from_elements_elements _ _ _ _ _ o) as h'''.
+  rewrite <- h'''.
+  rewrite <- h_all_equiv.  
+  (* rewrite hperm fails: problem unfiying equiv and eq_pair. *)
+  assert (hperm':PermutationA eq_pair (elements o) (l ++ l')).
+  { apply hperm. }
+  rewrite hperm'.
+  reflexivity.
+Qed.
+
+Lemma doubleton_subset: forall o (x y : location),
+    x =/= y
+    -> Subset (from_elements ((x,o[x]):: (y,o[y]) :: nil)) o.
+Proof using Type.
+  intros o x y H0. 
+  assert (filter (fun a:location => mem (location_EqDec) a (x :: (y :: nil))) o == (from_elements ((x,o[x]):: (y,o[y]) :: nil))).
+  { cbn [from_elements].
+    unfold equiv.
+    cbn -[filter add].
+    intros x1. 
+    rewrite filter_spec; try typeclasses eauto.
+    match goal with
+      |- (if ?X then _ else _) = _ => destruct X eqn:heqmem
+    end.
+    - rewrite mem_true_iff in heqmem.
+      inversion heqmem as [| _a _b hinA _c];subst.
+      + rewrite add_same,add_other;auto.
+      + inversion hinA as [| _a _b hinA' _c];subst.
+        * rewrite add_other,add_same;auto.
+        * inversion hinA'.
+    - rewrite mem_false_iff in heqmem.
+      rewrite add_other.
+      2:{ intro abs.
+          apply heqmem.
+          rewrite abs. constructor 1. reflexivity. }
+      rewrite add_other.
+      2:{ intro abs.
+          apply heqmem.
+          rewrite abs. constructor 2. constructor 1. reflexivity. }
+      reflexivity. }
+  rewrite <- H.
+  apply filter_subset.
+  typeclasses eauto.
+Qed.
+
+Lemma tripleton_subset: forall o (x y z : location),
+    x =/= y
+    -> x =/= z
+    -> y =/= z
+    -> Subset (from_elements ((x,o[x]):: (y,o[y]) :: (z,o[z]):: nil)) o.
+Proof using Type.
+  intros o x y z hxy hxz hyz. 
+  assert
+    (h_filter:
+      filter (fun a:location => mem (location_EqDec) a (x :: (y :: (z :: nil)))) o
+      == (from_elements ((x,o[x]):: (y,o[y]) ::  (z,o[z]) :: nil))).
+  { cbn [from_elements].
+    unfold equiv.
+    cbn -[filter add].
+    intros x1. 
+    rewrite filter_spec; try typeclasses eauto.
+    match goal with
+      |- (if ?X then _ else _) = _ => destruct X eqn:heqmem
+    end.
+    - rewrite mem_true_iff in heqmem.
+      inversion heqmem as [| _a _b hinA _c];subst.
+      + rewrite add_same.
+        do 2 (rewrite add_other;auto).
+      + inversion hinA as [| _a _b hinA' _c];subst.
+        * rewrite add_other;try now symmetry;auto.
+          rewrite add_same;try now symmetry;auto.
+          rewrite add_other;try now symmetry;auto.
+        * inversion hinA' as [| _a _b hinA'' _c];subst.
+          -- do 2 (rewrite add_other;try now symmetry;auto).
+             rewrite add_same;try now symmetry;auto.
+          -- inversion hinA''.
+    - rewrite mem_false_iff in heqmem.
+      rewrite add_other.
+      2:{ intro abs.
+          apply heqmem.
+          rewrite abs. constructor 1. reflexivity. }
+      rewrite add_other.
+      2:{ intro abs.
+          apply heqmem.
+          rewrite abs. constructor 2. constructor 1. reflexivity. }
+      rewrite add_other.
+      2:{ intro abs.
+          apply heqmem.
+          rewrite abs. constructor 2. constructor 2. constructor 1. reflexivity. }
+      reflexivity. }
+  rewrite <- h_filter.
+  apply filter_subset.
+  typeclasses eauto.
+Qed.
+
+(*
+Ltac rename_hyp_3 n th :=
+  match th with
+  | nil => name(`_nil`)
+  | complement equiv ?x ?y => name(`_nequiv` ++ x#n ++ y#n)
+  | complement ?r ?x ?y => name(`_not` ++ x#n ++ y#n)
+  | NoDupA ?equiv ?l => name(`_NodupA` ++ l#n)
+  | InA ?equiv ?x ?l => name(`_InA` ++ x#n ++ l#n)
+  | removeA ?equiv ?x ?l => name(`_remA` ++ x#n ++ l#n)
+  | support ?ms => name(`_support` ++ ms#n)
+  | cardinal ?ms => name(`_card` ++ ms#n)
+  | remove_all ?x ?ms => name(`_remAll` ++ x#n ++ ms#n)
+  | PermutationA ?equiv ?l1 ?l2 => name(`_PermutA` ++ l1#n ++ l2#n)
+  | (@eq _ ?x ?y) => name (`_eq` ++ x#n ++ y#n)
+  | _ => rename_hyp_2 n th
+  end.
+
+(* Then overwrite the customization hook of the naming tactic *)
+Ltac rename_hyp ::= rename_hyp_3.
+Ltac rename_depth ::= constr:(3).
+*)
+
+Lemma support_mult_pos:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall ms x,
+    InA equiv x (support ms) <-> ms[x]>0.
+Proof using Type.
+  clear lt_0n.
+  intros.
+  split.
+  { intros h_InA_x_support_ms_ .
+    apply support_spec in h_InA_x_support_ms_.
+    destruct (eq_0_lt_dec (ms[x])) as [h_eq_mult_x_ms_0_|?].
+    - apply not_In in h_eq_mult_x_ms_0_.
+      contradiction.
+    - lia. }
+  { intros h_gt_mult_x_ms_0_ .
+    destruct (InA_dec H x (support ms)) as [|h_not_InA_x_support_].
+    - assumption.
+    - exfalso.
+      rewrite support_spec in h_not_InA_x_support_.
+      apply not_In in h_not_InA_x_support_.
+      foldR2.
+      lia. }
+Qed.
+
+
+Lemma remove_all_cardinal_plus:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall (x : elt) (m : multiset elt), cardinal (remove_all x m) + m[x] = cardinal m.
+Proof using Type.
+  clear lt_0n delta_g0 delta n.
+  intros elt S0 H Ops h_FMultisetsOn_Ops_ x m .
+  assert (m[x] <= cardinal m).
+  { apply cardinal_lower. }
+  specialize remove_all_cardinal with (x:=x)(m:=m) as h.
+  lia.
+Qed.
+
+Lemma obs_invalid_doubleton1:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall o x y,
+    x =/= y
+    -> o[x] + o[y] >= (cardinal o)
+    -> forall z, z =/= x -> z =/= y -> o[z]=0.
+Proof using Type.
+  clear lt_0n.
+  intros elt S0 H Ops h_FMultisetsOn_Ops_ o x y h_nequiv_x_y_ h_ge_add_mult_mult_card_o_ z h_nequiv_z_x_ h_nequiv_z_y_ .
+  specialize remove_all_cardinal_plus with (x:=x)(m:=o).
+  intro h_eq_add_card_mult_card_o_ .
+  specialize remove_all_cardinal_plus with (x:=y)(m:=(remove_all x o)).
+  intro h_eq_add_card_mult_card_remAll_ .
+  foldR2.
+  rewrite  remove_all_other with (x:=x)(y:=y)(m:=o) in h_eq_add_card_mult_card_remAll_.
+  2:{ symmetry;assumption. }
+  rewrite <- h_eq_add_card_mult_card_remAll_ in h_eq_add_card_mult_card_o_.
+  assert (cardinal (remove_all y (remove_all x o)) = 0) as h_eq_card_remAll_0_.
+  { lia. }
+  assert (o[z] = (remove_all y (remove_all x o))[z]) as h_eq_mult_z_o_mult_z_remAll_.
+  { rewrite remove_all_other;auto.
+    rewrite remove_all_other;auto. }
+  rewrite h_eq_mult_z_o_mult_z_remAll_.
+  apply Nat.le_0_r.
+  rewrite <- h_eq_card_remAll_0_.
+  apply cardinal_lower.
+Qed.
+
+Lemma obs_invalid_doubleton2:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall o x y,
+    x =/= y
+    -> o[x] > 0
+    -> o[y] > 0
+    -> o[x] + o[y] >= (cardinal o)
+    -> PermutationA equiv (support o) (x::y::nil).
+Proof using Type.
+  clear lt_0n delta_g0 delta n.
+  intros elt S0 H Ops h_FMultisetsOn_Ops_ o x y h_nequiv_x_y_ h_gt_mult_x_o_0_ h_gt_mult_y_o_0_ h_ge_add_mult_mult_card_o_ .
+  specialize (obs_invalid_doubleton1 _ h_nequiv_x_y_ h_ge_add_mult_mult_card_o_).
+  intro h_all_eq_mult_0_ .
+
+  specialize support_mult_pos with (x:=x)(ms:=o).
+  intros [h_impl_gt_mult_0_ h_impl_InA_x_support_].
+  specialize h_impl_InA_x_support_ with (1:=h_gt_mult_x_o_0_).
+  
+  specialize support_mult_pos with (x:=y)(ms:=o).
+  intros [h_impl_gt_mult_0_0 h_impl_InA_y_support_].
+  specialize h_impl_InA_y_support_ with (1:=h_gt_mult_y_o_0_).
+
+  specialize (@PermutationA_split _ equiv _ x (support o) h_impl_InA_x_support_) as h_permut.
+  destruct h_permut as [l' hl'].
+  transitivity (x :: l');auto.
+  constructor 2;auto.
+
+  assert (InA equiv y l') as h_InA_y_l'_.
+  { rewrite  hl' in h_impl_InA_y_support_.
+    inversion h_impl_InA_y_support_.
+    - exfalso. symmetry in h_nequiv_x_y_. contradiction.
+    - assumption. }
+
+  specialize (@PermutationA_split _ equiv _ y l' h_InA_y_l'_) as h_permut.
+  destruct h_permut as [l'' hl''].
+  transitivity (y :: l'').
+  { assumption. }
+  constructor 2;auto.
+  rewrite hl'' in hl'.
+  destruct l''.
+  - reflexivity.
+  - absurd (o[e]=0).
+    all:cycle 1.
+    + assert (NoDupA equiv (e :: x :: y :: l'')) as h_NodupA_cons_e_x_.
+      { eapply PermutationA_preserves_NoDupA with (l₁ := (x :: e :: y :: l''));try typeclasses eauto.
+        constructor 3.
+        eapply PermutationA_preserves_NoDupA with (l₁ := (x :: y :: e :: l''));try typeclasses eauto.
+        constructor 2;auto. 
+        constructor 3. 
+        apply PermutationA_preserves_NoDupA with (2:=hl');try typeclasses eauto.
+        eapply support_NoDupA. }
+      inversion h_NodupA_cons_e_x_ as [ | x0 l h_not_InA_e_cons_ h_NodupA_cons_x_y_ [h_eq_x0_e_ h_eq_l_cons_x_y_] ].
+      apply h_all_eq_mult_0_.
+      * intro abs.
+        apply h_not_InA_e_cons_.
+        rewrite abs.
+        now constructor 1.
+      * intro abs.
+        apply h_not_InA_e_cons_.
+        rewrite abs.
+        constructor 2.
+        now constructor 1.
+    + apply Nat.neq_0_lt_0.
+      apply -> gt_lt_iff.
+      apply support_mult_pos.
+      rewrite hl'.
+      intuition.
+Qed.
+
+
+Lemma support_div2_split:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall o x y, o[x] = Nat.div2 (cardinal o + 1)
+                -> o[y] = Nat.div2 (cardinal o + 1)
+                -> o[x] + o[y] >= cardinal o.
+Proof using Type.
+  clear lt_0n delta_g0 delta n.
+  intros elt S0 H Ops h_FMultisetsOn_Ops_ o x y h_eq_mult_x_o_div2_add_ h_eq_mult_y_o_div2_add_ .
+  specialize div2_sum_p1_ge with (n:=cardinal o).
+  intro h_le_card_o_add_div2_div2_.
+  rewrite <- h_eq_mult_x_o_div2_add_ in h_le_card_o_add_div2_div2_ at 1.
+  rewrite <- h_eq_mult_x_o_div2_add_ in h_le_card_o_add_div2_div2_ at 1.
+  lia.
+Qed.
+
+Lemma obs_invalid_doubleton (o:observation):
+  invalid_obs o
+  -> exists x y, PermutationA equiv (support o) (x :: y :: nil)
+                  /\ o[x] = Nat.div2 (cardinal o + 1)
+                  /\ o[y] = Nat.div2 (cardinal o + 1)
+                  /\ forall z, z =/= x -> z =/= y -> o[z]=0.
+Proof using Type.
+  clear lt_0n.
+  intros h. 
+  unfold invalid_obs in h.
+  destruct h as [x [y h]].
+  (* decompose [and] h; clear h /n. *)
+  destruct h as [h_nequiv_x_y_ [h_neq_card_o_0_ [h_eq_mult_x_o_div2_add_ h_eq_mult_y_o_div2_add_]]].
+  specialize @obs_invalid_doubleton1 with (o:=o)  (2:=h_nequiv_x_y_).
+  intro h_impl_eq_mult_0_.
+  specialize (h_impl_eq_mult_0_ _).
+
+  assert(o[x] + o[y] >= cardinal o) as h_ge_add_mult_mult_card_o_.
+  { apply support_div2_split;auto. }
+  specialize h_impl_eq_mult_0_ with (1:=h_ge_add_mult_mult_card_o_).
+
+  specialize @obs_invalid_doubleton2 with (o:=o)  (2:=h_nequiv_x_y_).
+  intro h_impl_PermutA_support_cons_.
+  assert (h_div2_pos: 0 < Nat.div2 (cardinal o + 1)) by (apply card_div2;eauto).
+  assert (o[x] > 0) as h_gt_mult_x_o_0_ by lia  .
+  assert (o[y] > 0) as h_gt_mult_y_o_0_ by lia.
+  specialize (h_impl_PermutA_support_cons_ _ h_gt_mult_x_o_0_ h_gt_mult_y_o_0_).
+  exists x, y.
+  repeat split;auto.
+Qed.
+
+Lemma NoDupA_2fst: forall [A : Type] (eqA : relation A) (x y: A) l, NoDupA eqA (x::y::l) ->  ~ eqA x y.
+Proof using Type.
+  intros A eqA x y l H. 
+  inversion H;subst.
+  intro abs.
+  apply H2.
+  constructor 1.
+  assumption.
+Qed.
+
+Definition obs_invalid_dec (o:observation): {invalid_obs o} + {~ invalid_obs o}.
+Proof.
+  clear lt_0n.
+  (* specialize (cardinal_map_support o) as h. *)
+  destruct (support o) as [|a l] eqn:heq.
+  { apply support_nil in heq.
+    right.
+    intro abs.
+    rewrite heq in *.
+    specialize obs_invalid_doubleton with (1:=abs).
+    intros h_ex_and_PermutA_and_.
+    rewrite support_empty in h_ex_and_PermutA_and_.
+    destruct h_ex_and_PermutA_and_
+      as [ x [ x0 [ h_PermutA_nil_cons_x_x0_ [
+                    h_eq_mult_x_empty_div2_add_ 
+                      [ h_eq_mult_x0_empty_div2_add_ h_all_eq_mult_0_]]]]].
+    
+    assert (length (@nil location) = length (x::x0::nil)) as h_eq_length_nil_length_cons_ .
+    { rewrite h_PermutA_nil_cons_x_x0_ at 1;auto. }
+    cbn in h_eq_length_nil_length_cons_.
+    discriminate h_eq_length_nil_length_cons_. }
+(*  assert (h_card_neq_0:cardinal o <> 0).
+  { assert (0 < cardinal o).
+    { eapply cardinal_In with (x:=a).
+      apply support_spec.
+      rewrite heq.
+      constructor 1;auto. }
+    lia. }*)
+  destruct l as [ | b l'].
+  { right.
+    intro abs.
+    specialize obs_invalid_doubleton with (1:=abs).
+    intro h_ex_and_PermutA_and_ .
+    rewrite heq in *.
+    destruct h_ex_and_PermutA_and_
+      as [x [ x0 [ h_PermutA_cons_a_cons_x_x0_
+                     [ h_eq_mult_x_o_div2_add_ [ h_eq_mult_x0_o_div2_add_ h_all_eq_mult_0_]]]]].
+    assert (length (a::nil) = length (x::x0::nil)) as h_eq_length_cons_length_cons_.
+    { rewrite h_PermutA_cons_a_cons_x_x0_;auto. }
+    cbn in h_eq_length_cons_length_cons_.
+    discriminate h_eq_length_cons_length_cons_. }
+  destruct l' as [ | c l''].
+  all:swap 1 2.
+  { right.
+    intro abs.
+    specialize obs_invalid_doubleton with (1:=abs).
+    intro h_ex_and_PermutA_and_ .
+    rewrite heq in *.
+    destruct  h_ex_and_PermutA_and_
+              as [x [ x0 [ h_PermutA_cons_a_b_cons_x_x0_ 
+                             [ h_eq_mult_x_o_div2_add_
+                                 [ h_eq_mult_x0_o_div2_add_ h_all_eq_mult_0_]]]]].
+    assert (length (a::b::c::l'') = length (x::x0::nil)) as h_eq_length_cons_length_cons_.
+    { rewrite h_PermutA_cons_a_b_cons_x_x0_;auto. }
+    cbn in h_eq_length_cons_length_cons_.
+    discriminate h_eq_length_cons_length_cons_. }
+  (* Il y a deux positions occupées exactement. Il faut maintenant
+     vérifier qu'elles ont exactement la même multiplicité. *)
+
+  specialize (cardinal_map_support o) as h_card.
+  rewrite heq in h_card.
+  cbn [fold_left] in h_card.
+  rewrite Nat.add_0_r in h_card.
+  destruct (Nat.eq_dec (o[a]) (o[b])).
+  - left.
+    assert (Nat.Even (cardinal o)) as h_Even_card_o_ .
+    { rewrite e in h_card.
+      replace (o[b] + o[b]) with (2 * o[b]) in h_card;[|lia].
+      red.
+      eauto. }
+    assert (cardinal o = (2 * (Nat.div2 (cardinal o + 1)))).
+    { specialize even_div2 with (1:=h_Even_card_o_) as h_eq_add_div2_div2_card_o_ .
+      rewrite <- even_div2_p1 in h_eq_add_div2_div2_card_o_;auto.
+      lia. }
+    assert (o[a] = (Nat.div2 (cardinal o + 1))) as h_eq_mult_a_o_div2_add_.
+    { lia. }
+    assert (o[b] = (Nat.div2 (cardinal o + 1))) as h_eq_mult_b_o_div2_add_.
+    { lia. }
+    red. exists a, b;repeat (split;auto).
+    + apply NoDupA_2fst with (l:=nil).
+      rewrite <- heq.
+      apply support_NoDupA.
+    + symmetry.
+       apply Nat.lt_neq.
+       unfold lt.
+       transitivity (Nat.div2 (cardinal o + 1)).
+       { apply card_div2.
+         symmetry.
+         apply Nat.lt_neq.
+         eapply cardinal_In with (x:=a).
+         apply support_spec.
+         rewrite heq.
+         intuition. }
+       rewrite <- h_eq_mult_a_o_div2_add_.
+       apply cardinal_lower.
+  - right.
+    intro h_abs.
+    specialize obs_invalid_doubleton with (1:=h_abs) as h_ex_and_PermutA_and_ .
+    rewrite heq in *.
+    destruct h_ex_and_PermutA_and_
+      as [ x [ x0 
+                 [ h_PermutA_cons_a_b_cons_x_x0_
+                     [ h_eq_mult_x_o_div2_add_
+                         [ h_eq_mult_x0_o_div2_add_ h_all_eq_mult_0_]]]]].
+    assert (support o = [x ;x0] \/ support o = [x0 ;x]) as h_or_eq_support_cons_eq_support_cons_.
+     { rewrite heq.
+       specialize (@PermutationA_2 _ equiv _ a b x x0) as h_iff_PermutA_cons_cons_or_and_and_.
+       destruct h_iff_PermutA_cons_cons_or_and_and_ as [ h_impl_or_and_and_ ]; intros.
+       specialize h_impl_or_and_and_ with (1:=h_PermutA_cons_a_b_cons_x_x0_).
+       (* decompose [and or] h_impl_or_and_and_ /n . *)
+       destruct h_impl_or_and_and_ as [ [ h_equiv_a_x_ h_equiv_b_x0_ ] | [ h_equiv_a_x0_ h_equiv_b_x_] ].
+       - left.
+         cbn in h_equiv_a_x_,h_equiv_b_x0_.
+         rewrite h_equiv_a_x_,h_equiv_b_x0_.
+         reflexivity.
+       - right.
+         cbn in h_equiv_a_x0_,h_equiv_b_x_.
+         rewrite h_equiv_a_x0_,h_equiv_b_x_.
+         reflexivity. }
+     destruct h_or_eq_support_cons_eq_support_cons_
+       as [h_eq_support_o_cons_x_x0_  | h_eq_support_o_cons_x0_x_ ].
+    + rewrite heq in h_eq_support_o_cons_x_x0_.
+      inversion h_eq_support_o_cons_x_x0_;subst.
+      rewrite <- h_eq_mult_x0_o_div2_add_ in h_eq_mult_x_o_div2_add_.
+      contradiction.
+    + rewrite heq in h_eq_support_o_cons_x0_x_.
+      inversion h_eq_support_o_cons_x0_x_;subst.
+      rewrite <- h_eq_mult_x_o_div2_add_ in h_eq_mult_x0_o_div2_add_.
+      contradiction.
+Defined.
+
+(* Ltac rename_depth ::= constr:(2). *)
+
+
+(* Local Instance gatherW_pgm_compat : Proper (@equiv _ observation_Setoid ==> ???) obs_invalid_dec. *)
+Lemma obs_invalid_dec_compat_left: forall o o': observation,
+    equiv o o' ->
+    forall h,
+      obs_invalid_dec o = left h ->
+      exists h', obs_invalid_dec o' = left h'.
+Proof using Type.
+  intros o o' h_equiv_o_o'_ h_invalid_obs_o_ h_eq_obs_invalid_dec_left_ .
+  destruct (obs_invalid_dec o') as [h_invalid_obs_o'_ | h_not_invalid_obs_ ] eqn:h_eq_obs_invalid_dec_ ;eauto.
+  foldR2.
+  clear h_eq_obs_invalid_dec_left_ h_eq_obs_invalid_dec_.
+  setoid_rewrite h_equiv_o_o'_ in h_invalid_obs_o_.
+  contradiction.
+Qed.
+
+(* Local Instance gatherW_pgm_compat : Proper (@equiv _ observation_Setoid ==> ???) obs_invalid_dec. *)
+Lemma obs_invalid_dec_compat_right: forall o o': observation,
+    equiv o o' ->
+    forall h,
+      obs_invalid_dec o = right h ->
+      exists h', obs_invalid_dec o' = right h'.
+Proof using Type.
+  intros o o' h_equiv_o_o'_ h_not_invalid_obs_ h_eq_obs_invalid_dec_right_ .
+  destruct (obs_invalid_dec o') eqn:h_eq_obs_invalid_dec_;eauto.
+  foldR2.
+  clear h_eq_obs_invalid_dec_  h_eq_obs_invalid_dec_right_.
+  setoid_rewrite h_equiv_o_o'_ in h_not_invalid_obs_.
+  contradiction.
+Qed.
+
+
+(* The main algorithm.
+ * Note that [obs] describes the positions of the robots in the 
+ * LOCAL frame of reference.
+ * I chose to express the robogram with segments, and then switch to lines 
+ * in the 'round_simplify_phaseX' lemmas. *)
+Definition gatherW_pgm (obs : observation) : location := 
+  if obs_invalid_dec obs then 0%VS 
+  else
+    if (List.length (multi_support obs) =?= 0) then 0%VS
+    else
+      let '(w1, w2) := weber_calc_seg (multi_support obs) in
+      let L := line_calc (multi_support obs) in 
+      let m :=
+        let r1 := L^-P (L^min (multi_support obs)) in 
+        let r2 := L^-P (L^max (multi_support obs)) in        
+        middle r1 r2 in
+      (* phase 1 : move to the unique weber point *)   
+      if w1 =?= w2 then w1 else 
+        (* From now on [n] is even and the robots are aligned 
+         * (since the weber point isn't unique). *)
+        (* We use cardinal obs instead of n, for respectfulness. We
+           also use the fact that robot_with_mult is applied to
+           multi_support obs which always return the same list (not a
+           permutation). *)
+        match robot_with_mult (Nat.div2 (cardinal obs +1)) (multi_support obs) with
+        | Some t =>
+            (* phase 2 : safe move to [t] *)
+            if existsb (strict_segment_decb 0%VS t) (multi_support obs) 
+            then 0%VS else t
+        | None =>
+            if segment_decb w1 w2 m then 
+              (* phase 3 : robots on [w1] or [w2] move to the middle *)
+              if (w1 ==b 0%VS) || (w2 ==b 0%VS) then m else 0%VS
+            else if segment_decb m w1 w2 then 
+                   (* phase 4 : robots on [w1] move to [w2] *)
+                   if w1 ==b 0%VS then w2 else 0%VS 
+                 else if segment_decb m w2 w1 then
+                        (* also phase 4 : robots on [w2] move to [w1] *)
+                        if w2 ==b 0%VS then w1 else 0%VS
+                      else 
+                        (* this should not happen *) 
+                        0%VS                    
+        end.
+
+
+Lemma opt_eq_leibnitz: (forall T (x y: option T), opt_eq eq x y -> x = y).
+Proof using Type.
+  destruct x;destruct y;cbn;intros;subst;auto;contradiction.
+Qed.
+
+
+Lemma multi_support_add {A : Type} `{EqDec A} s x k : ~ In x s -> k > 0 ->
+  PermutationA equiv (multi_support (add x k s)) (alls x k ++ multi_support s).
+Proof using . 
+intros Hin Hk. unfold multi_support. 
+transitivity (flat_map (fun '(x0, mx) => alls x0 mx) ((x, k) :: elements s)).
++ apply flat_map_compat_perm with eq_pair ; autoclass ; [|].
+  - intros [a ka] [b kb] [H0 H1]. cbn in H0, H1. now rewrite H0, H1.
+  - apply elements_add_out ; auto.
++ now cbn -[elements].
+Qed.
+
+Lemma multi_support_countA {A : Type} `{eq_dec : EqDec A} s x :
+  countA_occ equiv eq_dec x (multi_support s) == s[x]. 
+Proof using .
+pattern s. apply MMultisetFacts.ind.
++ intros m m' Hm. f_equiv. 
+  - apply countA_occ_compat ; autoclass. now rewrite Hm.
+  - now rewrite Hm.
++ intros m x' n' Hin Hn IH. rewrite add_spec, multi_support_add, countA_occ_app by auto.
+  destruct_match.
+  - now rewrite <-H, countA_occ_alls_in, Nat.add_comm, IH ; autoclass.
+  - now rewrite countA_occ_alls_out, IH, Nat.add_0_l ; auto.  
++ now reflexivity.
+Qed.
+
+(* This is the main result about multi_support. *)
+(* RMK : typeclass instance resolution seems to be VERY slow in this proof.
+ * Thankfully I found that the 'change' tactic is fast here. *)
+Lemma multi_support_config (config : configuration) : 
+  PermutationA equiv 
+    (multi_support (!! config)) 
+    (pos_list config).
+Proof using . 
+pose (l := pos_list config). fold l.
+change (!! config) with (make_multiset l).
+apply PermutationA_countA_occ with R2_EqDec ; autoclass ; [].
+intros x. rewrite multi_support_countA. now apply make_multiset_spec.
+Qed. 
+
+Remove Hints FMapFacts.eq_key_elt_Setoid FMapFacts.eq_key_Setoid FMapFacts.eq_key_elt_Equivalence FMapFacts.eq_key_Equivalence FMapFacts.Equal_ST FMapFacts.eq_key_EqDec: typeclass_instances.
+
+Corollary multi_support_map f config : 
+  Proper (equiv ==> equiv) (projT1 f) ->
+  PermutationA equiv 
+    (multi_support (!! (map_config (lift f) config)))
+    (List.map (projT1 f) (pos_list config)).
+Proof using .  
+intros H. destruct f as [f Pf].
+rewrite multi_support_config. unfold pos_list. rewrite config_list_map, map_map.
++ apply eqlistA_PermutationA. rewrite map_map. f_equiv. intros [[s d] r] [[s' d'] r'] Hsdr.
+  inversion Hsdr. cbn in H0. inversion H0. subst.
+  cbn -[equiv straight_path]. destruct Pf as [sim Hsim].
+  rewrite <-Hsim.
+  rewrite H1.
+  apply straight_path_similarity.
+(* + intros [[s d] r] [[s' d'] r'] [[Hs Hd] Hr]. cbn -[equiv] in H, Hs, Hd, Hr |- *.  *)
+  (* repeat split ; cbn -[equiv] ; auto. *)
+Qed.
+
+Lemma multi_support_non_nil f config : 
+  Proper (equiv ==> equiv) (projT1 f) ->
+  multi_support (!! (map_config (lift f) config)) <> [].
+Proof using lt_0n.
+  intros h' abs.
+  specialize (@multi_support_map f config h').
+  intros H. 
+  rewrite abs in H.
+  symmetry in H.
+  apply Preliminary.PermutationA_nil in H.
+  apply map_eq_nil in H.
+  eelim pos_list_non_nil.
+  - eassumption.
+  - typeclasses eauto.
+Qed.
+
+Lemma multi_support_non_nil2 (config : configuration):
+  multi_support (!! config) <> [].
+Proof using lt_0n.
+  intro abs.
+  specialize (@multi_support_config config) as h.
+  foldR2.
+  rewrite abs in h.
+  apply pos_list_non_nil with config.
+  foldR2.
+  eapply Preliminary.PermutationA_nil.
+  2:{ symmetry.
+      eapply h. }
+  typeclasses eauto.
+Qed.
+
+
+Corollary multiplicity_countA config x : 
+  (!! config)[x] = countA_occ equiv R2_EqDec x (pos_list config).
+Proof using . now rewrite <-multi_support_config, multi_support_countA. Qed.
+
+Lemma multiplicity_pos_InA x config : 
+  (!!config)[x] > 0 <-> InA equiv x (pos_list config).
+Proof using . rewrite multiplicity_countA. apply countA_occ_pos ; autoclass. Qed.
+
+Lemma multiplicity_sum2_bound x y config : 
+  x =/= y -> (!!config)[x] + (!!config)[y] <= n.
+Proof using.
+  clear lt_0n delta_g0.
+  intros h_neq. 
+  repeat rewrite multiplicity_countA.
+  rewrite <- (pos_list_length config).
+  remember (pos_list config) as l.
+  symmetry in h_neq.
+  specialize (@countA_occ_removeA_other _ _ _ location_EqDec y x l h_neq) as h_eq_countA_occ_countA_occ_ .
+  foldR2.
+  rewrite <- h_eq_countA_occ_countA_occ_.
+  assert (countA_occ equiv location_EqDec y l <= length (removeA (eqA:=equiv) location_EqDec x l)).
+  { rewrite <- h_eq_countA_occ_countA_occ_.
+    apply countA_occ_length_le. }
+  assert (length (removeA (eqA:=equiv) location_EqDec x l) + countA_occ equiv location_EqDec x l = length l).
+  { specialize (@PermutationA_count_split _ _ _ location_EqDec x l) as h_PermutA_l_app_ .
+    rewrite h_PermutA_l_app_ at 3.
+    rewrite app_length.
+    rewrite alls_length.
+    lia. }
+  lia.
+Qed.
+
+Lemma multiplicity_sum3_bound x y z config : 
+  x =/= y -> 
+  y =/= z -> 
+  x =/= z ->
+  (!!config)[x] + (!!config)[y] + (!!config)[z] <= n.
+Proof using .
+  clear lt_0n delta_g0.
+  intros h_nequiv_x_y_ h_nequiv_y_z_ h_nequiv_x_z_ .
+  fold info0.
+  repeat rewrite multiplicity_countA.
+  remember (pos_list config) as l.
+  symmetry in h_nequiv_x_y_.
+  specialize (@countA_occ_removeA_other _ _ _ location_EqDec y x l h_nequiv_x_y_) as h_eq_countA_occ_countA_occ_ .
+  foldR2.
+  rewrite <- h_eq_countA_occ_countA_occ_.
+  symmetry in h_nequiv_x_z_.
+  specialize (@countA_occ_removeA_other _ _ _ location_EqDec z x l h_nequiv_x_z_) as h_eq_countA_occ_countA_occ_0 .
+  foldR2.
+  rewrite <- h_eq_countA_occ_countA_occ_0.
+  symmetry in h_nequiv_y_z_.
+  specialize (@countA_occ_removeA_other _ _ _ location_EqDec z y (removeA (eqA:=equiv) location_EqDec x l) h_nequiv_y_z_) as h_eq_countA_occ_countA_occ_1 .
+  foldR2.
+  rewrite <- h_eq_countA_occ_countA_occ_1.
+  remember (removeA (eqA:=equiv) location_EqDec x l) as l_nox.
+  remember (removeA (eqA:=equiv) location_EqDec y l_nox) as l_noxy.
+  assert (countA_occ equiv location_EqDec y l <= length l_nox) as h_le_countA_occ_length_ .
+  { rewrite <- h_eq_countA_occ_countA_occ_.
+    apply countA_occ_length_le. }
+  assert (countA_occ equiv location_EqDec z l_nox <= length l_noxy) as h_le_countA_occ_length_0 .
+  { rewrite <- h_eq_countA_occ_countA_occ_1.
+    apply countA_occ_length_le. }
+  assert (length (removeA (eqA:=equiv) location_EqDec x l) + countA_occ equiv location_EqDec x l = length l) as h_eq_add_length_ .
+  { specialize (@PermutationA_count_split _ _ _ location_EqDec x l) as h_PermutA_l_app_ .
+    rewrite h_PermutA_l_app_ at 3.
+    rewrite app_length.
+    rewrite alls_length.
+    lia. }
+  assert (length (removeA (eqA:=equiv) location_EqDec y l_nox) + countA_occ equiv location_EqDec y l_nox = length l_nox) as h_eq_add_length_0 .
+  { specialize (@PermutationA_count_split _ _ _ location_EqDec y l_nox) as h_PermutA_l_nox_app_ .
+    rewrite h_PermutA_l_nox_app_ at 3.
+    rewrite app_length.
+    rewrite alls_length.
+    lia. }
+  remember (countA_occ equiv location_EqDec x l) as h_l.
+  remember (countA_occ equiv location_EqDec y l_nox) as h_l_nox.
+  remember (countA_occ equiv location_EqDec z l_noxy) as h_l_noxy.
+  specialize (n_cardinal config info0) as h_eq_card_n_ .
+  rewrite <- (pos_list_length config).
+  rewrite <- Heql in *.
+  subst.
+  lia.
+Qed.
+
+
+Lemma invalid_refine_endpoints c L :
+  invalid c -> 
+  aligned_on L (pos_list c) -> 
+    (L^min (pos_list c) < L^max (pos_list c))%R /\
+    (!! c)[L^-P (L^min (pos_list c))] = Nat.div2 (n+1) /\
+    (!! c)[L^-P (L^max (pos_list c))] = Nat.div2 (n+1).
+Proof using lt_0n delta_g0.
+pose (r1 := L^-P (L^min (pos_list c))).
+pose (r2 := L^-P (L^max (pos_list c))).
+fold r1 r2.
+assert (Hr1 : L^min (pos_list c) == L^P r1). { unfold r1. now rewrite line_P_iP_min. }
+assert (Hr2 : L^max (pos_list c) == L^P r2). { unfold r2. now rewrite line_P_iP_max. }
+intros Hinvalid Halign. assert (H := Hinvalid). destruct H as [x [y [NExy [Hx_mult Hy_mult]]]].
+assert (Hsum : (!! c)[x] + (!! c)[y] = n).
+{ rewrite Hx_mult, Hy_mult. rewrite even_div2_p1, even_div2 by eapply invalid_even, Hinvalid. reflexivity. }
+assert (Htowers : forall z, InA equiv z (pos_list c) -> z == x \/ z == y).
+{
+  intros z Hinz. 
+  case (x =?= z) as [Hid_x | Hid_x] ; [now left|].
+  case (y =?= z) as [Hid_y | Hid_y] ; [now right|].
+  exfalso. 
+  assert (Hbound := @multiplicity_sum3_bound x y z c).
+  feed_n 3 Hbound ; [assumption .. |].
+  rewrite Hx_mult, Hy_mult in Hbound. rewrite even_div2_p1, even_div2 in Hbound by eapply invalid_even, Hinvalid.
+  rewrite <-multiplicity_pos_InA in Hinz. lia. 
+}
+repeat split.
++ case (line_min_le_max L (pos_list c)) as [Hr12 | Hr12] ; [intuition|].
+  assert (Heq_r1 : forall z, InA equiv z (pos_list c) -> z == r1).
+  {
+    intros z Hinz.
+    assert (Hz := line_bounds L (pos_list c) z Hinz).
+    rewrite <-Hr12 in Hz.
+    assert (HzP : L^P z = L^min (pos_list c)). { lra. }
+    apply (f_equal (L^-P)) in HzP. rewrite Halign in HzP ; [exact HzP | exact Hinz].
+  }
+  exfalso. apply NExy. transitivity r1.
+  - apply Heq_r1. rewrite <-multiplicity_pos_InA, Hx_mult, gt_lt_iff.
+    apply Exp_prop.div2_not_R0 ; lia.
+  - symmetry. apply Heq_r1. rewrite <-multiplicity_pos_InA, Hy_mult, gt_lt_iff.
+    apply Exp_prop.div2_not_R0 ; lia.
++ specialize (Htowers r1). feed Htowers.
+  { unfold r1. apply line_iP_min_InA ; [|exact Halign]. rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+  now case Htowers as [Hr1x | Hr1y] ; [rewrite Hr1x, Hx_mult | rewrite Hr1y, Hy_mult].
++ specialize (Htowers r2). feed Htowers.
+  { unfold r2. apply line_iP_max_InA ; [|exact Halign]. rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+  now case Htowers as [Hr2x | Hr2y] ; [rewrite Hr2x, Hx_mult | rewrite Hr2y, Hy_mult].
+Qed.
+
+Lemma two_majority_invalid c x y : 
+  x =/= y ->  
+  Nat.div2 (n+1) <= (!! c)[x] -> 
+  Nat.div2 (n+1) <= (!! c)[y] ->
+  invalid c.
+Proof using lt_0n .
+intros NExy Hx Hy. exists x, y. split ; [assumption |].
+assert (Hbound := multiplicity_sum2_bound c NExy). 
+assert (Hle := div2_sum_p1_ge n). lia. 
+Qed.
+
+Lemma lift_update_swap da config1 config2 g target :
+  @equiv info _
+    (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                               (precondition_satisfied_inv da config1 g))
+          (update config2
+           g (change_frame da config1 g) target (choose_update da config2 g target)))
+    (update (map_config (lift (existT precondition (frame_choice_bijection (change_frame da config1 g ⁻¹))
+                                      (precondition_satisfied_inv da config1 g)))
+                        config2)
+            g Similarity.id
+            ((frame_choice_bijection (change_frame da config1 g ⁻¹)) target)
+            (choose_update da config2 g target)).
+Proof using .
+pose (sim := change_frame da config1 g). fold sim.
+cbn -[inverse equiv straight_path]. destruct (config2 (Good g)) as [[start dest] r].
+now rewrite straight_path_similarity.
+Qed.
+
+Lemma line_middle_sim L L' (sim : similarity R2) ps : 
+  ps <> nil -> aligned_on L ps -> aligned_on L' (List.map sim ps) -> 
+    middle (L'^-P (L'^min (List.map sim ps))) (L'^-P (L'^max (List.map sim ps))) ==
+    sim (middle (L^-P (L^min ps)) (L^-P (L^max ps))).
+Proof using . 
+intros Hnil Halign Halign'.
+pose (r1 := L^-P (L^min ps)). pose (r2 := L^-P (L^max ps)).
+pose (r1' := L'^-P (L'^min (List.map sim ps))). pose (r2' := L'^-P (L'^max (List.map sim ps))).
+foldR2. fold r1 r2 r1' r2'.
+cut (perm_pairA equiv (r1', r2') (sim r1, sim r2)).
+{ now intros [[-> ->] | [-> ->]] ; rewrite R2_middle_morph ; [|rewrite middle_comm]. } 
+assert (Hr12 : perm_pairA equiv (r1, r2) (L^-P (L^min ps), L^-P (L^max ps))).
+{ left. intuition. }
+symmetry. unfold r1', r2'. rewrite line_min_max_spec in Hr12 |- * ; auto.
++ repeat split.
+  - rewrite InA_map_iff ; autoclass ; []. exists r1. intuition.
+  - rewrite InA_map_iff ; autoclass ; []. exists r2. intuition.
+  - intros y. rewrite InA_map_iff ; autoclass ; []. intros [x [<- Hin]].
+    rewrite <-R2segment_similarity. now apply Hr12.
++ intros Hsim_nil. eapply Hnil, map_eq_nil ; eauto.  
+Qed.  
+
+
+
+
+Lemma robot_with_mult_unique_obs o ps: 
+  ~invalid_obs o ->
+  PermutationA equiv (multi_support o) ps ->
+  robot_with_mult (Nat.div2 (cardinal o+1)) (multi_support o) == robot_with_mult (Nat.div2 (cardinal o+1)) ps.
+Proof using Type.
+  intros Hvalid Hperm.
+  unfold robot_with_mult.
+  assert (find (fun x : R2 => countA_occ equiv R2_EqDec x (multi_support o) ==b Nat.div2 (cardinal o + 1))
+    (multi_support o) == find (fun x : R2 => countA_occ equiv R2_EqDec x ps ==b Nat.div2 (cardinal o + 1))
+    (multi_support o)).
+  { apply find_compat.
+    - intros x y h_equiv_x_y_ .
+      rewrite h_equiv_x_y_.
+      foldR2.
+      specialize countA_occ_compat with (eqA:=@eq (@location Loc)) as h.
+      specialize (h _).
+      foldR2.
+      unfold Proper, respectful in h.
+      foldR2.
+      cbn in h_equiv_x_y_.
+      specialize h with (x:=y)( y:=y)(1:=eq_refl)(eq_dec:=R2_EqDec).
+      specialize h with (x:=(@multi_support _ _ R2_EqDec o)).
+      specialize h with (y:=ps).
+      cbn -[countA_occ location_EqDec multi_support].
+      foldR2.
+      rewrite h;auto.
+    - reflexivity. }
+  rewrite H.
+  clear H.
+  remember (fun x : R2 => countA_occ equiv R2_EqDec x ps ==b Nat.div2 (cardinal o + 1)) as f.
+  match goal with
+    |- ?A == ?B => destruct A eqn:HeqA; destruct B eqn:HeqB ;try discriminate
+  end.
+  all:swap 1 4.
+  - reflexivity.
+  - erewrite list_perm_find_none in HeqA;try typeclasses eauto;eauto.
+    + symmetry. assumption.
+    + rewrite Heqf.
+      now intros i j ->.
+  - erewrite list_perm_find_none in HeqB;try typeclasses eauto;eauto.
+    rewrite Heqf.
+    now intros i j ->.
+  - cbn.
+    destruct (R2_EqDec r r0).
+    + auto.
+    + exfalso.
+      apply Hvalid.
+      unfold invalid.
+      exists r, r0.
+      specialize find_some with (1:=HeqA) as [hinr hfr]. 
+      specialize find_some with (1:=HeqB) as [hinr0 hfr0].
+      repeat rewrite obs_from_config_spec.
+      rewrite Heqf in hfr,hfr0.
+      rewrite eqb_true_iff in hfr,hfr0.
+      repeat constructor;auto.
+      * unfold multi_support in hinr.
+        apply in_flat_map_Exists in hinr.
+        apply Exists_exists in hinr.
+        destruct hinr as [ x [h_In_x_elements_ h_In_r_] ].
+        symmetry.
+        apply Nat.lt_neq.
+        apply Nat.lt_le_trans with (size o).
+        2: apply size_cardinal.
+
+        rewrite size_elements.
+        apply Nat.neq_0_lt_0.
+        apply list_in_length_n0 with (1:=h_In_x_elements_).
+      * rewrite <- hfr.
+        rewrite <- Hperm.
+        symmetry.
+        apply  multi_support_countA.
+      * rewrite <- hfr0.
+        rewrite <- Hperm.
+        symmetry.
+        apply  multi_support_countA.
+Qed.
+
+
+Lemma multi_set_length_cardinal:
+  forall {elt : Type} {S0 : Setoid elt} {H : EqDec S0} {Ops : FMOps elt H}
+         {FM: FMultisetsOn elt Ops},
+  forall o, length (multi_support o) = cardinal o.
+Proof using lt_0n.
+
+  intros elt S0 H Ops FM o.
+  (* rewrite cardinal_map_support. *)
+  (* unfold multi_support. *)
+  (* rewrite flat_map_concat_map. *)
+
+  foldR2.
+  
+  specialize (from_elements_elements o) as h_equiv_from_elements_o_ .
+  foldR2.
+  rewrite <- h_equiv_from_elements_o_ at 2.
+  rewrite cardinal_from_elements.
+  rewrite  <- fold_left_length.
+
+  unfold multi_support.
+  rewrite flat_map_concat_map.
+  assert (forall l k,
+             fold_left (fun (x : nat) (_ : elt) => S x) (concat (List.map (fun '(x, mx) => alls x mx) l)) k =
+               fold_left (fun (acc : nat) (xn : elt * nat) => snd xn + acc) l k).
+  { induction l as [ | a l h_all_eq_ ].
+    - simpl;intros;auto.
+    - simpl;intros.
+      destruct a.
+      cbn.
+      rewrite fold_left_app.
+      assert ((fold_left (fun (x : nat) (_ : elt) => S x) (alls e n0) k) = n0 + k) as h_eq_fold_left_add_.
+      { assert (forall l k, (fold_left (fun (x : nat) (_ : elt) => S x) l k) = length l + k) as h_all_eq_0.
+        { induction l0;simpl;auto.
+          intros k0. 
+          rewrite IHl0.
+          lia. }
+        rewrite h_all_eq_0.
+        rewrite alls_length.
+        reflexivity.
+      }
+      rewrite h_eq_fold_left_add_.
+      apply h_all_eq_.
+  }
+  apply H0.
+Qed.
+
+(* FIXED: This was only true for valid position, we modified pgm  to fix this. *)
+Local Instance gatherW_pgm_compat : Proper (@equiv _ observation_Setoid ==> @equiv _ location_Setoid) gatherW_pgm.
+Proof using lt_0n. 
+  clear delta_g0.
+  intros x y H .
+  unfold observation,multiset_observation in x.
+
+  specialize @multi_support_compat as hsupp.
+  unfold Proper,respectful in hsupp.
+  specialize hsupp with (1:=H).
+  specialize weber_calc_seg_compat as h.
+  unfold Proper,respectful in h.
+  specialize h with (1:=hsupp).
+  lazy beta delta [gatherW_pgm]. (* unfold reduces letins *)
+  destruct (obs_invalid_dec x) eqn:h_eq_obs_inv.
+  (* If the observation is invalid, robogram does nothing *)
+  { destruct (obs_invalid_dec_compat_left H h_eq_obs_inv) as [i' h'].
+    rewrite h'.
+    reflexivity. }
+
+  destruct (obs_invalid_dec_compat_right H h_eq_obs_inv) as [i' h'].
+  rewrite h'.
+  clear h' h_eq_obs_inv.
+  
+  foldR2.
+  assert (length (multi_support x) == length (multi_support y)) as h_equiv_length_length_ .
+  { rewrite hsupp.
+    reflexivity. }
+
+  foldR2.
+  rewrite <- h_equiv_length_length_.
+  destruct (length (multi_support x) =?= 0) as [ h_equiv_length_0_ | h_nequiv_length_0_ ] eqn:h_eq_equiv_dec.
+  { reflexivity. }
+  clear h_eq_equiv_dec.
+
+  assert (multi_support x <> []) as h_neq_multi_support_nil_ .
+  { intro abs.
+    rewrite abs in h_nequiv_length_0_.
+    cbn in h_nequiv_length_0_.
+    now apply h_nequiv_length_0_. }
+
+  foldR2.
+  destruct (weber_calc_seg (multi_support x)) eqn: hcalcx.
+  destruct (weber_calc_seg (multi_support y)) eqn: hcalcy.
+  foldR2.
+  red in h.
+  destruct (r =?= r0) as [h_equiv_r_r0_ | h_nequiv_r_r0_];
+    destruct (r1 =?= r2) as [ h_equiv_r1_r2_ | h_nequiv_r1_r2_ ].
+  
+
+  { decompose [or and] h ;clear h.
+    + assumption.
+    + now rewrite h_equiv_r_r0_. }
+  { exfalso.
+    destruct h as [ [h_equiv_r_r1_  h_equiv_r0_r2_] | [h_equiv_r_r2_  h_equiv_r0_r1_]].
+    + rewrite <- h_equiv_r_r0_ in h_equiv_r0_r2_.
+      rewrite h_equiv_r_r1_ in h_equiv_r0_r2_.
+      contradiction.
+    + symmetry in h_equiv_r0_r1_.
+      rewrite <- h_equiv_r_r0_ in h_equiv_r0_r1_.
+      rewrite h_equiv_r_r2_ in h_equiv_r0_r1_.
+      contradiction. }
+  { exfalso.
+    destruct h as [[h_equiv_r_r1_ h_equiv_r0_r2_] | [h_equiv_r_r2_ h_equiv_r0_r1_]].
+    + symmetry in h_equiv_r0_r2_.
+      rewrite <- h_equiv_r1_r2_ in h_equiv_r0_r2_.
+      rewrite <- h_equiv_r_r1_ in h_equiv_r0_r2_.
+      contradiction.
+    + symmetry in h_equiv_r0_r1_.
+      rewrite h_equiv_r1_r2_ in h_equiv_r0_r1_. 
+      rewrite <- h_equiv_r_r2_ in h_equiv_r0_r1_.
+      contradiction. }
+  (* From here the Weber segment is not a single point, therefore:
+     - all points are aligned. *)
+  assert (cardinal x = cardinal y) as h_eq_card_card_ .
+  { rewrite <- multi_set_length_cardinal.
+    rewrite <- multi_set_length_cardinal.
+    eapply Preliminary.PermutationA_length;eauto;try typeclasses eauto. }
+  rewrite <- h_eq_card_card_.
+
+  specialize robot_with_mult_unique_obs with (1:=n0) (2:=hsupp)
+    as h_equiv_robot_with_mult_robot_with_mult_ .
+  cbn -[robot_with_mult cardinal multi_support] in h_equiv_robot_with_mult_robot_with_mult_.
+  apply opt_eq_leibnitz in h_equiv_robot_with_mult_robot_with_mult_.
+  foldR2.
+  rewrite <- h_equiv_robot_with_mult_robot_with_mult_.
+  destruct (robot_with_mult (Nat.div2 (cardinal x + 1)) (multi_support x))
+    as [l | ] eqn:h_eq_robot_with_mult_l_  .
+  { cbn -[multi_support].
+    now rewrite hsupp. }
+  remember (line_calc (multi_support x)) as Lx.
+  remember (line_calc (multi_support y)) as Ly.
+  lazy zeta.
+  (remember ((Lx ^-P) ((Lx ^min) (multi_support x))) as r3).
+  remember ((Lx ^-P) ((Lx ^max) (multi_support x))) as r4.
+  remember ((Ly ^-P) ((Ly ^min) (multi_support y))) as r3'.
+  remember ((Ly ^-P) ((Ly ^max) (multi_support y))) as r4'.
+
+  assert (aligned (multi_support x)).
+  { assert (Weber (multi_support x) r) as h_Weber_multi_support_r_ .
+    { specialize weber_calc_seg_correct with (ps:=(multi_support x))(w:=r)(1:=h_neq_multi_support_nil_)
+        as h_webercalc.
+      rewrite hcalcx in h_webercalc.
+      apply h_webercalc.
+      apply segment_start. }
+    assert (Weber (multi_support x) r0) as h_Weber_multi_support_r0_ .
+    { specialize weber_calc_seg_correct with (ps:=(multi_support x))(w:=r0)(1:=h_neq_multi_support_nil_)
+        as h_webercalc.
+      rewrite hcalcx in h_webercalc.
+      apply h_webercalc.
+      apply segment_end. }
+    
+    destruct (aligned_dec (multi_support x)); auto.
+    exfalso.
+    assert (OnlyWeber (multi_support x) r) as h_OnlyWeber_multi_support_r_ .
+    { eapply weber_Naligned_unique;eauto. }
+    red in h_OnlyWeber_multi_support_r_.
+    destruct h_OnlyWeber_multi_support_r_ as [ h_Weber_multi_support_r_0 h_all_equiv_ ].
+    specialize h_all_equiv_ with (x:=r0) (1:=h_Weber_multi_support_r0_).
+    symmetry in h_all_equiv_.
+    contradiction. }
+  assert (aligned (multi_support y)).
+  { now rewrite <- hsupp. }
+
+  specialize @line_endpoints_invariant with (ps := multi_support x) (L1:=Lx)(L2:=Ly)
+    as h_all_perm_pairA_ .
+
+  specialize h_all_perm_pairA_ with (1:=h_neq_multi_support_nil_).
+  assert (aligned_on Lx (multi_support x)) as h_aligned_on_Lx_multi_support_ .
+  { rewrite HeqLx.
+    rewrite <- line_calc_spec.
+    eassumption. }
+  specialize h_all_perm_pairA_ with (1:=h_aligned_on_Lx_multi_support_).
+  assert (aligned_on Ly (multi_support x)) as h_aligned_on_Ly_multi_support_ .
+  { rewrite hsupp.
+    rewrite HeqLy.
+    rewrite <- line_calc_spec.
+    eassumption. }
+  specialize h_all_perm_pairA_ with (1:=h_aligned_on_Ly_multi_support_).
+
+  foldR2.
+  rewrite <- Heqr3 in h_all_perm_pairA_.
+  rewrite <- Heqr4 in h_all_perm_pairA_.
+  foldR2.
+  rewrite <- hsupp in Heqr3', Heqr4'.
+  rewrite <- Heqr3' in h_all_perm_pairA_.
+  rewrite <- Heqr4' in h_all_perm_pairA_.
+
+
+  assert ((middle r3 r4) = (middle r3' r4')) as h_eq_middle_middle_ .
+  { red in h_all_perm_pairA_.
+    destruct h_all_perm_pairA_ as [[h_equiv_r3_r3'_ h_equiv_r4_r4'_] | [ h_equiv_r3_r4'_ h_equiv_r4_r3'_ ]].
+    - rewrite h_equiv_r3_r3'_,h_equiv_r4_r4'_.
+      reflexivity.
+    - rewrite <-  h_equiv_r3_r4'_, h_equiv_r4_r3'_.
+      apply middle_comm. }
+  foldR2.
+  rewrite <- h_eq_middle_middle_.
+
+  assert (segment_decb r r0 (middle r3 r4) == segment_decb r1 r2 (middle r3 r4)) as h_equiv_segment_decb_segment_decb_ .
+  {
+    destruct h as [[ h_equiv_r_r1_ h_equiv_r0_r2_ ] | [h_equiv_r_r2_ h_equiv_r0_r1_ ]].
+    + rewrite <- h_equiv_r_r1_,h_equiv_r0_r2_.
+      reflexivity.
+    + rewrite <- h_equiv_r_r2_,h_equiv_r0_r1_. 
+      apply segment_decb_sym. }
+  rewrite h_equiv_segment_decb_segment_decb_.
+
+  destruct (segment_decb r1 r2 (middle r3 r4)) eqn:h_eq_segment_decb.
+  { destruct h as[[h_equiv_r_r1_ h_equiv_r0_r2_]|[h_equiv_r_r2_ h_equiv_r0_r1_]].
+    - rewrite <- h_equiv_r_r1_,h_equiv_r0_r2_.
+      reflexivity.
+    - rewrite <- h_equiv_r_r2_,h_equiv_r0_r1_. 
+      rewrite orb_comm.
+      reflexivity. }
+
+
+    (* TODO: lemma *)
+    assert (forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} (a b x : T), ~ segment a b x -> ~ strict_segment a b x) as not_strict_segment_not_seg.
+    { intros.
+      intro abs.
+      apply strict_segment_incl in abs.
+      contradiction. }
+    (* TODO: lemma *)
+    (* assert ((forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {ES:EuclideanSpace T} (a b p : T) (L:@line T), *)
+    (*             a =/= b -> *)
+    (*             line_contains L a -> *)
+    (*             line_contains L b -> *)
+    (*             line_contains L p ->  *)
+    (*             ~ strict_segment a b p -> *)
+    (*             { strict_segment p a b /\ ~ segment p b a} + { strict_segment p b a /\ ~ segment p a b})) as not_strict_seg_conseq. *)
+    assert (forall (a b p : location),
+               segment a b p
+               -> p =/= b
+               -> ~segment p a b
+           ) as h_lemma2.
+    { intros a b p h_segment_a_b_p_ h_nequiv_p_b_ .
+      rewrite  segment_dist_sum in h_segment_a_b_p_.
+      intro abs.
+      rewrite  segment_dist_sum in abs.
+      rewrite dist_sym in abs at 1.
+      rewrite (dist_sym a b) in abs.
+      rewrite abs in h_segment_a_b_p_.
+      assert (0 <= dist a b)%R by now apply dist_nonneg.
+      assert (0 <= dist p b)%R by now apply dist_nonneg.
+      assert (dist p b = 0)%R as h_eq_dist_IZR_.
+      { foldR2.
+        remember (dist p b) as x1.
+        remember (dist a b) as x2.
+        foldR2.
+        lra. }
+      apply dist_defined in h_eq_dist_IZR_.
+      contradiction. }
+      
+   assert (forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {ES:EuclideanSpace T} (a b p : T) (L:@line T),
+                a =/= b ->
+                line_contains L a ->
+                line_contains L b ->
+                line_contains L p ->
+                (strict_segment p a b \/ strict_segment p b a) \/ segment a b p) as h_lemma.
+    { intros T S0 EQ0 VS ES a b p L h_nequiv_a_b_ h_line_contains_L_a_ h_line_contains_L_b_ h_line_contains_L_p_ .
+      specialize segment_line with (L:=L) as h_all_iff_ .  
+      specialize h_all_iff_ with (1:=h_line_contains_L_a_) (2:=h_line_contains_L_b_) (3:=h_line_contains_L_p_) as h_abp.
+      specialize h_all_iff_ with (2:=h_line_contains_L_a_) (3:=h_line_contains_L_b_) (1:=h_line_contains_L_p_) as h_pab.
+      specialize h_all_iff_ with (3:=h_line_contains_L_a_) (2:=h_line_contains_L_b_) (1:=h_line_contains_L_p_) as h_pba.
+      foldR2.
+      destruct (segment_dec a b p).
+      { right.
+        assumption. }
+      destruct (segment_dec p a b).
+      { left.
+        left.
+        red.
+        repeat split;auto.
+        - intro abs.
+          rewrite abs in n1.
+          apply n1.
+          apply segment_end. }
+      destruct (segment_dec p b a).
+      { left.
+        right.
+        red.
+        repeat split;auto.
+        - intro abs.
+          rewrite abs in n1.
+          apply n1.
+          apply segment_start. }
+      exfalso.
+      rewrite h_abp in n1.
+      rewrite h_pab in n2.
+      rewrite h_pba in n3.
+      lra. }
+   
+    specialize (@h_lemma _ _ _ LocVS LocES) with (L:=Lx)(p := (@middle _ _ _ LocVS r3 r4)) (a:=r) (b:=r0) (1:=h_nequiv_r_r0_).
+    destruct (h_lemma) as [[ h_strict_segment_middle_r_r0_  | h_strict_segment_middle_r0_r_]   | h_segment_r_r0_middle_ ].
+    all:cycle 3.
+    - red in h_strict_segment_middle_r_r0_.
+      destruct h_strict_segment_middle_r_r0_ as [ h_segment_middle_r_r0_ [h_nequiv_r0_middle_ h_nequiv_r0_r_ ]] .
+      (* decompose [and] h_strict_segment_middle_r_r0_ /n h_segment_middle_r_r0_ h_nequiv_r0_middle_ h_nequiv_r0_r_ . *)
+      specialize h_lemma2 with (1:=h_segment_middle_r_r0_) (2:=h_nequiv_r0_r_).
+      rewrite segment_sym in h_lemma2.
+      apply segment_decb_true in h_segment_middle_r_r0_.
+      apply segment_decb_false in h_lemma2.
+      foldR2.
+      destruct h as [[h_equiv_r_r1_ h_equiv_r0_r2_]|[h_equiv_r_r2_ h_equiv_r0_r1_]].
+      + repeat rewrite <- h_equiv_r_r1_.
+        repeat rewrite <-h_equiv_r0_r2_.
+        rewrite h_segment_middle_r_r0_.
+        reflexivity.
+      + 
+        repeat rewrite <- h_equiv_r0_r1_.
+        repeat rewrite <-h_equiv_r_r2_.
+        rewrite h_segment_middle_r_r0_.
+        rewrite h_lemma2.
+        reflexivity.
+    - red in h_strict_segment_middle_r0_r_.
+      destruct h_strict_segment_middle_r0_r_ as [ h_segment_middle_r0_r_ [ h_nequiv_r_middle_ h_nequiv_r_r0_0 ]] .
+      specialize h_lemma2 with (1:=h_segment_middle_r0_r_)(2:=h_nequiv_r_r0_0).
+      rewrite segment_sym in h_lemma2.
+      apply segment_decb_true in h_segment_middle_r0_r_.
+      apply segment_decb_false in h_lemma2.
+      foldR2.
+      destruct h as [[h_equiv_r_r1_ h_equiv_r0_r2_]|[h_equiv_r_r2_ h_equiv_r0_r1_]].
+      + repeat rewrite <- h_equiv_r_r1_.
+        repeat rewrite <-h_equiv_r0_r2_.
+        rewrite h_segment_middle_r0_r_.
+        reflexivity.
+      + 
+        repeat rewrite <- h_equiv_r0_r1_.
+        repeat rewrite <-h_equiv_r_r2_.
+        rewrite h_segment_middle_r0_r_.
+        rewrite h_lemma2.
+        reflexivity.
+    - specialize h_lemma2 with (1:=h_segment_r_r0_middle_).
+      assert (middle r3 r4 =/= r0) as h_nequiv_middle_r0_ .
+      { foldR2.
+        apply  segment_decb_false in h_equiv_segment_decb_segment_decb_.
+        intro abs.
+        rewrite abs in h_equiv_segment_decb_segment_decb_.
+        apply h_equiv_segment_decb_segment_decb_.
+        apply segment_end. }
+      specialize h_lemma2 with (1:=h_nequiv_middle_r0_).
+      apply segment_decb_true in h_segment_r_r0_middle_.
+      apply segment_decb_false in h_lemma2.
+      foldR2.
+      destruct h as [ [h_equiv_r_r1_ h_equiv_r0_r2_]  | [h_equiv_r_r2_ h_equiv_r0_r1_] ].
+      + repeat rewrite <- h_equiv_r_r1_.
+        repeat rewrite <-h_equiv_r0_r2_.
+        rewrite h_lemma2.
+        reflexivity.
+      + repeat rewrite <- h_equiv_r0_r1_.
+        repeat rewrite <-h_equiv_r_r2_.
+        rewrite h_lemma2.
+        reflexivity.
+    - eapply weber_aligned with (ps:=multi_support x);auto.
+      specialize weber_calc_seg_correct with (ps:=multi_support x) as h_let.
+      rewrite hcalcx in h_let.
+      apply h_let;auto.
+      apply segment_start.
+    - eapply weber_aligned with (ps:=multi_support x);auto.
+      specialize weber_calc_seg_correct with (ps:=multi_support x) as h_let.
+      rewrite hcalcx in h_let.
+      apply h_let;auto.
+      apply segment_end.
+    - apply line_contains_middle.
+      + rewrite Heqr3.
+        red.
+        rewrite line_P_iP_min.
+        reflexivity.
+      + rewrite Heqr4.
+        red.
+        rewrite line_P_iP_max.
+        reflexivity.
+Qed.
+
+Definition gatherW : robogram := {| pgm := gatherW_pgm |}.
+
+
+
+(* Simplify the [round] function and express it in the global frame of reference.
+ * All the proofs below use this simplified version. *)
+Lemma round_simplify da config : 
+  similarity_da_prop da ->
+  ~invalid config ->
+  exists (r : ident -> ratio),
+  round gatherW da config == 
+  fun id => if activate da id then  
+              let target := 
+                let '(w1, w2) := weber_calc_seg (pos_list config) in 
+                let L := line_calc (pos_list config) in 
+                let r1 := L^-P (L^min (pos_list config)) in 
+                let r2 := L^-P (L^max (pos_list config)) in 
+                let org := get_location (config id) in
+                (* phase 1 *)
+                if w1 =?= w2 then w1 else 
+                match robot_with_mult (Nat.div2 (n+1)) (pos_list config) with
+                | Some t =>
+                  (* phase 2 *)
+                  if existsb (strict_segment_decb org t) (pos_list config) 
+                  then org else t
+                | None =>
+                  if segment_decb w1 w2 (middle r1 r2) then
+                    (* phase 3 *)
+                    if (w1 ==b org) || (w2 ==b org) then (middle r1 r2) else org
+                  else if segment_decb (middle r1 r2) w1 w2 then 
+                    (* phase 4 *)
+                    if w1 ==b org then w2 else org 
+                  else if segment_decb (middle r1 r2) w2 w1 then
+                    (* phase 4 *)
+                    if w2 ==b org then w1 else org
+                  else (* this should not happen *) 
+                    org                    
+                end   
+              in (get_location (config id), target, ratio_0)
+            else inactive config id (r id).
+Proof using lt_0n delta_g0. 
+intros Hsim Hvalid. eexists ?[r]. intros id. unfold round. 
+destruct_match ; [|reflexivity]. rename H into Hactivate.
+destruct_match ; [|byz_exfalso]. rename H into Hgood.
+rewrite (lift_update_swap da config _ g). 
+pose (f := existT precondition
+  (change_frame da config g)
+  (precondition_satisfied da config g)). 
+pose (f_inv := existT precondition
+  ((change_frame da config g) ⁻¹)
+  (precondition_satisfied_inv da config g)).
+pose (obs := obs_from_config (map_config (lift f) config) (lift f (config (Good g)))).
+change_LHS (update
+  (map_config (lift f_inv) (map_config (lift f) config)) 
+  g
+  Similarity.id
+  (frame_choice_bijection (change_frame da config g ⁻¹) (gatherW_pgm obs))
+  (choose_update da (map_config (lift f) config) g (gatherW_pgm obs))).
+assert (Hcancel : map_config (lift f_inv) (map_config (lift f) config) == config).
+{ intros idt. cbn -[equiv]. destruct (config idt) as [[start dest] r]. now rewrite 2 Bijection.retraction_section. }
+rewrite Hcancel. clear Hcancel.
+assert (Proper (equiv ==> equiv) (projT1 f)) as f_compat.
+{ unfold f ; cbn -[equiv]. intros x y Hxy ; now rewrite Hxy. }
+unfold gatherW_pgm at 1. pose (sim := change_frame da config g). foldR2. fold sim.
+cbn -[segment_decb strict_segment_decb middle robot_with_mult Nat.div sim
+  config_list get_location multi_support equiv_dec line_proj line_proj_inv cardinal] ; repeat split ; [].
+apply (Similarity.injective sim). rewrite Bijection.section_retraction. foldR2.
+
+
+
+
+specialize weber_calc_seg_correct with (1:=(pos_list_non_nil config)) as Hw.
+destruct (weber_calc_seg (pos_list config)) as [w1 w2].
+foldR2.
+assert (Hw' := @weber_calc_seg_correct (multi_support obs)).
+destruct (weber_calc_seg (multi_support obs)) as [w1' w2']. 
+assert (Hw_perm : perm_pairA equiv (sim w1, sim w2) (w1', w2')).
+{
+  apply segment_eq_iff. intros x. rewrite <-(Bijection.section_retraction sim x).
+  rewrite <-R2segment_similarity.
+  rewrite <-Hw.
+  foldR2.
+  rewrite <-Hw'.
+  change obs with (!! (map_config (lift f) config)). rewrite multi_support_map by auto.
+  cbn [projT1 f]. fold sim. now rewrite <-weber_similarity. 
+  apply multi_support_non_nil2.
+}
+clear Hw'.
+
+pose (L := line_calc (pos_list config)). pose (L' := line_calc (multi_support obs)).
+pose (r1 := L^-P (L^min (pos_list config))). pose (r2 := L^-P (L^max (pos_list config))).
+pose (r1' := L'^-P (L'^min (multi_support obs))). pose (r2' := L'^-P (L'^max (multi_support obs))).
+foldR2. fold L L' r1 r2 r1' r2'.  
+
+assert (Hmiddle : aligned (pos_list config) -> middle r1' r2' = sim (middle r1 r2)).
+{ 
+  intros Halign. 
+  assert (Halign' : aligned (multi_support obs)).
+  { change obs with (!! (map_config (lift f) config)). rewrite multi_support_map by auto.
+    apply R2aligned_similarity, Halign. }
+  rewrite line_calc_spec in Halign, Halign'. foldR2. fold L L' in Halign, Halign'.
+  unfold r1', r2'. change obs with (!! (map_config (lift f) config)) in Halign' |- *.
+  rewrite multi_support_map in Halign' |- * by auto. change (projT1 f) with (Bijection.section sim) in Halign' |- *.
+  apply line_middle_sim ; auto. 
+  intros Hnil. apply (f_equal (@length R2)) in Hnil. rewrite pos_list_length in Hnil. cbn in Hnil. lia.
+}
+
+assert ( invalid_obs obs -> invalid_obs (obs_from_config config (config (Good g)))) as h_valid_sim.
+{ fold sim in f,f_inv.
+  intro abs.
+  unfold obs in abs.
+  unfold f in abs.
+  rewrite <- obs_from_config_map in abs.
+
+  red in abs|-*.
+  destruct abs as [x [y [hneq_xy [heq_card [h_multx h_multy]]]]].
+  exists (sim ⁻¹ x), (sim ⁻¹ y).
+  repeat split.
+  * intro abs.
+    apply hneq_xy.
+    apply (Bijection.injective (sim ⁻¹)).
+    assumption.
+  * rewrite n_cardinal.
+    lia.
+  * rewrite n_cardinal.
+    assert (x == (sim (sim ⁻¹ x))) as h_eq'.
+    { rewrite (Bijection.section_retraction sim x).
+      auto. }
+    rewrite h_eq' in h_multx.
+    rewrite map_injective_spec in h_multx;try typeclasses eauto.
+    -- rewrite h_multx.
+       rewrite map_cardinal;try typeclasses eauto.
+       rewrite n_cardinal.
+       reflexivity.
+    -- apply Bijection.injective.
+  * rewrite n_cardinal.
+    assert (y == (sim (sim ⁻¹ y))) as h_eq'.
+    { rewrite (Bijection.section_retraction sim y).
+      auto. }
+    rewrite h_eq' in h_multy.
+    rewrite map_injective_spec in h_multy;try typeclasses eauto.
+    -- rewrite h_multy.
+       rewrite map_cardinal;try typeclasses eauto.
+       rewrite n_cardinal.
+       reflexivity.
+    -- apply Bijection.injective.
+  * apply f_compat. }
+
+assert (Hmult : robot_with_mult (Nat.div2 (cardinal obs+1)) (multi_support obs) == option_map sim (robot_with_mult (Nat.div2 (n+1)) (pos_list config))).
+{ rewrite robot_with_mult_map;try typeclasses eauto.
+  - setoid_rewrite <-(n_cardinal config (config (Good g))).
+    rewrite (@robot_with_mult_unique_obs obs (List.map sim (pos_list config)));auto.
+    + unfold obs.
+      repeat rewrite cardinal_obs_from_config.
+      reflexivity.
+    + clear L L' r1 r2 r1' r2' Hmiddle.
+      clear Hw_perm w2' w1' Hw w2 w1.
+      fold sim in f,f_inv.
+      rewrite (invalid_config_obs config (config (Good g))) in Hvalid.
+      intro abs.
+      apply Hvalid.
+      now apply h_valid_sim.
+    + unfold obs, sim.
+      specialize (multi_support_map f config) as h.
+      unfold f in h at 3.
+      cbn [projT1] in h.
+      rewrite h;try typeclasses eauto.
+      reflexivity.
+  - unfold sim.
+    apply Bijection.injective. }
+
+assert (Hcenter : 0%VS == sim (get_location (config (Good g)))).
+
+{ rewrite <-Hsim. cbn. unfold sim. now rewrite Bijection.section_retraction. }
+
+assert (Hseg : forall t, existsb (strict_segment_decb 0%VS (sim t)) (multi_support obs) =
+  existsb (strict_segment_decb (get_location (config (Good g))) t) (pos_list config)).
+{
+  intros t. change obs with (!! (map_config (lift f) config)).
+  rewrite multi_support_map by auto. setoid_rewrite <-List.map_id at 3. 
+  rewrite 2 existsb_map. f_equiv. intros ? ? h. cbn -[get_location sim]. foldR2. fold sim.
+  rewrite strict_segment_decb_eq_iff. 
+  now rewrite Hcenter, <-R2strict_segment_similarity, h.
+}
+
+assert (Hseg_sim : forall x y z, segment_decb (sim x) (sim y) (sim z) = segment_decb x y z).
+{ intros x y z. rewrite segment_decb_eq_iff. now rewrite <-R2segment_similarity. }
+     
+assert (Heq_sim : forall x y, (sim x ==b sim y) = (x ==b y)).
+{ intros x y. unfold equiv_decb. case ifP_sumbool ; case ifP_sumbool ; try now tauto.
+  + intros HNeq Heq. apply Similarity.injective in Heq. intuition.
+  + intros ->. intuition. } 
+      
+assert (Hweb_align : w1 =/= w2 -> aligned (pos_list config)).
+{
+  intros Hw12. 
+  case (aligned_dec (pos_list config)) as [Halign | HNalign] ; [auto|].
+  apply weber_Naligned_unique with (w := w1) in HNalign.
+  + exfalso. apply Hw12. symmetry. apply HNalign.
+    rewrite Hw. apply segment_end.
+  + rewrite Hw. apply segment_start.
+}
+
+destruct (obs_invalid_dec obs) at 1.
+{ exfalso.
+  apply Hvalid.
+  setoid_rewrite -> (invalid_config_obs config (config (Good g))).
+  now apply h_valid_sim. }
+destruct (length (multi_support obs)) eqn:heq_supp.
+{ exfalso.
+  rewrite multi_set_length_cardinal in heq_supp.
+  unfold obs in heq_supp.
+  rewrite n_cardinal in heq_supp.
+  lia. }
+cbn [ equiv_dec nat_EqDec Nat.eq_dec nat_rec nat_rect].
+
+case Hw_perm as [[<- <-] | [<- <-]]. 
++ case ifP_sumbool; case ifP_sumbool ; intros Hw' Hw_sim.
+  - reflexivity.
+  - exfalso. apply Hw'. now apply (Similarity.injective sim).
+  - exfalso. apply Hw_sim. now f_equal.
+  - foldR2.
+
+    destruct (robot_with_mult (Nat.div2 (cardinal obs+1)) (multi_support obs)) as [t|];
+    destruct (robot_with_mult (Nat.div2 (n+1)) (pos_list config)) as [t'|];
+      cbn [option_map] in Hmult ; try now tauto.
+    cbn [option_map] in Hmult.
+    * rewrite Hmult, Hseg, Hcenter. now destruct_match.
+    * repeat rewrite Hmiddle by auto. repeat rewrite Hseg_sim.
+      repeat rewrite Hcenter. repeat rewrite Heq_sim.
+      foldR2. repeat destruct_match ; reflexivity.
++ case ifP_sumbool ; case ifP_sumbool ; intros Hw' Hw_sim.
+  - exact Hw_sim. 
+  - exfalso. apply Hw'. now apply (Similarity.injective sim).
+  - exfalso. apply Hw_sim. now f_equal.
+  - foldR2. 
+    destruct (robot_with_mult (Nat.div2 (cardinal obs+1)) (multi_support obs)) as [t|] ;
+    destruct (robot_with_mult (Nat.div2 (n+1)) (pos_list config)) as [t'|] ;
+      cbn in Hmult ; try now tauto.
+    * rewrite Hmult, Hseg, Hcenter. now destruct_match.
+    * repeat rewrite Hmiddle by auto. repeat rewrite Hseg_sim.
+      repeat rewrite Hcenter. repeat rewrite Heq_sim.
+      rewrite segment_decb_sym. foldR2. 
+      case ifP_bool ; [rewrite segment_decb_true | rewrite segment_decb_false] ; intros Hm1.
+      ++rewrite orb_comm. case ifP_bool ; reflexivity.
+      ++case ifP_bool ; [rewrite segment_decb_true | rewrite segment_decb_false] ; intros Hm2.
+        --assert (Hm3 : ~segment (middle r1 r2) w1 w2).
+          { 
+            intros Hm3. cut (w1 = w2). { intuition. }
+            assert (H : forall x, segment (middle r1 r2) w1 x <-> segment (middle r1 r2) w2 x).
+            { intros x. split ; revert x; rewrite segment_incl_iff; split;try apply segment_start;auto. }
+            rewrite segment_eq_iff in H. case H as [[H1 H2] | [<- ->]] ; auto.
+          }
+          rewrite <-segment_decb_false in Hm3. foldR2. rewrite Hm3.
+          destruct_match ; reflexivity.
+        --repeat destruct_match ; reflexivity.
+Qed.
+
+(* This is the property : all robots are looping. 
+ * This is what should be verified in the initial configuration. *)
+Definition config_stay (c : configuration) : Prop := 
+  forall id, get_dest (c id) == get_start (c id).
+
+Local Instance config_stay_compat : Proper (equiv ==> iff) config_stay.
+Proof using . intros ? ? H. unfold config_stay. now setoid_rewrite H. Qed.
+
+(* This is the property : all robots are looping OR 
+ * go towards point p. *)
+Definition config_stg (c : configuration) p : Prop := 
+  forall id, get_dest (c id) == get_start (c id) \/ get_dest (c id) == p.
+
+Local Instance config_stg_compat : Proper (equiv ==> equiv ==> iff) config_stg.
+Proof using . intros ? ? H1 ? ? H2. unfold config_stg. setoid_rewrite H1. setoid_rewrite H2. reflexivity. Qed. 
+
+(* This the property : robots on the endpoints of [L] stay where they are.
+ * RMK : this does NOT mean they are looping (their start could be different from their destination). *)
+Definition endpoints_stay c L := 
+  (forall id, get_location (c id) == L^-P (L^min (pos_list c)) -> get_dest (c id) == L^-P (L^min (pos_list c))) /\
+  (forall id, get_location (c id) == L^-P (L^max (pos_list c)) -> get_dest (c id) == L^-P (L^max (pos_list c))).
+
+Lemma config_stay_or_Nloop c : 
+  config_stay c \/ exists id, get_dest (c id) =/= get_start (c id).
+Proof using . 
+case (Sumbool.sumbool_of_bool (forallb (fun id => get_dest (c id) ==b get_start (c id)) names)).
++ intros HF. left. rewrite forallb_forall in HF. setoid_rewrite eqb_true_iff in HF.
+  intros id. apply HF, In_names.
++ intros HF. right. apply (f_equal negb) in HF. rewrite existsb_forallb, existsb_exists in HF.
+  destruct HF as [id [_ HNloop]]. rewrite negb_true_iff, eqb_false_iff in HNloop. eauto.
+Qed.
+
+Lemma config_stay_impl_config_stg c : 
+  config_stay c -> forall x, config_stg c x.
+Proof using . intros Hstay x id. now left. Qed. 
+
+Lemma config_stay_impl_endpoints_stay c : 
+  config_stay c -> forall L, endpoints_stay c L.
+Proof using . 
+intros Hstay L. split ; intros id <- ; specialize (Hstay id) ; 
+  destruct (c id) as [[s d] ri] ; cbn -[equiv straight_path] in Hstay |- * ; 
+  now rewrite Hstay, straight_path_same.
+Qed.
+
+Lemma config_stay_location_same c da : 
+  similarity_da_prop da -> 
+  ~invalid c -> 
+  config_stay c -> 
+  forall id, get_location (round gatherW da c id) == get_location (c id).
+Proof using lt_0n delta_g0.
+intros Hsim Hvalid Hstay id. destruct (round_simplify Hsim Hvalid) as [r Hround].
+rewrite (Hround id). clear Hround. destruct_match.
++ cbn zeta. now rewrite get_location_ratio_0.
++ cbn -[get_location equiv]. specialize (Hstay id). destruct (c id) as [[s d] ri].
+  cbn -[equiv straight_path] in Hstay |- *. now rewrite Hstay, 2 straight_path_same.
+Qed.  
+
+Lemma config_stay_pos_list_same c da : 
+  similarity_da_prop da -> 
+  ~invalid c -> 
+  config_stay c -> 
+  eqlistA equiv (pos_list (round gatherW da c)) (pos_list c).
+Proof using lt_0n delta_g0. 
+intros Hsim Hvalid Hstay.
+unfold pos_list. rewrite 2 config_list_spec.
+induction names as [|id l IH].
++ cbn. now left.
++ cbn [List.map]. right.
+  - apply config_stay_location_same ; assumption.
+  - apply IH.
+Qed.    
+
+(* Each phase of the algorithm is characterized by an invariant.
+ * We will proove lemmas of the form (inv2 config -> inv1 config \/ inv2 config).
+ * The invariants are pairwise disjoint, although this shouldn't be necessary.
+ * Additionally every initial configuration verifies at least one invariant. *)
+
+Definition inv1 c w : Prop := 
+  (* All robots are moving towards [w]. *)
+  config_stg c w /\ 
+  (* [w] is the unique weber point. *)
+  OnlyWeber (pos_list c) w.
+
+Definition inv2 c L : Prop :=   
+  let ps := pos_list c in 
+  (* The robots are aligned. *)
+  aligned_on L ps /\
+  (* The weber point isn't unique. *)
+  (let (w1, w2) := weber_calc_seg ps in w1 =/= w2) /\
+  (* There is a point with multiplicity (n+1)/2. 
+  * We can assume it is an endpoint (if it isn't, then we are in phase 1).
+  * With the right choice of [L], it is the minimum point. *)
+  (!! c)[L^-P (L^min ps)] = Nat.div2 (n+1) /\
+  (* The maximum has multiplicity less than [(n+1)/2], 
+   * i.e. the configuration is valid. *)
+  (!! c)[L^-P (L^max ps)] < Nat.div2 (n+1) /\
+  (* All robots are moving towards the minimum. *)
+  config_stg c (L^-P (L^min ps)) /\ 
+  (* The maximum isn't moving (this is implied by the fact
+   * that we perform a safe move). *)
+  endpoints_stay c L.
+
+Definition inv3 c L w1 w2 : Prop := 
+  let ps := pos_list c in
+  (* The robots are aligned. *)
+  aligned_on L ps /\
+  (* The weber segment is [w1, w2]. *)
+  (forall w, Weber ps w <-> segment w1 w2 w) /\
+  (* The weber point is not unique *)
+  w1 =/= w2 /\
+  (* The middle point is inside the weber segment *)
+  (L^P w1 <= (L^min ps + L^max ps) / 2 <= L^P w2)%R /\
+  (* All robots are moving towards the middle point *)
+  config_stg c (middle (L^-P (L^min ps)) (L^-P (L^max ps))) /\ 
+  (* The endpoints are not moving *)
+  endpoints_stay c L /\
+  (* The endpoints have multiplicity less than [(n+1)/2] *)
+  (!! c)[L^-P (L^min ps)] < Nat.div2 (n+1) /\
+  (!! c)[L^-P (L^max ps)] < Nat.div2 (n+1).
+    
+Definition inv4 c L w1 w2 : Prop :=
+  let ps := pos_list c in 
+  (* The robots are aligned. *)
+  aligned_on L ps /\  
+  (* The weber segment is [w1, w2]. *)
+  (forall w, Weber ps w <-> segment w1 w2 w) /\
+  (* The middle point is to the left of [w1] (we can enforce this 
+    * with the right choice for [L]). *)
+  ((L^min ps + L^max ps) / 2 < L^P w1 < L^P w2)%R /\
+  (* All robots are moving towards [w1]. *)
+  config_stg c w1 /\ 
+  (* The endpoints are not moving. *)
+  endpoints_stay c L /\
+  (* The endpoints have multiplicity less than [(n+1)/2] *)
+  (!! c)[L^-P (L^min ps)] < Nat.div2 (n+1) /\
+  (!! c)[L^-P (L^max ps)] < Nat.div2 (n+1).
+
+Lemma inv234_endpoints_distinct c L : 
+  inv2 c L \/ (exists w1 w2, inv3 c L w1 w2) \/ (exists w1 w2, inv4 c L w1 w2) -> 
+  L^-P (L^min (pos_list c)) =/= L^-P (L^max (pos_list c)).
+Proof using lt_0n. 
+intros Hinv234 Er12.
+assert (H : exists w1 w2,
+  aligned_on L (pos_list c) /\
+  (forall w, Weber (pos_list c) w <-> segment w1 w2 w) /\
+  w1 =/= w2).
+{ case Hinv234 as [Hinv2 | [[w1 [w2 Hinv3]] | [w1 [w2 Hinv4]]]].
+  + destruct Hinv2 as [Halign [NEw12 _]].
+    assert (Hweb := @weber_calc_seg_correct (pos_list c)).
+    destruct (weber_calc_seg (pos_list c)) as [w1 w2]. 
+    exists w1, w2.
+    split;auto.
+    split;auto.
+    intros w. 
+    apply Hweb.
+    apply pos_list_non_nil.
+  + destruct Hinv3 as [Halign [Hweb [NEw12 _]]]. now exists w1, w2.
+  + destruct Hinv4 as [Halign [Hweb [Hseg _]]]. exists w1, w2.
+    split ; [assumption | split ; [assumption|]].
+    intros ->. lra. }
+destruct H as [w1 [w2 [Halign [Hweb NEw12]]]].
+apply NEw12.
+cut (L^P w1 == L^P w2).
+{ intros HeqP. apply (f_equal (L^-P)) in HeqP. 
+  rewrite 2 Halign in HeqP.
+  - assumption.
+  - apply (weber_segment_InA (pos_list_non_nil _) NEw12 Hweb);auto.
+  - apply (weber_segment_InA (pos_list_non_nil _) NEw12 Hweb);auto. }
+assert (Hmin_max : L^max (pos_list c) == L^min (pos_list c)).
+{ rewrite <-line_P_iP_min, <-line_P_iP_max. foldR2. now rewrite Er12. }
+transitivity (L^min (pos_list c)).
++ assert (Hbounds := line_bounds L (pos_list c) w1).
+  feed Hbounds. { apply (weber_segment_InA (pos_list_non_nil _) NEw12 Hweb). }
+  rewrite Hmin_max in Hbounds. now apply Rle_antisym.
++ assert (Hbounds := line_bounds L (pos_list c) w2).
+  feed Hbounds. { apply (weber_segment_InA (pos_list_non_nil _) NEw12 Hweb). }
+  rewrite Hmin_max in Hbounds. now apply Rle_antisym.
+Qed.
+
+Lemma inv234_not_gathered c L x : 
+  inv2 c L \/ (exists w1 w2, inv3 c L w1 w2) \/ (exists w1 w2, inv4 c L w1 w2) ->
+  ~gathered_at x c.
+Proof using lt_0n.
+intros Hinv Hgather.
+assert (Halign : aligned_on L (pos_list c)).
+{ destruct Hinv as [Hinv | [[w1 [w2 Hinv]] | [w1 [w2 Hinv]]]] ; apply Hinv. }
+apply inv234_endpoints_distinct in Hinv. apply Hinv.
+assert (Hr1 : exists id, L^-P (L^min (pos_list c)) == get_location (c id)).
+{ rewrite <-pos_list_InA. apply line_iP_min_InA ; [|assumption].
+  rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hr2 : exists id, L^-P (L^max (pos_list c)) == get_location (c id)).
+{ rewrite <-pos_list_InA. apply line_iP_max_InA ; [|assumption].
+  rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+destruct Hr1 as [id1 ->]. destruct Hr2 as [id2 ->].
+transitivity x.
++ rewrite <-good_unpack_good. apply Hgather.
++ symmetry. rewrite <-good_unpack_good. apply Hgather.
+Qed. 
+
+
+(* In phases 3 and 4, the endpoints aren't weber points 
+ * (because they have multiplicity < (n+1)/2). *)
+Lemma inv34_endpoints_not_weber c L w1 w2 : 
+  line_dir L <> 0%VS ->
+  let ps := pos_list c in 
+  inv3 c L w1 w2 \/ inv4 c L w1 w2 -> 
+  ~Weber ps (L^-P (L^min ps)) /\ ~Weber ps (L^-P (L^max ps)).
+Proof using lt_0n. 
+specialize (pos_list_non_nil c) as h_ps_nonnil.
+cbn zeta. intros Hnonnil Hinv34.
+pose (ps := pos_list c). pose (r1 := L^-P (L^min ps)). pose (r2 := L^-P (L^max ps)).
+assert (H :
+  aligned_on L ps /\ 
+  (forall w, Weber ps w <-> segment w1 w2 w) /\
+  w1 =/= w2 /\ 
+  (!! c)[L^-P (L^min ps)] < Nat.div2 (n+1) /\
+  (!! c)[L^-P (L^max ps)] < Nat.div2 (n+1)).
+{ case Hinv34 as [Hinv | Hinv] ; repeat split ; try apply Hinv. 
+  destruct Hinv as [_ [_ [Hseg _]]]. intros ->. lra. }
+destruct H as [Halign [Hweb [NEw12 [Hr1_mult Hr2_mult]]]].
+fold ps r1 r2 in Hr1_mult, Hr2_mult |- *.
+ split.
++ intros Hr1_web. 
+  assert (aligned_on L (r1::ps)) as Halign'.
+  { specialize (weber_aligned) as h.
+    apply aligned_on_cons_iff.
+    split;auto.
+    eapply h;eauto. }
+  rewrite (weber_aligned_spec_weak Hnonnil Halign') in Hr1_web. 
+  assert (Hperm := line_left_on_right_partition L r1 ps).
+  unfold r1 at 1 in Hperm. unfold r1 at 1 in Hr1_web. rewrite line_left_iP_min in Hperm, Hr1_web.
+  cbn [app] in Hperm. cbn [length INR] in Hr1_web. 
+  rewrite Rminus_0_l, Rabs_Ropp, <-Rle_abs in Hr1_web. apply INR_le in Hr1_web.
+  apply PermutationA_length in Hperm. unfold ps at 1 in Hperm. 
+  rewrite app_length, pos_list_length in Hperm.
+  assert (Hon : length (L^on r1 ps) < Nat.div2 (n+1)).
+  { rewrite line_on_length_aligned.
+    + unfold ps. now rewrite <-multiplicity_countA.
+    + rewrite aligned_on_cons_iff. split ; [apply Halign, line_iP_min_InA |] ; try assumption. }
+  assert (Hright : length (L^right r1 ps) < Nat.div2 (n+1)).
+  { eapply Nat.le_lt_trans ; eauto. }
+  case (Nat.Even_or_Odd n) as [[k Hk] | [k Hk]].
+  - rewrite Hk in Hon, Hright, Hperm.
+    rewrite Nat.add_1_r, Nat.div2_succ_double in Hon, Hright. lia.
+  - rewrite Hk in Hon, Hright, Hperm.
+    assert (Hrewrite : 2 * k + 1 + 1 = 2 * (k + 1)). { lia. }
+    rewrite Hrewrite, Nat.div2_double in Hon, Hright. lia.
++ intros Hr2_web. 
+  assert (aligned_on L (r2::ps)) as Halign'.
+  { specialize (weber_aligned) as h.
+    apply aligned_on_cons_iff.
+    split;auto.
+    eapply h;eauto. }
+  rewrite (weber_aligned_spec_weak Hnonnil Halign') in Hr2_web.
+  assert (Hperm := line_left_on_right_partition L r2 ps).
+  unfold r2 at 3 in Hperm. unfold r2 at 2 in Hr2_web. rewrite line_right_iP_max in Hperm, Hr2_web.
+  rewrite app_nil_r in Hperm. cbn [length INR] in Hr2_web. 
+  rewrite Rminus_0_r, <-Rle_abs in Hr2_web. apply INR_le in Hr2_web.
+  apply PermutationA_length in Hperm. unfold ps at 1 in Hperm. 
+  rewrite app_length, pos_list_length in Hperm.
+  assert (Hon : length (L^on r2 ps) < Nat.div2 (n+1)).
+  { rewrite line_on_length_aligned.
+    + unfold ps. now rewrite <-multiplicity_countA.
+    + rewrite aligned_on_cons_iff. split ; [apply Halign, line_iP_max_InA |] ; try assumption. }
+  assert (Hleft : length (L^left r2 ps) < Nat.div2 (n+1)).
+  { eapply Nat.le_lt_trans ; eauto. }
+  case (Nat.Even_or_Odd n) as [[k Hk] | [k Hk]].
+  - rewrite Hk in Hon, Hperm, Hleft.
+    rewrite Nat.add_1_r, Nat.div2_succ_double in Hon, Hleft. lia.
+  - rewrite Hk in Hon, Hleft, Hperm.
+    assert (Hrewrite : 2 * k + 1 + 1 = 2 * (k + 1)). { lia. }
+    rewrite Hrewrite, Nat.div2_double in Hon, Hleft. lia.
+Qed.
+
+Lemma inv1_valid config w : inv1 config w -> ~invalid config.
+Proof using . 
+intros [_ Hweb] [x [y [NExy [Hx_mult Hy_mult]]]].
+apply NExy. transitivity w. 
++ apply Hweb. apply weber_majority_weak. 
+  rewrite <-multiplicity_countA, Hx_mult, pos_list_length, Nat.div2_div, ge_le_iff. 
+  reflexivity.
++ symmetry. apply Hweb. apply weber_majority_weak.
+  rewrite <-multiplicity_countA, Hy_mult, pos_list_length, Nat.div2_div, ge_le_iff.
+  reflexivity.
+Qed.
+
+Lemma inv2_valid config L : inv2 config L -> ~invalid config.
+Proof using lt_0n delta_g0.
+intros [Halign [_ [Hr1 [Hr2 _]]]] Hinvalid. 
+apply invalid_refine_endpoints with (L := L) in Hinvalid ; [|exact Halign].
+destruct Hinvalid as [_ [_ Hr2']]. rewrite Hr2' in Hr2. intuition.
+Qed.
+
+Lemma inv3_valid config L w1 w2 : inv3 config L w1 w2 -> ~invalid config.
+Proof using lt_0n delta_g0.
+intros [Halign [_ [_ [_ [_ [Hr1 Hr2]]]]]] Hinvalid.
+apply invalid_refine_endpoints with (L := L) in Hinvalid ; [|exact Halign].
+destruct Hinvalid as [_ [Hr1' Hr2']]. rewrite Hr2' in Hr2. intuition.
+Qed.
+
+Lemma inv4_valid config L w1 w2 : inv4 config L w1 w2 -> ~invalid config.
+Proof using lt_0n delta_g0.
+intros [Halign [_ [_ [_ [_ [Hr1 Hr2]]]]]] Hinvalid.
+apply invalid_refine_endpoints with (L := L) in Hinvalid ; [|exact Halign].
+destruct Hinvalid as [_ [Hr1' Hr2']]. rewrite Hr2' in Hr2. intuition.
+Qed.
+
+
+(* If a tower has at least half of the robots, and this tower is not an endpoint, 
+ * then the number of robots to the left and right of it is constrained. *)
+Lemma majority_left_right_bounds t c L : 
+  aligned_on L (pos_list c) ->
+  (!! c)[t] >= Nat.div2 (n+1) -> 
+  (L^min (pos_list c) < L^P t < L^max (pos_list c))%R -> 
+    1 <= length (L^left t (pos_list c)) < Nat.div2 (n+1) /\
+    1 <= length (L^right t (pos_list c)) < Nat.div2 (n+1).
+Proof using lt_0n.
+pose (r1 := L^-P (L^min (pos_list c))). pose (r2 := L^-P (L^max (pos_list c))).
+intros Halign Ht_mult Ht_bounds.  
+assert (Hps_nil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }  
+assert (Ht_InA : InA equiv t (pos_list c)).
+{ rewrite <-multiplicity_pos_InA, gt_lt_iff. rewrite ge_le_iff in Ht_mult.
+  eapply Nat.lt_le_trans ; [|exact Ht_mult]. apply Exp_prop.div2_not_R0. lia. }
+assert (Hlr1 : 1 <= length (L^left t (pos_list c)) /\ 1 <= length (L^right t (pos_list c))).
+{ 
+split. 
++ cut (length (L^left t (pos_list c)) <> 0). { lia. }
+  rewrite length_zero_iff_nil.
+  cut (InA equiv r1 (L^left t (pos_list c))). { intros Hin Hnil. rewrite Hnil in Hin. inv Hin. }
+  rewrite line_left_spec. split ; [unfold r1 ; now apply line_iP_min_InA|].
+  destruct (line_bounds L (pos_list c) t) as [[H1 | H2] _] ; [assumption| |].
+  - unfold r1. now rewrite line_P_iP_min.
+  - lra.
++ cut (length (L^right t (pos_list c)) <> 0). { lia. }
+  rewrite length_zero_iff_nil.
+  cut (InA equiv r2 (L^right t (pos_list c))). { intros Hin Hnil. rewrite Hnil in Hin. inv Hin. }
+  rewrite line_right_spec. split ; [unfold r2 ; now apply line_iP_max_InA|].
+  destruct (line_bounds L (pos_list c) t) as [[H1 | H2] _] ; [assumption| |].
+  - unfold r2. now rewrite line_P_iP_max.
+  - lra.
+}
+destruct Hlr1 as [Hl1 Hr1].
+assert (Ht : line_contains L t).
+{ apply Halign. rewrite <-multiplicity_pos_InA, gt_lt_iff.
+  eapply Nat.lt_le_trans ; [|exact Ht_mult]. apply Exp_prop.div2_not_R0 ; lia. }
+split.
++ split ; [assumption|].
+  apply Nat.nlt_ge. intros Hle. rewrite Nat.lt_succ_r in Hle.
+  rewrite multiplicity_countA, <-(line_on_length_aligned L) in Ht_mult by (now rewrite aligned_on_cons_iff).
+  assert (Hpart := line_left_on_right_partition L t (pos_list c)).
+  apply PermutationA_length in Hpart. rewrite pos_list_length, 2 app_length in Hpart.
+  generalize (div2_sum_p1_ge n). lia.
++ split ; [assumption|].
+  apply Nat.nlt_ge. intros Hle. rewrite Nat.lt_succ_r in Hle.
+  rewrite multiplicity_countA, <-(line_on_length_aligned L) in Ht_mult by now rewrite aligned_on_cons_iff.
+  assert (Hpart := line_left_on_right_partition L t (pos_list c)).
+  apply PermutationA_length in Hpart. rewrite pos_list_length, 2 app_length in Hpart.
+  generalize (div2_sum_p1_ge n). lia.
+Qed. 
+
+(* This is very important, and not very clear considering how complex
+ * the invariants are. *)
+Lemma inv_initial c : 
+  ~invalid c -> 
+  config_stay c -> 
+  (exists w, inv1 c w) \/ 
+  (exists L, inv2 c L) \/ 
+  (exists L w1 w2, inv3 c L w1 w2) \/ 
+  (exists L w1 w2, inv4 c L w1 w2).
+Proof using lt_0n. 
+intros Hvalid Hstay.
+specialize weber_calc_seg_correct_2 with (c:=c) as Hweb.
+pose (ps := pos_list c).
+fold ps in Hweb.
+assert( ps <> []) as h_psnonnil.
+{ apply pos_list_non_nil. }
+destruct (weber_calc_seg ps) as [w1 w2] eqn:Ew.
+specialize (Hweb _ _ eq_refl).
+case (w1 =?= w2) as [Ew12 | NEw12].
++ (* Phase 1. *)  
+  left. exists w1.
+  split ; [now apply config_stay_impl_config_stg|]. fold ps. split.
+  - apply Hweb. apply segment_start.
+  - intros x Hx. apply Hweb in Hx. rewrite <-Ew12, segment_same in Hx. exact Hx.
++ right. 
+  assert (Halign : aligned ps).
+  { case (aligned_dec (pos_list c)) as [Halign | HNalign] ; [assumption|].
+    exfalso. apply NEw12. apply weber_Naligned_unique with (w := w1) in HNalign.
+    2:{ fold ps ; rewrite Hweb.
+        - apply segment_start. }
+    symmetry. apply HNalign. fold ps. rewrite Hweb. apply segment_end. }
+  destruct Halign as [L Halign].
+  assert (NEr12 : L^-P (L^min ps) =/= L^-P (L^max ps)).
+  { intros Er12. apply (f_equal (L^P)) in Er12. rewrite line_P_iP_min, line_P_iP_max in Er12. apply NEw12. 
+    specialize (weber_segment_InA h_psnonnil NEw12 Hweb) as Hweb'. destruct Hweb' as [Hw1_InA Hw2_InA].
+    cut (L^P w1 == L^P w2).
+    { intros Heq. apply (f_equal (L^-P)) in Heq. rewrite 2 Halign in Heq by assumption. exact Heq. }
+    generalize (line_bounds L ps w1 Hw1_InA), (line_bounds L ps w2 Hw2_InA).
+    change (@equiv R _) with (@eq R). foldR2. lra. }
+  assert (Hdir : line_dir L =/= 0%VS).
+  { apply line_dir_nonnul with (x := L^-P (L^min ps)) (y := L^-P (L^max ps)).
+    + exact NEr12.
+    + unfold line_contains. now rewrite line_P_iP_min.
+    + unfold line_contains. now rewrite line_P_iP_max. }
+  case (robot_with_mult (Nat.div2 (n+1)) ps) as [t|] eqn:Et.
+  - (* Phase 2. *) 
+    left. unfold robot_with_mult in Et. apply find_some in Et. destruct Et as [Ht_InA Ht_mult].
+    rewrite eqb_true_iff in Ht_mult. apply (@In_InA _ equiv) in Ht_InA ; autoclass. 
+    assert (Ht_endpoints : t == L^-P (L^min ps) \/ t == (L^-P (L^max ps))).
+    { 
+      cut (~(L^min ps < L^P t < L^max ps)%R).
+      { intros Ht1. 
+        assert (Ht2 : L^P t = L^min ps \/ L^P t = L^max ps). 
+        { generalize (line_bounds L ps t Ht_InA). foldR2. lra. }
+        case Ht2 as [Ht3 | Ht3] ; apply (f_equal (L^-P)) in Ht3 ; 
+          rewrite Halign in Ht3 by assumption ; tauto. }
+      intros Ht.
+      cut (OnlyWeber ps t).
+      { intros Ht_web. apply NEw12. transitivity t.
+        + apply Ht_web. rewrite Hweb. apply segment_start.
+        + symmetry. apply Ht_web. rewrite Hweb. apply segment_end. }
+      rewrite weber_aligned_spec with (L:=L);auto.
+      2:{ apply aligned_on_cons_iff.
+          split.
+          - apply aligned_on_InA_contains with (ps:=ps);auto.
+          - exact Halign. }
+      assert (Hlr_bounds := @majority_left_right_bounds t c L Halign). 
+      feed_n 2 Hlr_bounds.
+      { rewrite multiplicity_countA. fold ps. rewrite Ht_mult. intuition. }
+      { exact Ht. } 
+      fold ps in Hlr_bounds. destruct Hlr_bounds as [[Hl1 Hl2] [Hr1 Hr2]]. 
+      apply le_INR in Hl1. apply lt_INR in Hl2. 
+      apply le_INR in Hr1. apply lt_INR in Hr2.
+      rewrite line_on_length_aligned, Ht_mult. 
+      + cbn [INR] in Hl1, Hr1. foldR2. unfold Rabs. destruct_match ; lra.
+      + rewrite aligned_on_cons_iff. split ; [|assumption].
+        apply Halign. unfold ps. rewrite <-multiplicity_pos_InA, multiplicity_countA. fold ps. rewrite Ht_mult.
+        apply Exp_prop.div2_not_R0. lia.
+    }
+    case Ht_endpoints as [Ht_min | Ht_max].
+    * exists L. split ; [|split ; [|split ; [|split ; [|split]]]] ; 
+        auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+      ++fold ps. now rewrite Ew.
+      ++fold ps. rewrite <-Ht_min, multiplicity_countA. fold ps. exact Ht_mult.
+      ++fold ps. rewrite <-Nat.nle_gt. intros Hr2_mult. apply Hvalid.
+        eapply two_majority_invalid with (x := t) (y := L^-P (L^max ps)).
+        --rewrite Ht_min. assumption.
+        --rewrite multiplicity_countA. fold ps. now rewrite Ht_mult.
+        --assumption.
+    * exists (line_reverse L). split ; [|split ; [|split ; [|split ; [|split]]]] ;
+        auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+      ++now rewrite <-aligned_on_reverse.
+      ++fold ps. now rewrite Ew.
+      ++fold ps. rewrite line_reverse_iP_min, <-Ht_max, multiplicity_countA. fold ps. exact Ht_mult.
+      ++fold ps. rewrite line_reverse_iP_max, <-Nat.nle_gt. intros Hr1_mult. apply Hvalid.
+        eapply two_majority_invalid with (x := t) (y := L^-P (L^min ps)).
+        --rewrite Ht_max. intuition.
+        --rewrite multiplicity_countA. fold ps. now rewrite Ht_mult.
+        --assumption.
+  - right.
+    assert (HNmaj : forall x, (!! c)[x] < Nat.div2 (n+1)).
+    { 
+      intros x. rewrite Nat.lt_nge. intros Hx_mult. apply Nat.lt_eq_cases in Hx_mult.
+      destruct Hx_mult as [Hx_mult | Hx_mult].
+      + assert (HOW : OnlyWeber ps x). 
+        { apply weber_majority. unfold ps. rewrite <-multiplicity_countA, pos_list_length. intuition. }
+        apply NEw12. transitivity x.
+        - apply HOW. rewrite Hweb. apply segment_start.
+        - symmetry. apply HOW. rewrite Hweb. apply segment_end.
+      + unfold robot_with_mult in Et. apply find_none with (x := x) in Et.
+        - rewrite eqb_false_iff in Et. apply Et. unfold ps. rewrite <-multiplicity_countA. intuition.
+        - rewrite <-(In_InA_is_leibniz equiv) by now cbn. unfold ps. rewrite <-multiplicity_pos_InA.
+          rewrite <-Hx_mult. apply Exp_prop.div2_not_R0. lia. 
+    }
+    case (segment_dec w1 w2 (middle (L^-P (L^min ps)) (L^-P (L^max ps)))) as [Hmiddle | Hmiddle].
+    * (* Phase 3. *)
+      left. case (Rle_dec (L^P w1) (L^P w2)) as [Cw12 | Cw12].
+      ++exists L, w1, w2. split ; [|split ; [|split ; [|split ; [|split ; [|split]]]]] ;
+          auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+        --rewrite (segment_line L) in Hmiddle. 
+          ** repeat rewrite (line_middle L ps) in Hmiddle at 1.
+             repeat rewrite line_P_iP in Hmiddle at 1 by assumption.
+            fold ps. foldR2. case Hmiddle; lra.
+          **apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          **apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          **apply line_contains_middle ; unfold line_contains ; [now rewrite line_P_iP_min | now rewrite line_P_iP_max].
+      ++exists L, w2, w1. split ; [|split ; [|split ; [|split ; [|split ; [|split]]]]] ;
+          auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+        --now setoid_rewrite segment_sym.
+        (* --intuition. *)
+        --rewrite (segment_line L) in Hmiddle. 
+          ** repeat rewrite (line_middle L ps) in Hmiddle at 1.
+             repeat rewrite line_P_iP in Hmiddle at 1 by assumption.
+            fold ps. foldR2. case Hmiddle ; lra.
+          **apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          **apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          **apply line_contains_middle ; unfold line_contains ; [now rewrite line_P_iP_min | now rewrite line_P_iP_max].
+    * (* Phase 4. *)
+      right.
+      assert (Cw12 : (L^P w1 < L^P w2 \/ L^P w2 < L^P w1)%R).
+      {
+        case (Rle_dec (L^P w1) (L^P w2)) as [[Hlt | Heq] | Hgt] ; [left | exfalso | right].
+        + assumption.
+        + apply NEw12. apply (f_equal (L^-P)) in Heq. 
+          rewrite 2 Halign in Heq by apply (weber_segment_InA h_psnonnil NEw12 Hweb). assumption.
+        + lra.   
+      }
+      case Cw12 as [Cw12 | Cw12].
+      ++assert (Hlr : ((L^min ps + L^max ps) / 2 < L^P w1 \/ L^P w2 < (L^min ps + L^max ps) / 2)%R).
+        { rewrite (segment_line L) in Hmiddle.
+          + rewrite line_middle in Hmiddle. rewrite line_P_iP in Hmiddle by assumption. lra.
+          + apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          + apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          + apply line_contains_middle ; unfold line_contains ; [now rewrite line_P_iP_min | now rewrite line_P_iP_max]. }
+        clear Hmiddle. case Hlr as [Hmiddle | Hmiddle].
+        --exists L, w1, w2. split ; [|split ; [|split ; [|split ; [|split]]]] ; 
+            auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+        --exists (line_reverse L), w2, w1. split ; [|split ; [|split ; [|split ; [|split]]]] ; 
+            auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+          **now rewrite <-aligned_on_reverse.
+          **now setoid_rewrite segment_sym.
+          **rewrite line_reverse_min, line_reverse_max, 2 line_reverse_proj. 
+            fold ps. foldR2. lra.
+      ++assert (Hlr : ((L^min ps + L^max ps) / 2 < L^P w2 \/ L^P w1 < (L^min ps + L^max ps) / 2)%R).
+        { rewrite (segment_line L) in Hmiddle.
+          + rewrite line_middle in Hmiddle. rewrite line_P_iP in Hmiddle by assumption. lra.
+          + apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          + apply Halign, (weber_segment_InA h_psnonnil NEw12 Hweb).
+          + apply line_contains_middle ; unfold line_contains ; [now rewrite line_P_iP_min | now rewrite line_P_iP_max]. }
+        clear Hmiddle. case Hlr as [Hmiddle | Hmiddle].
+        --exists L, w2, w1. split ; [|split ; [|split ; [|split ; [|split]]]] ; 
+            auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+          **now setoid_rewrite segment_sym.
+        --exists (line_reverse L), w1, w2. split ; [|split ; [|split ; [|split ; [|split]]]] ; 
+            auto using config_stay_impl_config_stg, config_stay_impl_endpoints_stay.
+          **now rewrite <-aligned_on_reverse.
+          **rewrite line_reverse_min, line_reverse_max, 2 line_reverse_proj. 
+            fold ps. foldR2. lra.
+Qed.
+    
+
+Lemma round_simplify_phase1 c da w : 
+  inv1 c w ->
+  similarity_da_prop da ->
+  exists (r : ident -> ratio),
+  round gatherW da c == 
+  fun id => if activate da id 
+            then (get_location (c id), w, ratio_0)
+            else inactive c id (r id).
+Proof using lt_0n delta_g0. 
+intros Hinv Hsim. destruct (round_simplify Hsim (inv1_valid Hinv)) as [r Hround].
+exists r. rewrite Hround. intros id. destruct_match ; [|reflexivity].
+repeat split ; cbn [fst snd]. 
+assert (Hw12 := weber_calc_seg_correct_2 c).
+destruct (weber_calc_seg (pos_list c)) as [w1 w2].
+specialize (Hw12 w1 w2 eq_refl).
+assert (Hw1 : w1 == w).
+{ apply Hinv. rewrite Hw12. apply segment_start. }
+assert (Hw2 : w2 == w).
+{ apply Hinv. rewrite Hw12. apply segment_end. }
+rewrite Hw1, Hw2. destruct_match ; intuition.
+Qed.
+
+Lemma round_simplify_phase2 c da L :
+  inv2 c L -> 
+  similarity_da_prop da ->
+  exists (r : ident -> ratio),
+  round gatherW da c == 
+  fun id => if activate da id then  
+              let target := 
+                if existsb 
+                  (fun x => 
+                    (Rltb (L^min (pos_list c)) (L^P x)) && 
+                    (Rltb (L^P x) (L^P (get_location (c id))))) 
+                  (pos_list c) 
+                then get_location (c id) else L^-P (L^min (pos_list c))
+              in (get_location (c id), target, ratio_0)
+            else inactive c id (r id).
+Proof using lt_0n delta_g0.
+intros Hinv2 Hsim. assert (Hvalid := inv2_valid Hinv2). 
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+exists r. rewrite Hround. clear Hround. intros id. destruct_match ; [|reflexivity].
+repeat split ; cbn [fst snd]. unfold inv2 in Hinv2. 
+destruct (weber_calc_seg (pos_list c)) as [w1 w2]. case (w1 =?= w2) as [Ew12 | NEw12] ; [now intuition|].
+pose (t := L^-P (L^min (pos_list c))). foldR2. fold t in Hinv2 |- *.
+destruct (robot_with_mult (Nat.div2 (n+1)) (pos_list c)) as [t2|] eqn:E.
++ assert (Ht : t2 == t).
+  { 
+    unfold robot_with_mult in E. apply find_some in E.
+    destruct E as [E1 E2]. case (t2 =?= t) as [Ht | Ht] ; [intuition|].
+    exfalso. apply Hvalid. unfold invalid. exists t, t2. 
+    repeat split ; [intuition | apply Hinv2 |].
+    rewrite eqb_true_iff in E2. now rewrite multiplicity_countA, E2.
+  }
+  repeat rewrite Ht.
+  assert (Hseg : forall x, InA equiv x (pos_list c) -> 
+    strict_segment_decb (get_location (c id)) t x = 
+      Rltb (L^min (pos_list c)) (L^P x) &&
+      Rltb (L^P x) (L^P (get_location (c id)))).
+  {
+    intros x Hin. rewrite eq_iff_eq_true, strict_segment_decb_true, andb_true_iff, 2 Rltb_true.
+    rewrite strict_segment_sym, (strict_segment_line L). 
+    + unfold t. rewrite line_P_iP_min. split ; [|now intuition].
+      intros [H1 | H2] ; [now auto|]. exfalso. destruct H2 as [_ H2].
+      apply Rlt_not_le in H2. apply H2. apply line_bounds ; intuition. 
+    + unfold t. apply Hinv2. apply line_iP_min_InA ; [|intuition].
+      rewrite <-length_zero_iff_nil, pos_list_length. lia.
+    + apply Hinv2. rewrite pos_list_InA. now exists id.
+    + apply Hinv2. auto.
+  }
+  case ifP_bool ; case ifP_bool ; intros H1 H2 ; try reflexivity.
+  - exfalso. rewrite existsb_exists in H2. destruct H2 as [x [Hin H2]].
+    apply (f_equal negb) in H1. rewrite forallb_existsb, forallb_forall in H1.
+    rewrite Hseg in H2 by (apply In_InA ; autoclass).
+    apply (f_equal negb) in H2. now rewrite H1 in H2.
+  - exfalso. rewrite existsb_exists in H1. destruct H1 as [x [Hin H1]].
+    apply (f_equal negb) in H2. rewrite forallb_existsb, forallb_forall in H2.
+    rewrite <-Hseg in H1 by (apply In_InA ; autoclass).
+    apply (f_equal negb) in H1. now rewrite H2 in H1.
++ exfalso. apply find_none with (x := t) in E.
+  - rewrite eqb_false_iff in E. apply E. now rewrite <-multiplicity_countA.
+  - rewrite <-(In_InA_is_leibniz equiv) by now cbn.
+    apply line_iP_min_InA ; [|intuition].
+    rewrite <-length_zero_iff_nil, pos_list_length. lia.
+Qed.
+
+
+(* Ltac dolsps ps := *)
+(*   match goal with *)
+(*     H: let ps := _  |- _ => idtac H *)
+(*   end *)
+(* . *)
+Ltac fold_all t :=
+  (repeat progress match goal with
+      H  : _ |- _ => fold t in H
+    end); fold t.
+
+Ltac fold_all2 t :=
+  (repeat progress match goal with
+      H  : _ |- _ => progress fold t in H
+    end); fold t.
+
+Lemma line_dir_nonnul_sufficient L ps w1 w2: 
+ps <> nil ->
+  aligned_on L ps ->
+  (forall w : R2, Weber ps w <-> segment w1 w2 w) ->
+  w1 =/= w2 ->
+  line_dir L =/= 0%VS.
+Proof using Type.
+  intros h_nonnil h_align h_web_seg h_neq. 
+  apply (line_dir_nonnul _ w1 w2).
+    - apply h_neq.
+    - apply h_align.
+      apply (@weber_segment_InA _ w1 w2);auto.
+    - apply h_align. apply (@weber_segment_InA _ w1 w2);auto.
+Qed.
+
+Lemma round_simplify_phase3 da c L w1 w2 :
+  inv3 c L w1 w2 ->
+  similarity_da_prop da ->
+  exists (r : ident -> ratio),
+  round gatherW da c == 
+  fun id => if activate da id then  
+              let target := 
+                if (w1 ==b get_location (c id)) || (w2 ==b get_location (c id)) 
+                then middle (L^-P (L^min (pos_list c))) (L^-P (L^max (pos_list c))) else get_location (c id)   
+              in (get_location (c id), target, ratio_0)
+            else inactive c id (r id).
+Proof using lt_0n delta_g0.
+specialize (pos_list_non_nil c) as h_ps_nonnil.
+pose (ps := pos_list c).
+fold_all2 ps.
+intros Hinv3 Hsim. assert (Hvalid := inv3_valid Hinv3). 
+assert(line_dir L =/= 0%VS) as HlineDir.
+{ apply (@line_dir_nonnul_sufficient L ps w1 w2);auto; apply Hinv3. }
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+fold_all2 ps.
+exists r. rewrite Hround. clear Hround. intros id. destruct_match ; [|reflexivity].
+repeat split ; cbn [fst snd]. 
+destruct Hinv3 as [Halign [Hweb [NEw12 [Hmiddle [Hstg [Hends [Hr1_mult Hr2_mult]]]]]]].
+assert (Hweb' := weber_calc_seg_correct_2 c).
+fold_all2 ps.
+destruct (weber_calc_seg ps) as [w1' w2'].
+specialize (Hweb' w1' w2' eq_refl).
+assert (Hperm : perm_pairA equiv (w1', w2') (w1, w2)).
+{ apply segment_eq_iff. intros x. now rewrite <-Hweb'. }
+case (w1' =?= w2') as [Ew12' | NEw12'].
++ exfalso. apply NEw12. case Hperm as [[H1 H2] | [H1 H2]].
+  - rewrite <-H1, <-H2. exact Ew12'.
+  - rewrite <-H1, <-H2. symmetry. exact Ew12'.
++ pose (r1 := L^-P (L^min ps)). pose (r2 := L^-P (L^max ps)).
+  destruct (robot_with_mult (Nat.div2 (n+1)) ps) as [t|] eqn:E.
+  - exfalso.
+    unfold robot_with_mult in E. apply find_some in E. 
+    destruct E as [_ E]. rewrite eqb_true_iff in E.
+    assert (Hps : ps <> nil).
+    { unfold ps. rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+    assert (Ht : InA equiv t ps).
+    { unfold ps in *. rewrite <-multiplicity_pos_InA, multiplicity_countA, E, gt_lt_iff.
+      apply Exp_prop.div2_not_R0. lia. }
+    cut (OnlyWeber ps t).
+    { intros HOW. apply NEw12. transitivity t.
+      * apply HOW. fold_all2 ps. rewrite Hweb. apply segment_start.
+      * symmetry. apply HOW. rewrite Hweb. apply segment_end. } 
+    assert (aligned_on L (t::ps)) as halignt.
+    { apply aligned_on_cons_iff.
+      split.
+      - apply aligned_on_InA_contains with ps;auto.
+      - assumption. }
+    fold_all2 ps.
+    rewrite (weber_aligned_spec h_ps_nonnil HlineDir halignt).
+    assert (Ht_line : line_contains L t).
+    { apply Halign. unfold ps in *. rewrite <-multiplicity_pos_InA, multiplicity_countA, E.
+      apply Exp_prop.div2_not_R0 ; lia. }
+    fold_all2 ps.
+    rewrite line_on_length_aligned;auto.
+    rewrite E by now rewrite aligned_on_cons_iff.
+    assert (Hlr_bounds := @majority_left_right_bounds t c L Halign).
+    feed_n 2 Hlr_bounds.
+    { rewrite multiplicity_countA. fold ps. rewrite E. intuition. }
+    { fold ps in Hr2_mult, Hr1_mult,Halign.
+      destruct (line_bounds L ps t) as [[H1 | H2] [H3 | H4]] ; try assumption.
+      + tauto.
+      + exfalso. rewrite <-H4, Halign,multiplicity_countA (*, E*) in Hr2_mult by assumption.
+        fold ps in Hr2_mult.
+        rewrite E in Hr2_mult.
+        lia.
+      + exfalso. rewrite H2, Halign, multiplicity_countA in Hr1_mult by assumption.
+        fold ps in Hr1_mult.
+        rewrite E in Hr1_mult.
+        lia.
+      + exfalso. rewrite H2, Halign, multiplicity_countA in Hr1_mult by assumption.
+        fold ps in Hr1_mult.
+        rewrite E in Hr1_mult.
+        lia. } 
+    fold_all2 ps.
+    destruct Hlr_bounds as [[Hl1 Hl2] [Hr1 Hr2]]. 
+    apply le_INR in Hl1. apply lt_INR in Hl2. apply le_INR in Hr1. apply lt_INR in Hr2. 
+    cbn [INR] in Hl1, Hr1. foldR2. unfold Rabs. destruct_match; lra.
+- pose (L' := line_calc ps). fold L'.
+    assert (H_LL' := @line_endpoints_invariant _ _ _ _ _ L L' ps).
+    feed_n 3 H_LL'.
+    { unfold ps. rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+    { exact Halign. }
+    { unfold L'. apply line_calc_spec. exists L. exact Halign. }
+    pose (r1' := L'^-P (L'^min ps)). pose (r2' := L'^-P (L'^max ps)).
+    foldR2. fold r1 r2 r1' r2' in H_LL' |- *.
+    assert (Hmiddle_eq : middle r1 r2 == middle r1' r2').
+    { case H_LL' as [[-> ->] | [-> ->]] ; [|rewrite middle_comm] ; reflexivity. }
+    repeat rewrite <-Hmiddle_eq.
+    assert (Hseg : segment w1 w2 (middle r1 r2)).
+    { 
+      assert (Hps : ps <> nil).
+      { unfold ps. rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+      rewrite (segment_line L).
+      + left. unfold r1, r2. rewrite line_middle. rewrite line_P_iP ; [exact Hmiddle|].
+        apply (line_dir_nonnul L w1 w2) ; [assumption | |].
+        - apply weber_aligned with ps ; [assumption .. | rewrite Hweb ; apply segment_start].
+        - apply weber_aligned with ps ; [assumption .. | rewrite Hweb ; apply segment_end].
+      + apply weber_aligned with ps ; [assumption .. | rewrite Hweb ; apply segment_start].
+      + apply weber_aligned with ps ; [assumption .. | rewrite Hweb ; apply segment_end].
+      + apply line_contains_middle. 
+        - unfold r1. now apply Halign, line_iP_min_InA.
+        - unfold r2. now apply Halign, line_iP_max_InA. 
+    }
+    assert (Hperm' := Hperm). rewrite <-segment_eq_iff in Hperm'. 
+    rewrite <-Hperm', <-segment_decb_true in Hseg. foldR2. rewrite Hseg.
+    case Hperm as [[-> ->] | [-> ->]] ; [|rewrite orb_comm] ; reflexivity.
+Qed.
+   
+Lemma round_simplify_phase4 da c L w1 w2 : 
+  inv4 c L w1 w2 ->  
+  similarity_da_prop da ->
+  exists (r : ident -> ratio),
+  round gatherW da c == 
+  fun id => if activate da id then  
+              let target := 
+                if get_location (c id) ==b w2 then w1 else get_location (c id)       
+              in (get_location (c id), target, ratio_0)
+            else inactive c id (r id).
+Proof using lt_0n delta_g0.
+intros Hinv4 Hsim. assert (Hvalid := inv4_valid Hinv4). 
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+exists r. rewrite Hround. clear Hround. intros id. destruct_match ; [|reflexivity].
+repeat split ; cbn [fst snd]. 
+destruct Hinv4 as [Halign [Hweb [Hmiddle [Hstg [Hends [Hr1_mult Hr2_mult]]]]]].
+assert (NEw12 : w1 =/= w2).
+{ intros Ew12. rewrite Ew12 in Hmiddle. lra. }
+assert (Hweb' := @weber_calc_seg_correct (pos_list c)).
+destruct (weber_calc_seg (pos_list c)) as [w1' w2'].
+specialize pos_list_non_nil as h_nonnil.
+assert (Hperm : perm_pairA equiv (w1', w2') (w1, w2)).
+{ apply segment_eq_iff. intros x. now rewrite <-Hweb', Hweb. }
+case (w1' =?= w2') as [Ew12' | NEw12'].
++ exfalso. apply NEw12. case Hperm as [[H1 H2] | [H1 H2]].
+  - rewrite <-H1, <-H2. exact Ew12'.
+  - rewrite <-H1, <-H2. symmetry. exact Ew12'.
++ pose (r1 := L^-P (L^min (pos_list c))). pose (r2 := L^-P (L^max (pos_list c))).
+  destruct (robot_with_mult (Nat.div2 (n+1)) (pos_list c)) as [t|] eqn:E.
+  - exfalso.
+    unfold robot_with_mult in E. apply find_some in E. 
+    destruct E as [_ E]. rewrite eqb_true_iff in E.
+    assert (Hps : pos_list c <> nil).
+    { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+    assert (Ht : InA equiv t (pos_list c)).
+    { rewrite <-multiplicity_pos_InA, multiplicity_countA, E, gt_lt_iff.
+      apply Exp_prop.div2_not_R0. lia. }
+    cut (OnlyWeber (pos_list c) t).
+    { intros HOW. apply NEw12. transitivity t.
+      * apply HOW. rewrite Hweb. apply segment_start.
+      * symmetry. apply HOW. rewrite Hweb. apply segment_end. }
+    assert (aligned_on L (t :: pos_list c)).
+    { apply aligned_on_cons_iff.
+      split.
+      - eapply aligned_on_InA_contains;eauto.
+      - assumption. }
+    
+    rewrite (@weber_aligned_spec L (pos_list c) t);auto.
+    2:{ specialize weber_segment_InA with (3:=Hweb) as [h1 h2];auto.
+        apply line_dir_nonnul with (x:=w1) (y:=w2);auto. }
+
+    assert (Ht_line : line_contains L t).
+    { apply Halign;auto. }
+    rewrite line_on_length_aligned, E by now rewrite aligned_on_cons_iff.
+    assert (Hlr_bounds := @majority_left_right_bounds t c L Halign).
+    feed_n 2 Hlr_bounds.
+    { rewrite multiplicity_countA, E. intuition. }
+    { destruct (line_bounds L (pos_list c) t) as [[H1 | H2] [H3 | H4]] ; try assumption.
+      + tauto.
+      + exfalso. rewrite <-H4, Halign, multiplicity_countA, E in Hr2_mult by assumption. lia.
+      + exfalso. rewrite H2, Halign, multiplicity_countA, E in Hr1_mult by assumption. lia.
+      + exfalso. rewrite H2, Halign, multiplicity_countA, E in Hr1_mult by assumption. lia. } 
+    destruct Hlr_bounds as [[Hl1 Hl2] [Hr1 Hr2]]. 
+    apply le_INR in Hl1. apply lt_INR in Hl2. 
+    apply le_INR in Hr1. apply lt_INR in Hr2. 
+    cbn [INR] in Hl1, Hr1. foldR2. unfold Rabs. destruct_match ; lra.
+  - pose (L' := line_calc (pos_list c)). fold L'.
+    assert (H_LL' := @line_endpoints_invariant _ _ _ _ _ L L' (pos_list c)).
+    feed_n 3 H_LL'.
+    { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+    { exact Halign. }
+    { unfold L'. apply line_calc_spec. exists L. exact Halign. }
+    pose (r1' := L'^-P (L'^min (pos_list c))). pose (r2' := L'^-P (L'^max (pos_list c))).
+    foldR2. fold r1 r2 r1' r2' in H_LL' |- *.
+    assert (Hmiddle_eq : middle r1 r2 == middle r1' r2').
+    { case H_LL' as [[-> ->] | [-> ->]] ; [|rewrite middle_comm] ; reflexivity. }
+    repeat rewrite <-Hmiddle_eq. clear Hmiddle_eq.
+    assert (Hps : pos_list c <> nil).
+    { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+    assert (Hline_dir : line_dir L =/= 0%VS).
+    { apply (line_dir_nonnul L w1 w2) ; [assumption | |].
+      + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start].
+      + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end]. }
+    assert (Hline_r1 : line_contains L r1). 
+    { unfold r1. now apply Halign, line_iP_min_InA. }
+    assert (Hline_r2 : line_contains L r2). 
+    { unfold r2. now apply Halign, line_iP_max_InA. }
+    assert (Hline_w1 : line_contains L w1). 
+    { apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start]. }
+    assert (Hline_w2 : line_contains L w2). 
+    { apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end]. }
+    assert (Hseg1 : segment_decb w1 w2 (middle r1 r2) = false).
+    { rewrite segment_decb_false, (segment_line L) ; try assumption.
+      + unfold r1, r2.
+        repeat rewrite line_middle at 1.
+        repeat rewrite line_P_iP at 1 by assumption .
+        foldR2 ; lra.
+      + now apply line_contains_middle. }
+    assert (Hseg2 : segment_decb (middle r1 r2) w1 w2 = false).
+    { rewrite segment_decb_false, (segment_line L) ; try assumption.
+      + unfold r1, r2. rewrite line_middle, line_P_iP by assumption. foldR2. lra.
+      + now apply line_contains_middle. }
+    assert (Hseg3 : segment_decb (middle r1 r2) w2 w1 = true).
+    { rewrite segment_decb_true, (segment_line L) ; try assumption.
+      + unfold r1, r2. rewrite line_middle, line_P_iP by assumption. foldR2. lra.
+      + now apply line_contains_middle. }
+    case Hperm as [[H1 H2] | [H1 H2]] ; repeat rewrite H1, H2.
+    * foldR2. rewrite Hseg1, Hseg2, Hseg3. now rewrite eqb_sym.
+    * foldR2. rewrite (segment_decb_sym w2 w1), Hseg1, Hseg2, Hseg3. now rewrite eqb_sym.
+Qed.
+
+Lemma config_stg_mult_increase config da p : 
+  similarity_da_prop da ->
+  ~invalid config ->
+  config_stg config p -> 
+  (!!config)[p] <= (!! (round gatherW da config))[p].
+Proof using lt_0n delta_g0. 
+intros Hsim Hvalid Hstg. destruct (round_simplify Hsim Hvalid) as [r Hround].
+rewrite 2 multiplicity_countA. unfold config_stg in Hstg.
+unfold pos_list. rewrite 2 config_list_spec. induction names as [|id l IH] ; cbn -[get_location].
++ auto.
++ case ifP_sumbool ; case ifP_sumbool ; intros HRloc Hloc.
+  - now apply le_n_S.
+  - exfalso. apply HRloc. specialize (Hstg id). rewrite (Hround id). clear Hround. 
+    destruct_match.
+    * cbn [get_location]. now rewrite get_location_ratio_0. 
+    * destruct (config id) as [[s d] ri] eqn:E. cbn -[straight_path] in Hloc, Hstg |- *.
+      rewrite E. case Hstg as [Hstay | Hgo].
+      ++rewrite Hstay, straight_path_same in Hloc |- *. exact Hloc.
+      ++rewrite Hgo in Hloc |- *. change (@eq R2) with (@equiv R2 _) in Hloc |- *.
+        now apply straight_path_end_advance.
+  - now apply le_S.
+  - auto.
+Qed. 
+
+(* If all robots are going towards a point on the line, 
+ * then they stay aligned. *)
+Lemma aligned_preserved da L c x : 
+  let ps := pos_list c in  
+  let Rps := pos_list (round gatherW da c) in 
+  similarity_da_prop da -> 
+  ~invalid c ->
+  aligned_on L (x :: ps) ->
+  config_stg c x ->
+    aligned_on L Rps.
+Proof using lt_0n delta_g0. 
+cbn zeta. intros Hsim Hvalid Halign Hstg. 
+pose (ps := pos_list c). pose (Rps := pos_list (round gatherW da c)).
+foldR2. fold ps Rps in Halign |- *.
+intros y Hin. unfold Rps in Hin. rewrite pos_list_InA in Hin.
+destruct Hin as [id Hy]. revert Hy. specialize (Hstg id). 
+destruct (round_simplify Hsim Hvalid) as [r Hround]. rewrite (Hround id). clear Hround.
+destruct_match.
++ rewrite get_location_ratio_0. intros ->. apply Halign. right.
+  unfold ps. rewrite pos_list_InA. now exists id.
++ intros Hy. cbn -[equiv straight_path] in Hy, Hstg. destruct (c id) as [[s d] ri] eqn:E.
+  cbn -[equiv straight_path] in Hstg. case Hstg as [Hstay | Hgo].
+  - apply Halign. right. unfold ps. rewrite pos_list_InA. exists id. 
+    rewrite Hy, E. cbn -[equiv straight_path]. now rewrite Hstay, 2 straight_path_same.
+  - foldR2. 
+    rewrite Hgo in Hy. rewrite Hy. clear Hy. pose (R := add_ratio ri (r id)). fold R.
+    (* rewrite Hgo in *. rewrite Hy. clear Hy Hgo. pose (R := add_ratio ri (r id)). fold R. *)
+    case (proj_ratio ri =?= 1%R) as [Hri | Hri].
+    * rewrite straight_path_eq_1. 
+      2: unfold R ; now rewrite add_ratio_1_left.
+      apply Halign ; now left.
+    * cbn -[mul opp RealVectorSpace.add line_contains].  
+      apply line_contains_combine ; [|now apply Halign ; left].
+      assert (Hs : (s == x + (1 / (1 - ri)) * (straight_path s x ri - x))%VS).
+      { 
+        assert (Hrewrite : (straight_path s x ri - x == (1 - ri) * (s - x))%VS).
+        { cbn -[mul opp RealVectorSpace.add equiv]. unfold Rminus. 
+          rewrite <-add_morph, mul_1, minus_morph, mul_distr_add, mul_opp.
+          repeat rewrite <-RealVectorSpace.add_assoc. apply RealVectorSpace.add_compat ; auto.
+          rewrite mul_distr_add, opp_distr_add, mul_opp, opp_opp.
+          rewrite RealVectorSpace.add_comm, RealVectorSpace.add_assoc.
+          apply RealVectorSpace.add_compat ; auto. now rewrite RealVectorSpace.add_comm. }
+        rewrite Hrewrite. clear Hrewrite.
+        rewrite mul_morph. unfold Rdiv. rewrite Rmult_assoc, Rinv_l by (cbn in Hri ; lra).
+        now rewrite Rmult_1_l, mul_1, add_sub.
+      }
+      rewrite Hs. apply line_contains_combine ; apply Halign ; [now left | right].
+      unfold ps. rewrite pos_list_InA. exists id. rewrite E.
+      cbn -[straight_path].
+      now rewrite Hgo.
+Qed.
+
+(* Without the explicit type for a and b, type class resolution go
+   wrong. *)
+Lemma straight_path_rebase_start (a b:R2) r1 r2 :
+  exists r : ratio, straight_path a b (add_ratio r1 r2) == straight_path (straight_path a b r1) b r. 
+Proof using .
+assert (Hr1_bounds := ratio_bounds r1).
+assert (Hr2_bounds := ratio_bounds r2).
+unfold add_ratio. case (Rle_dec R1 (r1 + r2)) as [Hle | HNle].
++ exists ratio_1. now rewrite 2 straight_path_1.
++ change R1 with 1%R in HNle.
+  pose (r := (r2 / (1 - r1))%R).
+  assert (Hr : (0 <= r <= 1)%R).
+  { split ; [apply Rdiv_le_0_compat | apply Rdiv_le_1] ; lra. }
+  exists (exist _ r Hr). unfold r. cbn -[equiv mul opp RealVectorSpace.add].
+  apply mul_reg_l with (1 - r1)%R ; [lra|].
+  repeat rewrite (mul_distr_add (1 - r1)) at 1.
+  repeat rewrite mul_morph.
+  unfold Rdiv. rewrite <-Rmult_assoc, Rinv_r_simpl_m by lra.
+  repeat rewrite <-RealVectorSpace.add_assoc. apply RealVectorSpace.add_compat ; [reflexivity|].
+  rewrite Rmult_plus_distr_l, <-add_morph. apply RealVectorSpace.add_compat ; [reflexivity|]. 
+  rewrite Rmult_comm, <-mul_morph. apply mul_compat ; [reflexivity|].
+  unfold Rminus. rewrite <-add_morph, mul_1, opp_distr_add, minus_morph, RealVectorSpace.add_assoc.
+  reflexivity.
+Qed.
+
+(* If the endpoints aren't moving and the configuration is contracting
+ * to a point [x] that is inside the configuration segment, 
+ * then the endpoints don't move *)
+Lemma endpoints_preserved da L c x :
+  let ps := pos_list c in  
+  let Rps := pos_list (round gatherW da c) in 
+  similarity_da_prop da -> 
+  ~invalid c ->
+  aligned_on L (x :: ps) ->
+  endpoints_stay c L -> 
+  config_stg c x -> 
+  (L^min ps <= L^P x <= L^max ps)%R ->
+    L^min Rps == L^min ps /\ L^max Rps == L^max ps.
+Proof using lt_0n delta_g0. 
+cbn zeta. pose (ps := pos_list c). pose (Rps := pos_list (round gatherW da c)).
+fold ps Rps. rewrite aligned_on_cons_iff. intros Hsim Hvalid [Hx Halign] Hends Hstg Hseg.
+pose (r1 := L^-P (L^min ps)). pose (r2 := L^-P (L^max ps)).
+pose (R1 := L^-P (L^min Rps)). pose (R2 := L^-P (L^max Rps)).
+assert (Hr1 : L^min ps == L^P r1). { unfold r1. now rewrite line_P_iP_min. }
+assert (Hr2 : L^max ps == L^P r2). { unfold r2. now rewrite line_P_iP_max. }
+assert (HR1 : L^min Rps == L^P R1). { unfold R1. now rewrite line_P_iP_min. }
+assert (HR2 : L^max Rps == L^P R2). { unfold R2. now rewrite line_P_iP_max. }
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+split.
+(* min *)
++ setoid_rewrite line_min_spec_nonnil at 1 ; [|unfold Rps ; rewrite <-length_zero_iff_nil, pos_list_length ; lia].
+  apply Rle_antisym.
+  - apply fold_left_Rmin_le. 
+    exists (L^P r1). split ; [|now rewrite Hr1].
+    rewrite InA_map_iff ; autoclass. exists r1. split ; auto ; [].
+    assert (Hin : InA equiv r1 ps).
+    { apply line_iP_min_InA ; auto. unfold ps. rewrite <-length_zero_iff_nil.
+      rewrite pos_list_length ; lia. }
+    unfold ps in Hin. rewrite pos_list_InA in Hin. destruct Hin as [id Hid].
+    unfold Rps. rewrite pos_list_InA. exists id. symmetry in Hid |- *.
+    rewrite (Hround id). clear Hround. destruct_match ; [now rewrite get_location_ratio_0|].
+    assert (Hid' := Hid). apply Hends in Hid'. fold ps r1 in Hid'.
+    cbn -[get_location r1 equiv]. destruct (c id) as [[s d] ri].
+    cbn -[equiv straight_path r1] in Hid, Hid' |- *. rewrite Hid' in Hid |- *.
+    rewrite straight_path_end in Hid. 
+    case Hid as [-> | Hri] ; [now rewrite straight_path_same|].
+    rewrite straight_path_eq_1 ; auto. rewrite add_ratio_1_left ; auto.
+  - apply fold_left_Rmin_ge. intros t Hin'. 
+    assert (Hin : InA equiv t (List.map (L^P) Rps)).
+    { rewrite InA_cons in Hin'. case Hin' as [-> | Hin] ; [|exact Hin].
+      rewrite InA_map_iff ; autoclass. exists (nth 0 Rps 0%VS). split ; [reflexivity|].
+      apply nth_InA ; autoclass. unfold Rps. now rewrite pos_list_length. }
+    clear Hin'. rewrite InA_map_iff in Hin ; autoclass. destruct Hin as [y [<- Hin]].
+    unfold Rps in Hin. rewrite pos_list_InA in Hin. destruct Hin as [id ->].
+    rewrite Hr1. rewrite (Hround id). clear Hround. destruct_match. 
+    * rewrite get_location_ratio_0, <-Hr1. apply line_bounds ; auto. 
+      unfold ps. rewrite pos_list_InA. now exists id.
+    * cbn -[r1 straight_path line_proj]. specialize (Hstg id). 
+      destruct (c id) as [[s d] ri] eqn:E. cbn -[equiv straight_path] in Hstg.
+      case Hstg as [Hstay | Hgo].
+      ++rewrite <-Hr1. apply line_bounds ; auto ; [].
+        unfold ps. rewrite pos_list_InA. exists id. rewrite E.
+        cbn -[equiv straight_path]. now rewrite Hstay, 2 straight_path_same. 
+      ++rewrite Hgo.
+        destruct (@straight_path_rebase_start s x ri (r id)) as [rr ->]. 
+        pose (p := straight_path s x ri). foldR2. fold p. cbn -[r1 p line_proj mul opp RealVectorSpace.add].
+        rewrite line_proj_combine, Rmult_minus_distr_l. unfold Rminus.
+        rewrite Rplus_comm, Rplus_assoc, Ropp_mult_distr_l. setoid_rewrite <-(Rmult_1_l (L^P p)) at 2.
+        rewrite <-Rmult_plus_distr_r. replace (L^P r1) with (rr * L^P r1 + (- rr + 1) * L^P r1)%R ; [|nra].
+        apply Rplus_le_compat. 
+        --apply Rmult_le_compat_l ; [now apply ratio_bounds|].
+          rewrite <-Hr1. apply Hseg.
+        --apply Rmult_le_compat_l ; [generalize (ratio_bounds rr) ; lra|].
+          rewrite <-Hr1. apply line_bounds. 
+          unfold p, ps. rewrite pos_list_InA. exists id. now rewrite <-Hgo, E.
+(* same for max *)
++ setoid_rewrite line_max_spec_nonnil at 1 ; [|unfold Rps ; rewrite <-length_zero_iff_nil, pos_list_length ; lia].
+  apply Rle_antisym.
+  - apply fold_left_Rmax_le. intros t Hin'. 
+    assert (Hin : InA equiv t (List.map (L^P) Rps)).
+    { rewrite InA_cons in Hin'. case Hin' as [-> | Hin] ; [|exact Hin].
+      rewrite InA_map_iff ; autoclass. exists (nth 0 Rps 0%VS). split ; [reflexivity|].
+      apply nth_InA ; autoclass. unfold Rps. now rewrite pos_list_length. }
+    clear Hin'. rewrite InA_map_iff in Hin ; autoclass. destruct Hin as [y [<- Hin]].
+    unfold Rps in Hin. rewrite pos_list_InA in Hin. destruct Hin as [id ->].
+    rewrite Hr2. rewrite (Hround id). clear Hround. destruct_match. 
+    * rewrite get_location_ratio_0, <-Hr2. apply line_bounds ; auto. 
+      unfold ps. rewrite pos_list_InA. now exists id.
+    * cbn -[r2 straight_path line_proj]. specialize (Hstg id). 
+      destruct (c id) as [[s d] ri] eqn:E. cbn -[equiv straight_path] in Hstg.
+      case Hstg as [Hstay | Hgo].
+      ++rewrite <-Hr2. apply line_bounds ; auto ; [].
+        unfold ps. rewrite pos_list_InA. exists id. rewrite E.
+        cbn -[equiv straight_path]. now rewrite Hstay, 2 straight_path_same. 
+      ++rewrite Hgo.
+        destruct (@straight_path_rebase_start s x ri (r id)) as [rr ->]. 
+        pose (p := straight_path s x ri). foldR2. fold p. cbn -[r2 p line_proj mul opp RealVectorSpace.add].
+        rewrite line_proj_combine, Rmult_minus_distr_l. unfold Rminus.
+        rewrite Rplus_comm, Rplus_assoc, Ropp_mult_distr_l. setoid_rewrite <-(Rmult_1_l (L^P p)) at 2.
+        rewrite <-Rmult_plus_distr_r. replace (L^P r2) with (rr * L^P r2 + (- rr + 1) * L^P r2)%R ; [|nra].
+        apply Rplus_le_compat. 
+        --apply Rmult_le_compat_l ; [now apply ratio_bounds|].
+          rewrite <-Hr2. apply Hseg.
+        --apply Rmult_le_compat_l ; [generalize (ratio_bounds rr) ; lra|].
+          rewrite <-Hr2. apply line_bounds. 
+          unfold p, ps. rewrite pos_list_InA. exists id. now rewrite <-Hgo, E.
+  - apply fold_left_Rmax_ge. 
+    exists (L^P r2). split ; [|now rewrite Hr2].
+    rewrite InA_map_iff ; autoclass. exists r2. split ; auto ; [].
+    assert (Hin : InA equiv r2 ps).
+    { unfold r2. apply line_iP_max_InA ; auto ; []. unfold ps. rewrite <-length_zero_iff_nil.
+      rewrite pos_list_length ; lia. }
+    unfold ps in Hin. rewrite pos_list_InA in Hin. destruct Hin as [id Hid].
+    unfold Rps. rewrite pos_list_InA. exists id. symmetry in Hid |- *.
+    rewrite (Hround id). clear Hround. destruct_match ; [now rewrite get_location_ratio_0|].
+    assert (Hid' := Hid). apply Hends in Hid'. fold ps r1 in Hid'.
+    cbn -[get_location r1 equiv]. destruct (c id) as [[s d] ri].
+    cbn -[equiv straight_path r1] in Hid, Hid' |- *. rewrite Hid' in Hid |- *.
+    rewrite straight_path_end in Hid. 
+    case Hid as [-> | Hri] ; [now rewrite straight_path_same|].
+    rewrite straight_path_eq_1 ; auto. rewrite add_ratio_1_left ; auto.
+Qed.
+
+(* If the endpoints aren't moving and the configuration is contracting
+ * to a point [x] that is inside the configuration segment, 
+ * then the set of robots on [L^-P (L^min ps)] doesn't change in a round (provided that x <> r1).
+ * The same goes for [L^-P (L^max ps)]. *)
+Lemma endpoints_preserved_strong da L c x :
+  let ps := pos_list c in
+  similarity_da_prop da -> 
+  ~invalid c ->
+  aligned_on L (x :: ps) ->
+  endpoints_stay c L -> 
+  config_stg c x ->
+    ((L^min ps < L^P x)%R -> (forall id, get_location (round gatherW da c id) == L^-P (L^min ps) <-> get_location (c id) == L^-P (L^min ps))) /\
+    ((L^P x < L^max ps)%R -> (forall id, get_location (round gatherW da c id) == L^-P (L^max ps) <-> get_location (c id) == L^-P (L^max ps))).
+Proof using lt_0n delta_g0.
+cbn zeta. pose (ps := pos_list c). fold ps. rewrite aligned_on_cons_iff. 
+intros Hsim Hvalid [Hx Halign] Hends Hstg.
+pose (r1 := L^-P (L^min ps)). pose (r2 := L^-P (L^max ps)). fold r1 r2.
+assert (Hr1 : L^min ps == L^P r1). { unfold r1. now rewrite line_P_iP_min. }
+assert (Hr2 : L^max ps == L^P r2). { unfold r2. now rewrite line_P_iP_max. }
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+split.
++ intros Hseg.
+  assert (NExr1 : x =/= r1).
+  { intros Exr1. apply (f_equal (L^P)) in Exr1. rewrite <-Hr1 in Exr1. lra. }
+  intros id. rewrite (Hround id). clear Hround. destruct_match ; [now rewrite get_location_ratio_0|].
+  destruct (c id) as [[s d] ri] eqn:E. generalize (Hstg id). cbn -[equiv straight_path r1]. rewrite E.
+  cbn -[equiv r1 straight_path].
+  intros [h | h].
+  - rewrite h. now rewrite 2 straight_path_same.
+  - rewrite h. split.
+    * destruct (@straight_path_rebase_start s x ri (r id)) as [rr ->].
+      intros Hpath. apply (f_equal (L^P)) in Hpath.
+      cut (L^P (straight_path s x ri) == L^min ps). 
+      { intros Hcut. apply (f_equal (L^-P)) in Hcut. fold r1 in Hcut. rewrite Halign in Hcut ; [exact Hcut|].
+        unfold ps. rewrite pos_list_InA. exists id. now rewrite E,h. }
+      unfold r1 in Hpath. rewrite line_P_iP_min in Hpath. 
+      pose (p := straight_path s x ri). foldR2. fold p in Hpath |- *.
+      case (proj_ratio rr =?= 0%R) as [Hrr0 | Hrr0] ; [now rewrite straight_path_eq_0 in Hpath|]. 
+      case (proj_ratio rr =?= 1%R) as [Hrr1 | Hrr1].
+      ++exfalso. rewrite straight_path_eq_1 in Hpath ; auto ; []. apply NExr1.
+        apply (f_equal (L^-P)) in Hpath. rewrite Hx in Hpath. exact Hpath. 
+      ++cbn -[p ps line_proj mul opp RealVectorSpace.add] in Hpath.
+        rewrite line_proj_combine in Hpath. apply (Rplus_eq_compat_l (-L^P p)%R) in Hpath.
+        rewrite <-Rplus_assoc, Rplus_opp_l, Rplus_0_l, Rplus_comm in Hpath.
+        apply (f_equal (Ropp)) in Hpath. unfold Rminus in Hpath. 
+        rewrite Ropp_mult_distr_r, 2 Ropp_plus_distr, 2 Ropp_involutive, 2 (Rplus_comm _ (L^P p)) in Hpath.
+        cut (rr * ((L ^P) p + - (L ^P) x) < 1 * ((L ^P) p + - (L ^min) ps))%R. { foldR2. lra. }
+        apply Rmult_le_0_lt_compat.
+        --apply ratio_bounds.
+        --apply Rmult_le_reg_l with rr ; [generalize (ratio_bounds rr) ; cbn in Hrr0 ; lra|].
+          foldR2. rewrite Rmult_0_r, Hpath. apply Rplus_le_reg_r with (L^min ps).
+          rewrite Rplus_0_l, Rplus_assoc, Rplus_opp_l, Rplus_0_r.
+          apply line_bounds. unfold p, ps. rewrite pos_list_InA. exists id. now rewrite E,h.
+        --generalize (ratio_bounds rr). cbn in Hrr1. lra.
+        --apply Rplus_lt_compat_l, Ropp_lt_contravar. auto.
+  * intros Hid. unfold r1 in Hid. change (straight_path s x ri) with (get_location (s, x, ri)) in Hid.
+    rewrite <-h, <-E in Hid. apply Hends in Hid. rewrite E,h in Hid. fold ps r1 in Hid. cbn -[equiv r1] in Hid. intuition.
++ intros Hseg.
+  assert (NExr2 : x =/= r2).
+  { intros Exr2. apply (f_equal (L^P)) in Exr2. rewrite <-Hr2 in Exr2. lra. }
+    intros id. rewrite (Hround id). clear Hround. destruct_match ; [now rewrite get_location_ratio_0|].
+  destruct (c id) as [[s d] ri] eqn:E. generalize (Hstg id). cbn -[equiv straight_path r2]. rewrite E.
+  cbn -[equiv r2 straight_path].
+  intros [h | h].
+  - rewrite h. now rewrite 2 straight_path_same.
+  - rewrite h. split.
+    * destruct (@straight_path_rebase_start s x ri (r id)) as [rr ->].
+      intros Hpath. apply (f_equal (L^P)) in Hpath.
+      cut (L^P (straight_path s x ri) == L^max ps). 
+      { intros Hcut. apply (f_equal (L^-P)) in Hcut. fold r2 in Hcut. rewrite Halign in Hcut ; [exact Hcut|].
+        unfold ps. rewrite pos_list_InA. exists id. now rewrite E,h. }
+      unfold r2 in Hpath. rewrite line_P_iP_max in Hpath. 
+      pose (p := straight_path s x ri). foldR2. fold p in Hpath |- *.
+      case (proj_ratio rr =?= 0%R) as [Hrr0 | Hrr0] ; [now rewrite straight_path_eq_0 in Hpath|]. 
+      case (proj_ratio rr =?= 1%R) as [Hrr1 | Hrr1].
+      ++exfalso. rewrite straight_path_eq_1 in Hpath ; auto ; []. apply NExr2.
+        apply (f_equal (L^-P)) in Hpath. rewrite Hx in Hpath. exact Hpath. 
+      ++cbn -[p ps line_proj mul opp RealVectorSpace.add] in Hpath.
+        rewrite line_proj_combine in Hpath. apply (Rplus_eq_compat_l (-L^P p)%R) in Hpath.
+        rewrite <-Rplus_assoc, Rplus_opp_l, Rplus_0_l, Rplus_comm in Hpath.
+        cut (rr * (L^P x - L^P p) < 1 * (L^max ps + - L^P p))%R. { foldR2. lra. }
+        apply Rmult_le_0_lt_compat.
+        --apply ratio_bounds.
+        --apply Rmult_le_reg_l with rr ; [generalize (ratio_bounds rr) ; cbn in Hrr0 ; lra|].
+          foldR2. rewrite Rmult_0_r, Hpath. apply Rplus_le_reg_r with (L^P p).
+          rewrite Rplus_0_l, Rplus_assoc, Rplus_opp_l, Rplus_0_r.
+          apply line_bounds. unfold p, ps. rewrite pos_list_InA. exists id. now rewrite E,h.
+        --generalize (ratio_bounds rr). cbn in Hrr1. lra.
+        --apply Rplus_lt_compat_r ; auto.
+    * intros Hid. unfold r2 in Hid. change (straight_path s x ri) with (get_location (s, x, ri)) in Hid.
+      rewrite <-h,<-E in Hid. apply Hends in Hid. rewrite E,h in Hid. fold ps r1 in Hid. cbn -[equiv r1] in Hid. intuition.
+Qed.
+
+(* If all robots are aiming for the same point,
+ * then each round the configuration contracts towards this point. *)
+Lemma config_stg_contract c da x : 
+  similarity_da_prop da -> 
+  ~invalid c ->
+  config_stg c x -> 
+  contract (pos_list c) (pos_list (round gatherW da c)) x.
+Proof using lt_0n delta_g0.
+intros Hsim Hvalid Hstg.
+unfold contract. rewrite Forall2_Forall, Forall_forall by now rewrite 2 pos_list_length.
+intros [y y'] Hin. apply In_InA with (eqA := equiv) in Hin ; autoclass.
+rewrite pos_list_InA_combine in Hin. destruct Hin as [id [-> ->]].
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+rewrite (Hround id). clear Hround. destruct_match.
++ rewrite get_location_ratio_0. apply segment_end. 
++ specialize (Hstg id). cbn -[straight_path] in Hstg |- *.
+  destruct (c id) as [[s d] ri]. cbn -[straight_path] in Hstg.
+  case Hstg as [Hstay | Hgo].
+  - rewrite Hstay, 2 straight_path_same. apply segment_end.
+  - rewrite Hgo, segment_sym. destruct (straight_path_rebase_start s x ri (r id)) as [rr ->].
+    apply segment_straight_path.
+Qed.   
+
+(* This describes the transitions we can make from phase 1. *)
+Lemma phase1_transitions c da w : 
+  similarity_da_prop da ->
+  inv1 c w -> 
+  inv1 (round gatherW da c) w.
+Proof using lt_0n delta_g0. 
+intros Hsim Hinv1. destruct (round_simplify_phase1 Hinv1 Hsim) as [r Hround].
+assert (Hinv1' := Hinv1). destruct Hinv1' as [Hstg Hweb]. split.
++ rewrite Hround. intros i. destruct_match ; [now right|].
+  specialize (Hstg i). cbn -[equiv get_location]. destruct (c i) as [[s d] ri]. 
+  cbn -[equiv straight_path] in Hstg |- *. case Hstg as [Hstg | ->] ; [now left | now right].
++ revert Hweb. apply weber_contract_unique. apply config_stg_contract ; [| eapply inv1_valid ; eauto|] ; assumption.
+Qed.
+
+(* This describes the transitions we can make from phase 2. *)
+Lemma phase2_transitions c da L :
+  similarity_da_prop da -> 
+  inv2 c L -> 
+  (exists w, inv1 (round gatherW da c) w) \/ 
+  inv2 (round gatherW da c) L.
+Proof using lt_0n delta_g0.
+intros Hsim Hinv2. destruct (round_simplify_phase2 Hinv2 Hsim) as [r Hround].
+assert (Hinv2' := Hinv2).
+unfold inv2 in Hinv2'. destruct Hinv2' as [Halign [Hw12 [Hr1_mult [Hr2_mult [Hstg Hr2_stay]]]]].
+pose (r1 := L^-P (L^min (pos_list c))). pose (r2 := L^-P (L^max (pos_list c))).
+fold r1 r2 in Hround, Hr1_mult, Hr2_mult, Hr2_stay, Hstg |- *.
+assert (Hweb := @weber_calc_seg_correct (pos_list c) (pos_list_non_nil _)).
+assert (Hr1_multR : (!! (round gatherW da c))[r1] >= Nat.div2 (n+1)).
+{ rewrite ge_le_iff. transitivity (!! c)[r1].
+  + now rewrite Hr1_mult.
+  + apply config_stg_mult_increase ; auto ; []. eapply inv2_valid ; eauto. }
+assert (HRstg : config_stg (round gatherW da c) r1).
+{ rewrite Hround. intros i. destruct_match.
+  + cbn zeta. cbn [get_dest get_start]. destruct_match ; [now left | now right].
+  + cbn -[equiv straight_path r1]. specialize (Hstg i). destruct (c i) as [[s d] ri].
+    cbn -[equiv r1 straight_path] in Hstg |- *. case Hstg as [Hstay | Hgo] ; [now left | now right]. }
+assert (Hr1_webR : Weber (pos_list (round gatherW da c)) r1).
+{ apply weber_majority_weak. now rewrite pos_list_length, <-multiplicity_countA. }
+assert (NEr12 : r1 =/= r2).
+{ apply inv234_endpoints_distinct. auto. }
+assert (Hends_strong := @endpoints_preserved_strong da L c r1 Hsim (inv2_valid Hinv2)).
+feed_n 3 Hends_strong.
+{ rewrite aligned_on_cons_iff. split ; auto ; [].
+  unfold line_contains, r1. now rewrite line_P_iP_min. }
+{ exact Hr2_stay. }
+{ exact Hstg. }
+destruct Hends_strong as [? Hends_strong].
+feed Hends_strong.
+{ case (line_min_le_max L (pos_list c)).
+  + rewrite <-line_P_iP_min, <-line_P_iP_max. fold r1 r2. auto.
+  + intros Er12. exfalso. apply NEr12. unfold r1, r2. now rewrite Er12. }
+apply Nat.le_lteq in Hr1_multR. case Hr1_multR as [Hr1_multR | Hr1_multR].
++ (* We go to phase 1. *)
+  left. exists (L^-P (L^min (pos_list c))). split ; auto.
+  apply weber_majority. rewrite pos_list_length, <-multiplicity_countA. rewrite gt_lt_iff.
+  eapply Nat.le_lt_trans ; [reflexivity | exact Hr1_multR].
++ unfold inv2. 
+  assert (HwebR := @weber_calc_seg_correct (pos_list (round gatherW da c)) (pos_list_non_nil _)).
+  destruct (weber_calc_seg (pos_list (round gatherW da c))) as [Rw1 Rw2].
+  case (Rw1 =?= Rw2) as [ERw | NERw].
+  - (* We go to phase 1. This case is actually impossible, but proving it would be rather tedious. *)
+    left. exists (L^-P (L^min (pos_list c))). split ; auto. split ; auto. intros x Hx. rewrite ERw in *.
+    rewrite HwebR, segment_same in Hx, Hr1_webR. fold r1. now rewrite Hr1_webR, Hx.
+  - (* We stay in phase 2. *) 
+    right. 
+    assert (Hends := @endpoints_preserved da L c r1 Hsim (inv2_valid Hinv2)).
+    feed_n 4 Hends ; auto.
+    { rewrite aligned_on_cons_iff. split.
+      + unfold line_contains, r1. now rewrite line_P_iP_min.
+      + apply Halign. }
+    { apply line_bounds ; auto ; []. rewrite <-multiplicity_pos_InA by now cbn.
+      rewrite Hr1_mult, gt_lt_iff. apply Exp_prop.div2_not_R0. lia. }
+    destruct Hends as [Hmin Hmax].
+    repeat split.
+    * apply aligned_preserved with r1 ; auto. 
+      ++eapply inv2_valid ; eauto.
+      ++rewrite aligned_on_cons_iff. split ; auto ; []. 
+        unfold line_contains, r1. now rewrite line_P_iP_min.
+    * exact NERw.
+    * rewrite Hmin. fold r1. now rewrite <-Hr1_multR.
+    * rewrite Hmax. fold r2.
+      assert (Hr2_multR : (!! (round gatherW da c))[r2] = (!! c)[r2]).
+      { 
+        rewrite 2 multiplicity_countA. unfold pos_list. rewrite 2 config_list_spec.
+        induction names as [|id l IH] ; [reflexivity|].
+        cbn -[equiv equiv_dec get_location r2]. foldR2. rewrite IH.
+        case ifP_sumbool ; case ifP_sumbool ; intros Hid HidR ; (try reflexivity) ; exfalso.
+        + apply Hid. unfold r2. now rewrite <-Hends_strong.
+        + apply HidR. unfold r2. now rewrite Hends_strong. 
+      }
+      now rewrite Hr2_multR.
+    * rewrite Hmin. apply HRstg.
+    * intros id. rewrite Hmin. fold r1. case (HRstg id) as [Hstay | ->] ; [|now auto].
+      destruct (round gatherW da c id) as [[rs rd] rri]. cbn -[equiv r1 straight_path] in Hstay |- *.
+      now rewrite Hstay, straight_path_same.
+    * intros id. rewrite Hmax. fold r2. intros Hid. unfold r2 in Hid. 
+      rewrite Hends_strong in Hid. fold r2 in Hid.
+      generalize (Hround id). destruct_match.
+      ++case ifP_bool ; cbn zeta.
+        --intros _ ->. intuition.
+        
+        -- intros Hex ->. exfalso. rewrite Hid in Hex. 
+          apply (f_equal negb) in Hex. rewrite forallb_existsb, forallb_forall in Hex.
+          clear Hid id H0.
+          assert (Htowers : forall id, get_location (c id) == r1 \/ get_location (c id) == r2).
+          { 
+            intros id. specialize (Hex (get_location (c id))). feed Hex.
+            { rewrite <-(In_InA_is_leibniz equiv), pos_list_InA by now cbn. now exists id. }
+            rewrite negb_true_iff, andb_false_iff, 2 Rltb_false in Hex.
+            case Hex as [Hex | Hex] ; [left | right] ; apply Rnot_lt_le in Hex.
+            + cut (L^P (get_location (c id)) == L^P r1).
+              { intros Hid. apply (f_equal (L^-P)) in Hid. unfold r1 in Hid.
+                rewrite line_P_iP_min, Halign in Hid ; [exact Hid | rewrite pos_list_InA ; now exists id]. }
+              unfold r1. rewrite line_P_iP_min. apply Rle_antisym ; auto ; [].
+              apply line_bounds ; auto ; []. rewrite pos_list_InA. now exists id.
+            + cut (L^P (get_location (c id)) == L^P r2).
+              { intros Hid. apply (f_equal (L^-P)) in Hid. unfold r2 in Hid.
+                rewrite line_P_iP_max, Halign in Hid ; [exact Hid | rewrite pos_list_InA ; now exists id]. }
+              unfold r2 in Hex |- *. rewrite line_P_iP_max in Hex |- *. apply Rle_antisym ; auto ; [].
+              apply line_bounds ; auto ; []. rewrite pos_list_InA. now exists id.
+          }
+          assert (Hsum : (!! c)[r1] + (!! c)[r2] = n).
+          { 
+            rewrite <-(pos_list_length c), 2 multiplicity_countA, 2 countA_occ_length_filter.
+            setoid_rewrite (split_filter_perm (equiv_decb r1) (pos_list c)) at 3.
+            rewrite app_length. f_equal. 
+            unfold pos_list. rewrite config_list_spec, map_map, 2 filter_map.
+            apply PermutationA_length with equiv. apply eqlistA_PermutationA. eapply map_eqlistA_compat ; autoclass ; [].
+            apply filter_extensionalityA_compat ; try reflexivity.
+            intros ? id ->. 
+            assert (Heqb_r12 : r1 ==b r2 = false). { rewrite eqb_false_iff. intuition. }
+            case (Htowers id) as [-> | ->].
+            + foldR2. now rewrite eqb_same, eqb_sym, Heqb_r12.
+            + foldR2. now rewrite eqb_same, Heqb_r12.
+          }
+          rewrite Hr1_mult in Hsum.
+          cut (n = Nat.div2 (n+1) + Nat.div2 (n+1)). { lia. }
+          case (Nat.Even_or_Odd n) as [Heven | Hodd].
+          **rewrite even_div2_p1 by auto. now rewrite even_div2.
+          **exfalso. rewrite <-(pos_list_length (round gatherW da c)) in Hodd.
+            apply weber_odd_unique with (w := Rw1) in Hodd.
+            2: rewrite HwebR ; apply segment_start.
+            apply NERw. symmetry. apply Hodd. rewrite HwebR. apply segment_end.
+      ++intros ->.
+        apply Hr2_stay in Hid. revert Hid. fold r2. cbn -[r2 straight_path equiv]. now destruct (c id) as [[s d] ri].
+Qed.
+
+Lemma phase3_transitions c da L w1 w2 : 
+  (* line_dir L <> 0%VS -> *)
+  similarity_da_prop da ->
+  inv3 c L w1 w2 -> 
+  (exists w, inv1 (round gatherW da c) w) \/ 
+  (exists w1' w2', inv3 (round gatherW da c) L w1' w2').
+Proof using lt_0n delta_g0.
+intros Hsim Hinv3. assert (Hvalid := inv3_valid Hinv3). assert (Hinv3' := Hinv3).
+destruct (round_simplify_phase3 Hinv3 Hsim) as [r Hround].
+destruct Hinv3' as [Halign [Hweb [NEw12 [Hmiddle [Hstg [Hends [Hr1_mult Hr2_mult]]]]]]].
+pose (r1 := L^-P (L^min (pos_list c))). 
+pose (r2 := L^-P (L^max (pos_list c))).
+assert (Hr1 : L^P r1 == L^min (pos_list c)). { unfold r1. now rewrite line_P_iP_min. }
+assert (Hr2 : L^P r2 == L^max (pos_list c)). { unfold r2. now rewrite line_P_iP_max. }
+fold r1 r2 in Hround, Hr1_mult, Hr2_mult, Hends, Hstg |- *.
+assert (Hps : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hline_dir : line_dir L =/= 0%VS).
+{ apply line_dir_nonnul with w1 w2.
+  + assumption.
+  + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start].
+  + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end].  }
+assert (Hline_r1 : line_contains L r1). { apply Halign. unfold r1. now apply line_iP_min_InA. }
+assert (Hline_r2 : line_contains L r2). { apply Halign. unfold r2. now apply line_iP_max_InA. }
+assert (HRstg : config_stg (round gatherW da c) (middle r1 r2)).
+{ rewrite Hround. intros i. destruct_match.
+  + cbn zeta. cbn [get_dest get_start]. destruct_match ; intuition.
+  + cbn -[equiv straight_path r1 r2 middle]. specialize (Hstg i). destruct (c i) as [[s d] ri].
+    cbn -[equiv straight_path r1 r2 middle] in Hstg |- *. 
+    case Hstg as [Hstay | Hgo] ; [now left | now right]. }
+assert (NEr12 : r1 =/= r2).   
+{ apply inv234_endpoints_distinct. eauto. }
+assert (Hmiddle_webR : Weber (pos_list (round gatherW da c)) (middle r1 r2)).
+{ apply weber_contract with (pos_list c).
+  + apply config_stg_contract ; assumption. 
+  + rewrite Hweb, (segment_line L).
+    - left. unfold r1, r2. rewrite line_middle. now rewrite line_P_iP by assumption.
+    - apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start].
+    - apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end].
+    - now apply line_contains_middle. }
+assert (Hends_stay := @endpoints_preserved da L c (middle r1 r2) Hsim Hvalid).
+feed_n 4 Hends_stay.
+{ rewrite aligned_on_cons_iff. split ; [apply line_contains_middle|] ; assumption. }
+{ assumption. }
+{ assumption. }
+{ unfold r1, r2. rewrite line_middle. rewrite line_P_iP by assumption. 
+  generalize (line_min_le_max L (pos_list c)). lra. }
+destruct Hends_stay as [Hr1_stay Hr2_stay].
+assert (Hends_strong := @endpoints_preserved_strong da L c (middle r1 r2) Hsim Hvalid).
+feed_n 3 Hends_strong.
+{ rewrite aligned_on_cons_iff. now split ; [apply line_contains_middle | assumption]. }
+{ exact Hends. }
+{ exact Hstg. }
+destruct Hends_strong as [Hends_min Hends_max].
+feed Hends_min.
+{ case (line_min_le_max L (pos_list c)).
+  + unfold r1, r2. rewrite line_middle, <-Hr1, <-Hr2.
+    rewrite line_P_iP by assumption. foldR2. lra.
+  + rewrite <-Hr1, <-Hr2. intros EPr12. exfalso. apply NEr12.
+    apply (f_equal (L^-P)) in EPr12. rewrite Hr1, Hr2 in EPr12. exact EPr12. }
+feed Hends_max.
+{ case (line_min_le_max L (pos_list c)).
+  + unfold r1, r2. rewrite line_middle, <-Hr1, <-Hr2.
+    rewrite line_P_iP by assumption. foldR2. lra.
+  + rewrite <-Hr1, <-Hr2. intros EPr12. exfalso. apply NEr12.
+    apply (f_equal (L^-P)) in EPr12. rewrite Hr1, Hr2 in EPr12. exact EPr12. }
+fold r1 r2 in Hends_max, Hends_min.
+assert (HalignR : aligned_on L (pos_list (round gatherW da c))).
+{ apply aligned_preserved with (middle r1 r2) ; try assumption.
+  rewrite aligned_on_cons_iff. split ; [now apply line_contains_middle | assumption]. }
+assert (HwebR := @weber_calc_seg_correct (pos_list (round gatherW da c)) (pos_list_non_nil _)).
+destruct (@weber_calc_seg (pos_list (round gatherW da c))) as [Rw1 Rw2].
+case (Rw1 =?= Rw2) as [ERw12 | NERw12].
+(* We go to phase 1. *)
++ left. exists (middle r1 r2). split ; [assumption|]. split ; [assumption|].
+  intros x Hx. rewrite HwebR, <-ERw12, segment_same in Hx, Hmiddle_webR.
+  now rewrite Hx, Hmiddle_webR.
+(* We go to phase 2. *)
++ right. 
+  pose (Rw_min := if Rle_dec (L^P Rw1) (L^P Rw2) then Rw1 else Rw2).
+  pose (Rw_max := if Rle_dec (L^P Rw1) (L^P Rw2) then Rw2 else Rw1).
+  assert (HRw_perm : perm_pairA equiv (Rw_min, Rw_max) (Rw1, Rw2)).
+  { unfold Rw_min, Rw_max. case (Rle_dec (L^P Rw1) (L^P Rw2)) ; intros _ ; [now left | now right]. }
+  assert (Hseg : segment Rw_min Rw_max (middle r1 r2)).
+  { rewrite <-segment_eq_iff in HRw_perm. rewrite HRw_perm, <-HwebR. assumption. }
+  assert (Hline_Rw1 : line_contains L Rw1).
+  { apply weber_aligned with (pos_list (round gatherW da c)).
+    + rewrite <-length_zero_iff_nil, pos_list_length. lia.
+    + assumption.
+    + rewrite HwebR. apply segment_start. }
+  assert (Hline_Rw2 : line_contains L Rw2).
+  { apply weber_aligned with (pos_list (round gatherW da c)).
+    + rewrite <-length_zero_iff_nil, pos_list_length. lia.
+    + assumption.
+    + rewrite HwebR. apply segment_end. }
+  assert (HRw_lt : (L^P Rw_min < L^P Rw_max)%R).
+  { unfold Rw_min, Rw_max. case ifP_sumbool.
+    + intros [? | Heq] ; [assumption | exfalso]. apply NERw12.
+      apply (f_equal (L^-P)) in Heq. 
+      rewrite Hline_Rw1, Hline_Rw2 in Heq. exact Heq.
+    + lra. }
+  exists Rw_min, Rw_max. split ; [|split ; [|split ; [|split ; [|split ; [|split]]]]]. 
+  - assumption.
+  - rewrite <-segment_eq_iff in HRw_perm. setoid_rewrite HRw_perm. exact HwebR.
+  - case HRw_perm as [[-> ->] | [-> ->]] ; intuition.
+  - rewrite Hr1_stay, Hr2_stay. rewrite (segment_line L) in Hseg.
+    * unfold r1, r2 in Hseg.
+      (* rewrite (...). seems to freeze. *)
+      repeat rewrite (line_middle L (pos_list c)) in Hseg at 1. 
+      repeat rewrite line_P_iP in Hseg at 1 by assumption.
+      case Hseg as [Hseg | Hseg] ; [apply Hseg | lra].
+    * case HRw_perm as [[-> ->] | [-> ->]] ; assumption.
+    * case HRw_perm as [[-> ->] | [-> ->]] ; assumption.
+    * now apply line_contains_middle.
+  - rewrite Hr1_stay, Hr2_stay. assumption.
+  - split. 
+    * intros id. rewrite Hr1_stay. fold r1. rewrite Hends_min. intros Hid.
+      rewrite (Hround id). destruct_match.
+      ++case ifP_bool ; [|intuition]. 
+        intros Hcond. exfalso. rewrite orb_true_iff, 2 eqb_true_iff in Hcond.
+        rewrite Hid in Hcond. 
+        assert (Hr_Nweb := @inv34_endpoints_not_weber c L w1 w2 Hline_dir). cbn zeta in Hr_Nweb.
+        feed Hr_Nweb. { auto. } destruct Hr_Nweb as [Hr1_Nweb Hr2_Nweb].
+        fold r1 r2 in Hr1_Nweb, Hr2_Nweb.
+        case Hcond as [Hcond | Hcond] ; apply Hr1_Nweb ; rewrite <-Hcond, Hweb ; 
+          [apply segment_start | apply segment_end].
+      ++apply Hends in Hid. fold r1 in Hid. cbn -[r1 equiv].
+        destruct (c id) as [[s d] ri]. cbn -[r1 equiv] in Hid |- *. exact Hid.
+    * intros id. rewrite Hr2_stay. fold r2. rewrite Hends_max. intros Hid.
+      rewrite (Hround id). destruct_match.
+      ++case ifP_bool ; [|intuition]. 
+        intros Hcond. exfalso. rewrite orb_true_iff, 2 eqb_true_iff in Hcond.
+        rewrite Hid in Hcond. 
+        assert (Hr_Nweb := @inv34_endpoints_not_weber c L w1 w2 Hline_dir). cbn zeta in Hr_Nweb.
+        feed Hr_Nweb. { auto. } destruct Hr_Nweb as [Hr1_Nweb Hr2_Nweb].
+        fold r1 r2 in Hr1_Nweb, Hr2_Nweb.
+        case Hcond as [Hcond | Hcond] ; apply Hr2_Nweb ; rewrite <-Hcond, Hweb ; 
+          [apply segment_start | apply segment_end].
+      ++apply Hends in Hid. fold r2 in Hid. cbn -[r2 equiv].
+        destruct (c id) as [[s d] ri]. cbn -[r2 equiv] in Hid |- *. exact Hid.
+  - split.
+    * rewrite Hr1_stay. fold r1.  
+      assert (Hr1_multR : (!! (round gatherW da c))[r1] = (!! c)[r1]).
+      { 
+        rewrite 2 multiplicity_countA. unfold pos_list. rewrite 2 config_list_spec.
+        induction names as [|id l IH] ; [reflexivity|].
+        cbn -[equiv equiv_dec get_location r1]. foldR2. rewrite IH.
+        case ifP_sumbool ; case ifP_sumbool ; intros Hid HidR ; (try reflexivity) ; exfalso.
+        + apply Hid. now rewrite <-Hends_min.
+        + apply HidR. now rewrite Hends_min. 
+      }
+      now rewrite Hr1_multR.
+    * rewrite Hr2_stay. fold r2.  
+      assert (Hr2_multR : (!! (round gatherW da c))[r2] = (!! c)[r2]).
+      { 
+        rewrite 2 multiplicity_countA. unfold pos_list. rewrite 2 config_list_spec.
+        induction names as [|id l IH] ; [reflexivity|].
+        cbn -[equiv equiv_dec get_location r2]. foldR2. rewrite IH.
+        case ifP_sumbool ; case ifP_sumbool ; intros Hid HidR ; (try reflexivity) ; exfalso.
+        + apply Hid. now rewrite <-Hends_max.
+        + apply HidR. now rewrite Hends_max. 
+      }
+      now rewrite Hr2_multR.
+Qed.
+
+Lemma phase4_transitions c da L w1 w2 : 
+  similarity_da_prop da ->
+  inv4 c L w1 w2 -> 
+  (exists w, inv1 (round gatherW da c) w) \/ 
+  (* Notice that w1 doesn't change. *)
+  (exists w2', inv4 (round gatherW da c) L w1 w2').
+Proof using lt_0n delta_g0. 
+intros Hsim Hinv4. assert (Hvalid := inv4_valid Hinv4). assert (Hinv4' := Hinv4).
+destruct (round_simplify_phase4 Hinv4 Hsim) as [r Hround].
+destruct Hinv4' as [Halign [Hweb [Hmiddle [Hstg [Hends [Hr1_mult Hr2_mult]]]]]].
+pose (r1 := L^-P (L^min (pos_list c))). 
+pose (r2 := L^-P (L^max (pos_list c))).
+assert (Hr1 : L^P r1 == L^min (pos_list c)). { unfold r1. now rewrite line_P_iP_min. }
+assert (Hr2 : L^P r2 == L^max (pos_list c)). { unfold r2. now rewrite line_P_iP_max. }
+fold r1 r2 in Hround, Hr1_mult, Hr2_mult, Hends, Hstg |- *.
+assert (Hps : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (NEw12 : w1 =/= w2).
+{ intros Ew12. rewrite Ew12 in Hmiddle. lra. }
+assert (Hline_dir : line_dir L =/= 0%VS).
+{ apply line_dir_nonnul with w1 w2.
+  + assumption.
+  + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start].
+  + apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end].  }
+assert (Hline_w1 : line_contains L w1). 
+{ apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_start]. }
+assert (Hline_w2 : line_contains L w2). 
+{ apply weber_aligned with (pos_list c) ; [assumption .. | rewrite Hweb ; apply segment_end]. }
+assert (HRstg : config_stg (round gatherW da c) w1).
+{ rewrite Hround. intros i. destruct_match.
+  + cbn zeta. cbn [get_start get_dest]. destruct_match ; intuition.
+  + cbn -[equiv straight_path r1 r2 middle]. specialize (Hstg i). destruct (c i) as [[s d] ri].
+    cbn -[equiv straight_path r1 r2 middle] in Hstg |- *. 
+    case Hstg as [Hstay | Hgo] ; [now left | now right]. }
+assert (NEr12 : r1 =/= r2).
+{ apply inv234_endpoints_distinct. eauto. }
+assert (Hw1_webR : Weber (pos_list (round gatherW da c)) w1).
+{ apply weber_contract with (pos_list c).
+  + apply config_stg_contract ; assumption. 
+  + rewrite Hweb, (segment_line L) ; [|assumption ..].
+    left. lra. }
+assert (Hends_stay := @endpoints_preserved da L c w1 Hsim Hvalid).
+feed_n 4 Hends_stay.
+{ rewrite aligned_on_cons_iff. intuition. }
+{ assumption. }
+{ assumption. }
+{ apply line_bounds. now apply (@weber_segment_InA _ w1 w2). }
+destruct Hends_stay as [Hr1_stay Hr2_stay].
+assert (Hends_strong := @endpoints_preserved_strong da L c w1 Hsim Hvalid).
+feed_n 3 Hends_strong.
+{ rewrite aligned_on_cons_iff. intuition. }
+{ exact Hends. }
+{ exact Hstg. }
+destruct Hends_strong as [Hends_min Hends_max].
+feed Hends_min.
+{ destruct (line_bounds L (pos_list c) w1) as [Hbounds _] ; [now apply (@weber_segment_InA _ w1 w2)|].
+  case Hbounds ; [intuition|]. rewrite <-Hr1. intros Heq. exfalso. 
+  cut ( Weber (pos_list c) r1). { unfold r1. apply (@inv34_endpoints_not_weber c L w1 w2 Hline_dir). now right. }
+  apply (f_equal (L^-P)) in Heq. rewrite Hr1 in Heq. fold r1 in Heq. rewrite Hline_w1 in Heq.
+  rewrite Heq, Hweb. apply segment_start. }
+feed Hends_max.
+{ destruct (line_bounds L (pos_list c) w1) as [_ Hbounds] ; [now apply (@weber_segment_InA _ w1 w2)|].
+  case Hbounds ; [intuition|]. rewrite <-Hr2. intros Heq. exfalso. 
+  cut ( Weber (pos_list c) r2). { unfold r2. apply (@inv34_endpoints_not_weber c L w1 w2 Hline_dir). now right. }
+  apply (f_equal (L^-P)) in Heq. rewrite Hr2 in Heq. fold r2 in Heq. rewrite Hline_w1 in Heq.
+  rewrite <-Heq, Hweb. apply segment_start. }
+fold r1 r2 in Hends_max, Hends_min.
+assert (HalignR : aligned_on L (pos_list (round gatherW da c))).
+{ apply aligned_preserved with w1 ; try assumption.
+  rewrite aligned_on_cons_iff. intuition. }
+assert (HwebR := @weber_calc_seg_correct (pos_list (round gatherW da c)) (pos_list_non_nil _)).
+destruct (weber_calc_seg (pos_list (round gatherW da c))) as [Rw1 Rw2].
+case (Rw1 =?= Rw2) as [ERw12 | NERw12].
+(* We go to phase 1. *)
++ left. exists w1. split ; [assumption|]. split ; [assumption|].
+  intros x Hx. rewrite HwebR, <-ERw12, segment_same in Hx, Hw1_webR.
+  now rewrite Hx, Hw1_webR.
+(* We go to phase 2. *)
++ right. 
+  pose (Rw_min := if Rle_dec (L^P Rw1) (L^P Rw2) then Rw1 else Rw2).
+  pose (Rw_max := if Rle_dec (L^P Rw1) (L^P Rw2) then Rw2 else Rw1).
+  assert (HRw_perm : perm_pairA equiv (Rw_min, Rw_max) (Rw1, Rw2)).
+  { unfold Rw_min, Rw_max. case (Rle_dec (L^P Rw1) (L^P Rw2)) ; intros _ ; [now left | now right]. }
+  assert (Hline_Rw1 : line_contains L Rw1).
+  { apply weber_aligned with (pos_list (round gatherW da c)).
+    + rewrite <-length_zero_iff_nil, pos_list_length. lia.
+    + assumption.
+    + rewrite HwebR. apply segment_start. }
+  assert (Hline_Rw2 : line_contains L Rw2).
+  { apply weber_aligned with (pos_list (round gatherW da c)).
+    + rewrite <-length_zero_iff_nil, pos_list_length. lia.
+    + assumption.
+    + rewrite HwebR. apply segment_end. }
+  assert (HRw_lt : (L^P Rw_min < L^P Rw_max)%R).
+  { unfold Rw_min, Rw_max. case ifP_sumbool.
+    + intros [? | Heq] ; [assumption | exfalso]. apply NERw12.
+      apply (f_equal (L^-P)) in Heq. 
+      rewrite Hline_Rw1, Hline_Rw2 in Heq. exact Heq.
+    + lra. }
+  assert (Hw1_stay : Rw_min == w1).
+  { 
+    cut (L^P Rw_min == L^P w1). 
+    { intros Heq. apply (f_equal (L^-P)) in Heq. revert Heq.
+      unfold Rw_min. destruct_match ; [now rewrite Hline_w1, Hline_Rw1 | now rewrite Hline_w1, Hline_Rw2]. }
+    apply Rle_antisym.
+    + rewrite HwebR, (segment_line L) in Hw1_webR by assumption.
+      unfold Rw_min. case ifP_sumbool ; case Hw1_webR ; foldR2 ; lra.
+    + assert (HRw_min_webR : Weber (pos_list (round gatherW da c)) Rw_min).
+      { unfold Rw_min. destruct_match ; rewrite HwebR ; [apply segment_start | apply segment_end]. }
+      rewrite (@weber_contract_strong (pos_list c) _ w1) in HRw_min_webR.
+      - destruct HRw_min_webR as [HRw_min_web _]. rewrite Hweb, (segment_line L) in HRw_min_web ; [|assumption ..|].
+        * destruct HRw_min_web as [H1 | H2] ; [tauto | lra]. 
+        * unfold Rw_min. destruct_match ; assumption.
+      - apply config_stg_contract ; assumption.
+      - rewrite Hweb. apply segment_start.
+  }
+  exists Rw_max. split ; [|split ; [|split ; [|split ; [|split]]]].
+  - assumption.
+  - rewrite <-segment_eq_iff in HRw_perm. rewrite <-Hw1_stay. setoid_rewrite HRw_perm. exact HwebR.
+  - rewrite <-Hw1_stay, Hr1_stay, Hr2_stay. split ; [|assumption]. rewrite Hw1_stay. apply Hmiddle.
+  - assumption.
+  - assert (Hr_Nweb := @inv34_endpoints_not_weber c L w1 w2 Hline_dir). cbn zeta in Hr_Nweb.
+    feed Hr_Nweb. { auto. } destruct Hr_Nweb as [Hr1_Nweb Hr2_Nweb].
+    fold r1 r2 in Hr1_Nweb, Hr2_Nweb.
+    split. 
+    * intros id. rewrite Hr1_stay. fold r1. rewrite Hends_min. intros Hid.
+      rewrite (Hround id). destruct_match.
+      ++case ifP_bool ; [|intuition]. 
+        intros Hcond. exfalso. rewrite eqb_true_iff, Hid in Hcond.
+        apply Hr1_Nweb. rewrite Hcond, Hweb. apply segment_end.
+      ++apply Hends in Hid. fold r1 in Hid. cbn -[r1 equiv].
+        destruct (c id) as [[s d] ri]. cbn -[r1 equiv] in Hid |- *. exact Hid.
+    * intros id. rewrite Hr2_stay. fold r2. rewrite Hends_max. intros Hid.
+      rewrite (Hround id). destruct_match.
+      ++case ifP_bool ; [|intuition]. 
+        intros Hcond. exfalso. rewrite eqb_true_iff, Hid in Hcond.
+        apply Hr2_Nweb. rewrite Hcond, Hweb. apply segment_end.
+      ++apply Hends in Hid. fold r2 in Hid. cbn -[r2 equiv].
+        destruct (c id) as [[s d] ri]. cbn -[r2 equiv] in Hid |- *. exact Hid.
+  - split.
+    * rewrite Hr1_stay. fold r1.  
+      assert (Hr1_multR : (!! (round gatherW da c))[r1] = (!! c)[r1]).
+      { 
+        rewrite 2 multiplicity_countA. unfold pos_list. rewrite 2 config_list_spec.
+        induction names as [|id l IH] ; [reflexivity|].
+        cbn -[equiv equiv_dec get_location r1]. foldR2. rewrite IH.
+        case ifP_sumbool ; case ifP_sumbool ; intros Hid HidR ; (try reflexivity) ; exfalso.
+        + apply Hid. now rewrite <-Hends_min.
+        + apply HidR. now rewrite Hends_min. 
+      }
+      now rewrite Hr1_multR.
+    * rewrite Hr2_stay. fold r2.  
+      assert (Hr2_multR : (!! (round gatherW da c))[r2] = (!! c)[r2]).
+      { 
+        rewrite 2 multiplicity_countA. unfold pos_list. rewrite 2 config_list_spec.
+        induction names as [|id l IH] ; [reflexivity|].
+        cbn -[equiv equiv_dec get_location r2]. foldR2. rewrite IH.
+        case ifP_sumbool ; case ifP_sumbool ; intros Hid HidR ; (try reflexivity) ; exfalso.
+        + apply Hid. now rewrite <-Hends_max.
+        + apply HidR. now rewrite Hends_max. 
+      }
+      now rewrite Hr2_multR.
+Qed.
+
+
+(* This measure counts how many robots are looping on a given point [x]. *)
+Definition measure_loops x c : R :=
+  list_sum (List.map 
+    (fun id => BtoR ((get_dest (c id) ==b get_start (c id)) && (get_location (c id) ==b x))) 
+    names).
+
+(* This measure counts the total distance from a given point [x] to 
+ * the start position of each robot. 
+ * RMK : this is NOT the distance from [x] to the 
+ * current position of each robot. *)
+Definition measure_dist x c : R := 
+  list_sum (List.map 
+    (fun id => dist x (get_start (c id)))
+    names).
+
+(* The measure is the same for each phase, but the point [x] depends on the phase. 
+ * For instance in phase 1 it is the unique weber point. *)
+Definition measure x c : R := measure_dist x c + (INR n - measure_loops x c).
+
+Local Instance measure_compat : Proper (equiv ==> equiv ==> equiv) measure.
+Proof using . 
+intros ? ? H1 ? ? H2. unfold measure.
+assert (Rplus_compat : Proper (equiv ==> equiv ==> equiv) Rplus).
+{ intros ? ? Heq1 ? ? Heq2. now rewrite Heq1, Heq2. }
+assert (Rminus_compat : Proper (equiv ==> equiv ==> equiv) Rminus).
+{ intros ? ? Heq1 ? ? Heq2. now rewrite Heq1, Heq2. } 
+apply Rplus_compat.
++ unfold measure_dist. apply list_sum_compat, eqlistA_PermutationA.
+  f_equiv. intros ? ? H3. now rewrite H1, H2, H3.
++ apply Rminus_compat ; [reflexivity|].
+  unfold measure_loops. apply list_sum_compat, eqlistA_PermutationA.
+  f_equiv. intros ? ? H3. 
+  foldR2.
+  now rewrite H1, H2, H3.
+Qed.
+
+(* The measure is trivially non-negative. *)
+Lemma measure_nonneg x c : (0 <= measure x c)%R.
+Proof using .
+unfold measure. apply Rplus_le_le_0_compat.
++ unfold measure_dist. apply list_sum_ge_0.
+  rewrite Forall_map, Forall_forall. intros id _. apply dist_nonneg.
++ cut (measure_loops x c <= INR n)%R. { lra. }
+  unfold measure_loops. transitivity (list_sum (alls 1%R (length names))).
+  - induction names as [|id l IH].
+    * cbn. reflexivity.
+    * cbn -[get_location]. apply Rplus_le_compat.
+      ++unfold BtoR. case ifP_bool ; lra.
+      ++apply IH.
+  - assert (Hlen : length names = n). { now rewrite names_length, Nat.add_0_r. }
+    rewrite <-Hlen. clear Hlen. induction names as [|id l IH].
+    * cbn. reflexivity.
+    * cbn -[INR]. rewrite S_INR. lra.
+Qed.
+
+Lemma Rplus_le_compat_eps_1 x y x' y' eps : 
+  (x <= x' - eps)%R -> (y <= y')%R -> (x + y <= x' + y' - eps)%R.
+Proof using . lra. Qed.
+
+Lemma Rplus_le_compat_eps_2 x y x' y' eps : 
+  (x <= x')%R -> (y <= y' - eps)%R -> (x + y <= x' + y' - eps)%R.
+Proof using . lra. Qed.
+
+Section MeasureDistLemmas.
+Variables (c : configuration) (da : demonic_action) (x : location).
+Hypothesis (Hsim : similarity_da_prop da).
+Hypothesis (Hflex : flex_da_prop da).
+Hypothesis (Hvalid : ~invalid c).
+Hypothesis (Hstg : config_stg c x).
+
+Lemma dist_decrease id :
+  (dist x (get_start (round gatherW da c id)) <= dist x (get_start (c id)))%R.
+Proof using All.
+destruct (round_simplify Hsim Hvalid) as [r Hround].
+rewrite (Hround id). clear Hround. destruct_match. 
++ cbn -[dist straight_path]. specialize (Hstg id).
+  destruct (c id) as [[s d] ri]. cbn -[equiv straight_path dist] in Hstg |- *.
+  case Hstg as [Hstay | Hgo].
+  - rewrite <-Hstay, straight_path_same. reflexivity.
+  - rewrite Hgo. rewrite straight_path_dist_end, dist_sym. 
+    rewrite <-Rmult_1_l. apply Rmult_le_compat_r ; [apply dist_nonneg|].
+    generalize (ratio_bounds ri). lra.
++ cbn -[dist straight_path]. destruct (c id) as [[s d] ri]. cbn [get_start]. reflexivity.
+Qed.
+
+Corollary measure_dist_decrease : 
+  (measure_dist x (round gatherW da c) <= measure_dist x c)%R.
+Proof using All.
+unfold measure_dist. induction names as [|id l IH].
++ cbn. reflexivity.
++ cbn -[dist]. apply Rplus_le_compat ; [apply dist_decrease | apply IH].
+Qed.
+
+Lemma dist_decrease_strong id :
+  activate da id = true -> 
+  get_dest (c id) =/= get_start (c id) ->
+  get_location (c id) =/= x ->
+  (dist x (get_start (round gatherW da c id)) <= dist x (get_start (c id)) - delta)%R.
+Proof using All.
+intros Hact HNloop Hloc. specialize (Hflex id c Hact).
+destruct (round_simplify Hsim Hvalid) as [r Hround]. rewrite (Hround id). clear Hround. 
+foldR2. change FrameChoiceSimilarity with FrameC in *.
+case ifP_bool ; [intros _ | rewrite Hact ; discriminate].
+cbn [get_start]. specialize (Hstg id).
+destruct (c id) as [[s d] ri]. cbn -[dist equiv straight_path] in HNloop, Hloc, Hflex, Hstg |- *.
+destruct Hstg as [Hstay | Hgo] ; [tauto|].
+rewrite Hgo in *. case Hflex as [Hreached | Hdelta] ; [tauto|].
+transitivity (dist s x - dist s (straight_path s x ri))%R.
++ rewrite straight_path_dist_end, straight_path_dist_start, Rmult_minus_distr_r, Rmult_1_l. 
+  repeat rewrite (dist_sym w s). reflexivity.
++ rewrite dist_sym. unfold Rminus. apply Rplus_le_compat_l, Ropp_le_contravar. exact Hdelta.
+Qed.
+
+Corollary measure_dist_decrease_strong id : 
+  activate da id = true -> 
+  get_dest (c id) =/= get_start (c id) ->
+  get_location (c id) =/= x ->
+  (measure_dist x (round gatherW da c) <= measure_dist x c - delta)%R.
+Proof using All.
+intros Hact HNloop Hloc. unfold measure_dist. 
+assert (Hin := In_names id). induction names as [|id' l IH].
++ exfalso. inv Hin.
++ cbn -[dist]. case Hin as [-> | Hin].
+  - apply Rplus_le_compat_eps_1.
+    * apply dist_decrease_strong ; assumption.
+    * clear IH. induction l as [|id' l IH] ; [reflexivity | cbn -[dist]].
+      now apply Rplus_le_compat ; [apply dist_decrease | apply IH].
+  - now apply Rplus_le_compat_eps_2 ; [apply dist_decrease | apply IH].
+Qed.
+    
+End MeasureDistLemmas.
+
+Section MeasureLoopsLemmas.
+Variables (c : configuration) (da : demonic_action) (x : location).
+Hypothesis (Hsim : similarity_da_prop da).
+Hypothesis (Hflex : flex_da_prop da).
+Hypothesis (Hvalid : ~invalid c).
+Hypothesis (Hstg : config_stg c x).
+Hypothesis (HstgR : config_stg (round gatherW da c) x).
+
+Lemma loops_increase id : 
+  (BtoR ((get_dest (round gatherW da c id) ==b get_start (round gatherW da c id)) && (get_location (round gatherW da c id) ==b x)) >=  
+  BtoR ((get_dest (c id) ==b get_start (c id)) && (get_location (c id) ==b x)))%R.
+Proof using All. 
+specialize (HstgR id). apply Rle_ge. apply BtoR_le. 
+rewrite 2 andb_true_iff, 4 eqb_true_iff. intros [Hloop Hloc].
+destruct (round_simplify Hsim Hvalid) as [r Hround]. 
+rewrite (Hround id) at 2 3. rewrite (Hround id) in HstgR at 2. clear Hround.
+destruct_match.
++ cbn [get_start get_dest] in HstgR |- *. rewrite get_location_ratio_0.
+  split ; [|assumption]. case HstgR as [-> | ->] ; intuition.
++ split.
+  - case HstgR as [-> | ->] ; [reflexivity|]. 
+    cbn -[equiv straight_path] in Hloop, Hloc |- *. destruct (c id) as [[s d] ri].
+    cbn [get_dest get_start] in Hloop |- *. rewrite Hloop, straight_path_same in Hloc. 
+    now rewrite Hloc.
+  - cbn -[equiv straight_path] in Hloop, Hloc |- *. destruct (c id) as [[s d] ri].
+    cbn [get_dest get_start] in Hloop. now rewrite Hloop, straight_path_same in Hloc |- *.
+Qed. 
+
+Corollary measure_loops_increase : 
+  (measure_loops x (round gatherW da c) >= measure_loops x c)%R.
+Proof using All.
+unfold measure_loops. induction names as [|id l IH].
++ cbn. intuition.
++ cbn -[get_location]. apply Rplus_ge_compat ; [apply loops_increase | apply IH].
+Qed.
+
+Lemma loops_increase_strong id :
+  activate da id = true -> 
+  get_dest (c id) =/= get_start (c id) ->
+  get_location (c id) == x ->
+  (BtoR ((get_dest (round gatherW da c id) ==b get_start (round gatherW da c id)) && (get_location (round gatherW da c id) ==b x)) >=  
+  BtoR ((get_dest (c id) ==b get_start (c id)) && (get_location (c id) ==b x)) + 1)%R.
+Proof using All.
+intros Hact HNloop Hloc. apply Rle_ge. apply BtoR_1_le.
+rewrite andb_false_iff, andb_true_iff, 2 eqb_false_iff, 2 eqb_true_iff.
+split ; [now left|]. specialize (HstgR id).
+destruct (round_simplify Hsim Hvalid) as [r Hround]. 
+rewrite (Hround id) at 2 3. rewrite (Hround id) in HstgR at 2. clear Hround.
+foldR2. change FrameChoiceSimilarity with FrameC in *.
+revert HstgR. case ifP_bool ; [intros _ | rewrite Hact ; discriminate]. cbn [get_start].
+rewrite get_location_ratio_0. intros HstgR. split ; [|assumption].
+case HstgR as [-> | ->] ; intuition.
+Qed.
+
+Corollary measure_loops_increase_strong id : 
+  activate da id = true -> 
+  get_dest (c id) =/= get_start (c id) ->
+  get_location (c id) == x ->
+  (measure_loops x (round gatherW da c) >= measure_loops x c + 1)%R.
+Proof using All.
+intros Hact HNloop Hloc. unfold measure_loops. 
+assert (Hin := In_names id). induction names as [|id' l IH].
++ exfalso. inv Hin.
++ cbn -[get_location]. case Hin as [-> | Hin].
+  - rewrite Rplus_assoc, (Rplus_comm _ 1), <-Rplus_assoc.
+    apply Rplus_ge_compat.
+    * apply loops_increase_strong ; assumption.
+    * clear IH. induction l as [|id' l IH] ; [intuition | cbn -[get_location]].
+      now apply Rplus_ge_compat ; [apply loops_increase | apply IH].
+  - rewrite Rplus_assoc. now apply Rplus_ge_compat ; [apply loops_increase | apply IH].
+Qed.
+  
+End MeasureLoopsLemmas.
+
+Lemma phase1_decrease c da w : 
+  similarity_da_prop da ->
+  flex_da_prop da ->
+  inv1 c w -> 
+  inv1 (round gatherW da c) w ->
+  (exists id, activate da id = true /\
+              get_dest (c id) =/= get_start (c id)) ->
+  (measure w (round gatherW da c) <= measure w c - Rmin 1%R delta)%R.
+Proof using lt_0n delta_g0.
+intros Hsim Hflex Hinv1 Hinv1R [id [Hact HNloop]].
+assert (Hvalid := inv1_valid Hinv1).
+unfold measure. case (get_location (c id) =?= w) as [Hloc | Hloc].
++ apply Rplus_le_compat_eps_2. 
+  - apply measure_dist_decrease ; solve [assumption | apply Hinv1].
+  - cut (measure_loops w (round gatherW da c) >= measure_loops w c + 1)%R.
+    { generalize (Rmin_l 1%R delta). lra. }
+    apply measure_loops_increase_strong with id ; solve [assumption | apply Hinv1 | apply Hinv1R].
++ apply Rplus_le_compat_eps_1.
+  - cut (measure_dist w (round gatherW da c) <= measure_dist w c - delta)%R.
+    { generalize (Rmin_r 1%R delta). lra. }
+    apply measure_dist_decrease_strong with id ; solve [assumption | apply Hinv1].
+  - unfold Rminus. apply Rplus_le_compat ; [reflexivity|]. apply Ropp_le_contravar, Rge_le.
+    apply measure_loops_increase ; solve [assumption | apply Hinv1 | apply Hinv1R].
+Qed.
+
+Lemma phase2_decrease c da L : 
+  let r1 := L^-P (L^min (pos_list c)) in 
+  similarity_da_prop da ->
+  flex_da_prop da ->
+  inv2 c L -> 
+  inv2 (round gatherW da c) L ->
+  (exists id, activate da id = true /\
+              get_dest (c id) =/= get_start (c id)) ->
+  (measure r1 (round gatherW da c) <= measure r1 c - Rmin 1%R delta)%R.
+Proof using lt_0n delta_g0.
+cbn zeta. pose (r1 := L^-P (L^min (pos_list c))). fold r1.
+intros Hsim Hflex Hinv2 Hinv2R [id [Hact HNloop]].
+assert (Hvalid := inv2_valid Hinv2).
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hends := @endpoints_preserved da L c r1 Hsim Hvalid).
+feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [apply Hinv2 | now apply Hinv2]. 
+  unfold r1. apply line_iP_min_InA ; [assumption | apply Hinv2]. }
+{ apply Hinv2. }
+{ apply Hinv2. }
+{ apply line_bounds. unfold r1. apply line_iP_min_InA ; [assumption | apply Hinv2]. }
+destruct Hends as [Hends_min Hends_max].
+unfold measure. case (get_location (c id) =?= r1) as [Hloc | Hloc].
++ apply Rplus_le_compat_eps_2. 
+  - apply measure_dist_decrease ; solve [assumption | apply Hinv2].
+  - cut (measure_loops r1 (round gatherW da c) >= measure_loops r1 c + 1)%R.
+    { generalize (Rmin_l 1%R delta). lra. }
+    apply measure_loops_increase_strong with id ; 
+      solve [assumption | apply Hinv2 | unfold r1 ; rewrite <-Hends_min ; apply Hinv2R].
++ apply Rplus_le_compat_eps_1.
+  - cut (measure_dist r1 (round gatherW da c) <= measure_dist r1 c - delta)%R.
+    { generalize (Rmin_r 1%R delta). lra. }
+    apply measure_dist_decrease_strong with id ; solve [assumption | apply Hinv2].
+  - unfold Rminus. apply Rplus_le_compat ; [reflexivity|]. apply Ropp_le_contravar, Rge_le.
+    apply measure_loops_increase ; solve [assumption | apply Hinv2 | unfold r1 ; rewrite <-Hends_min ; apply Hinv2R].
+Qed.
+
+Lemma phase3_decrease c da L w1 w2 : 
+  let r1 := L^-P (L^min (pos_list c)) in 
+  let r2 := L^-P (L^max (pos_list c)) in 
+  similarity_da_prop da ->
+  flex_da_prop da ->
+  inv3 c L w1 w2 -> 
+  (exists w1' w2', inv3 (round gatherW da c) L w1' w2') ->
+  (exists id, activate da id = true /\
+              get_dest (c id) =/= get_start (c id)) ->
+  (measure (middle r1 r2) (round gatherW da c) <= measure (middle r1 r2) c - Rmin 1%R delta)%R.
+Proof using lt_0n delta_g0.
+cbn zeta. pose (r1 := L^-P (L^min (pos_list c))). pose (r2 := L^-P (L^max (pos_list c))). 
+fold r1 r2.
+intros Hsim Hflex Hinv3 [w1' [w2' Hinv3R]] [id [Hact HNloop]].
+assert (Hvalid := inv3_valid Hinv3).
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hends := @endpoints_preserved da L c (middle r1 r2) Hsim Hvalid).
+feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [apply line_contains_middle | now apply Hinv3].
+  + unfold r1. apply Hinv3, line_iP_min_InA ; [assumption | apply Hinv3].
+  + unfold r2. apply Hinv3, line_iP_max_InA ; [assumption | apply Hinv3]. }
+{ apply Hinv3. }
+{ apply Hinv3. }
+{ unfold r1, r2. rewrite line_middle. rewrite line_P_iP. 
+  + generalize (line_min_le_max L (pos_list c)). lra.
+  + apply line_dir_nonnul with w1 w2.
+    - apply Hinv3.
+    - apply Hinv3, (@weber_segment_InA _ w1 w2);auto;apply Hinv3.
+    - apply Hinv3, (@weber_segment_InA _ w1 w2);auto;apply Hinv3. }
+destruct Hends as [Hends_min Hends_max].
+unfold measure. case (get_location (c id) =?= (middle r1 r2)) as [Hloc | Hloc].
++ apply Rplus_le_compat_eps_2. 
+  - apply measure_dist_decrease ; solve [assumption | apply Hinv3].
+  - cut (measure_loops (middle r1 r2) (round gatherW da c) >= measure_loops (middle r1 r2) c + 1)%R.
+    { generalize (Rmin_l 1%R delta). lra. }
+    apply measure_loops_increase_strong with id ; try assumption. 
+    * apply Hinv3.
+    * unfold r1, r2. rewrite <-Hends_min, <-Hends_max. apply Hinv3R.
++ apply Rplus_le_compat_eps_1.
+  - cut (measure_dist (middle r1 r2) (round gatherW da c) <= measure_dist (middle r1 r2) c - delta)%R.
+    { generalize (Rmin_r 1%R delta). lra. }
+    apply measure_dist_decrease_strong with id ; solve [assumption | apply Hinv3].
+  - unfold Rminus. apply Rplus_le_compat ; [reflexivity|]. apply Ropp_le_contravar, Rge_le.
+    apply measure_loops_increase ; try assumption.
+    * apply Hinv3.
+    * unfold r1, r2. rewrite <-Hends_min, <-Hends_max. apply Hinv3R.
+Qed.
+
+Lemma inv4_diff c L w1 w2: inv4 c L w1 w2 ->  w1 =/= w2.
+Proof using Type.
+  intros Hinv4.
+  intros Ew12. rewrite Ew12 in Hinv4.
+  unfold inv4 in Hinv4.
+  decompose [and] Hinv4; clear Hinv4.
+  lra.
+Qed. 
+
+Lemma inv3_diff c L w1 w2: inv3 c L w1 w2 ->  w1 =/= w2.
+Proof using Type.
+  intros Hinv3.
+  red in Hinv3.
+  decompose [and] Hinv3; clear Hinv3.
+  assumption.
+Qed. 
+
+Lemma phase4_decrease c da L w1 w2 : 
+  let r1 := L^-P (L^min (pos_list c)) in 
+  let r2 := L^-P (L^max (pos_list c)) in 
+  similarity_da_prop da ->
+  flex_da_prop da ->
+  inv4 c L w1 w2 -> 
+  (exists w2', inv4 (round gatherW da c) L w1 w2') ->
+  (exists id, activate da id = true /\
+              get_dest (c id) =/= get_start (c id)) ->
+  (measure w1 (round gatherW da c) <= measure w1 c - Rmin 1%R delta)%R.
+Proof using lt_0n delta_g0.
+cbn zeta. pose (r1 := L^-P (L^min (pos_list c))). pose (r2 := L^-P (L^max (pos_list c))). 
+fold r1 r2.
+intros Hsim Hflex Hinv4 [w1' [w2' Hinv4R]] [id [Hact HNloop]].
+assert (Hvalid := inv4_valid Hinv4).
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hdiff : w1 =/= w2) by apply (inv4_diff Hinv4).
+
+assert (Hends := @endpoints_preserved da L c w1 Hsim Hvalid).
+feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [apply Hinv4 | now apply Hinv4].
+  apply (@weber_segment_InA _ w1 w2);auto;try apply Hinv4. }
+{ apply Hinv4. }
+{ apply Hinv4. }
+
+{ apply line_bounds. apply (@weber_segment_InA _ w1 w2);auto. apply Hinv4. }
+destruct Hends as [Hends_min Hends_max].
+unfold measure. foldR2. case (get_location (c id) =?= w1) as [Hloc | Hloc].
++ apply Rplus_le_compat_eps_2. 
+  - apply measure_dist_decrease ; solve [assumption | apply Hinv4].
+  - cut (measure_loops w1 (round gatherW da c) >= measure_loops w1 c + 1)%R.
+    { generalize (Rmin_l 1%R delta). lra. }
+    apply measure_loops_increase_strong with id ; try assumption. 
+    * apply Hinv4.
+    * apply Hinv4R.
++ apply Rplus_le_compat_eps_1.
+  - cut (measure_dist w1 (round gatherW da c) <= measure_dist w1 c - delta)%R.
+    { generalize (Rmin_r 1%R delta). lra. }
+    apply measure_dist_decrease_strong with id ; solve [assumption | apply Hinv4].
+  - unfold Rminus. apply Rplus_le_compat ; [reflexivity|]. apply Ropp_le_contravar, Rge_le.
+    apply measure_loops_increase ; try assumption.
+    * apply Hinv4.
+    * apply Hinv4R. 
+Qed.
+
+(* This inductive proposition counts how many turns are left before
+ * there is a robot that isn't looping.
+ * This is only used as an intermediate step to prove [FirstMove]. *)
+Inductive FirstMoveWeak : demon -> configuration -> Prop :=
+| FirstMoveWeak_Now : forall d c id, 
+    get_dest (c id) =/= get_start (c id) -> 
+    FirstMoveWeak d c
+| FirstMoveWeak_Later : forall d c,
+    FirstMoveWeak (Stream.tl d) (round gatherW (Stream.hd d) c) -> 
+    FirstMoveWeak d c.
+
+(* This inductive proposition counts how many turns are left before
+ * a robot that isn't looping is activated. *)
+Inductive FirstMove : demon -> configuration -> Prop :=
+| FirstMove_Now : forall d c id, 
+    activate (Stream.hd d) id = true -> 
+    get_dest (c id) =/= get_start (c id) -> 
+    FirstMove d c
+| FirstMove_Later : forall d c,
+    FirstMove (Stream.tl d) (round gatherW (Stream.hd d) c) -> 
+    FirstMove d c.
+
+(* This is what ensures that the measure eventually decreases. *)
+Lemma phase1_progress c d w : 
+  Fair d -> 
+  (Stream.forever (Stream.instant similarity_da_prop)) d ->
+  inv1 c w -> 
+  ~gathered_at w c ->
+  FirstMoveWeak d c.
+Proof using lt_0n delta_g0.
+intros Hfair Hsim Hinv1 HNgather.
+assert (Hid : exists id, get_location (c id) =/= w).
+{ 
+  case (Sumbool.sumbool_of_bool (forallb (fun id => get_location (c id) ==b w) names)).
+  + intros HF. exfalso. apply HNgather. rewrite forallb_forall in HF.
+    setoid_rewrite eqb_true_iff in HF. intros g. apply HF. apply In_names.
+  + intros HF. rewrite <-negb_true_iff, existsb_forallb, existsb_exists in HF.
+    destruct HF as [id [_ Hid]]. rewrite negb_true_iff, eqb_false_iff in Hid.
+    exists id. exact Hid. 
+}
+destruct Hid as [id Hid].
+case (get_dest (c id) =?= get_start (c id)) as [Hloop | HNloop] ;
+  [|now apply FirstMoveWeak_Now with id].
+destruct Hfair as [Hfair_now Hfair_later].
+specialize (Hfair_now id). apply eventually_strengthen in Hfair_now ; [|intros ? ; cbn ; apply bool_dec].
+generalize dependent c.
+induction Hfair_now as [d Hact | d HNact Hlater IH] ; intros c Hinv1 HNgather Hloc Hloop.
++ cbn in Hact. destruct (Hsim) as [Hsim_hd _].
+  destruct (round_simplify_phase1 Hinv1 Hsim_hd) as [r Hround].
+  apply FirstMoveWeak_Later, FirstMoveWeak_Now with id.
+  rewrite (Hround id). foldR2. rewrite Hact. cbn [get_dest get_start]. intuition.
++ cbn in HNact. apply not_true_is_false in HNact. destruct (Hsim) as [Hsim_hd _].
+  destruct (round_simplify_phase1 Hinv1 Hsim_hd) as [r Hround].
+  specialize (Hround id). cbn beta in Hround. foldR2. rewrite HNact in Hround.
+  cbn -[equiv straight_path] in Hround. destruct (c id) as [[start dest] ri].
+  cbn [get_start get_dest] in Hloop.
+  apply FirstMoveWeak_Later, IH.
+  - apply Hfair_later.
+  - apply Hsim.
+  - apply phase1_transitions ; [apply Hsim | exact Hinv1].
+  - intros HgatherR. apply Hloc. rewrite <-(HgatherR (unpack_good id)), good_unpack_good.
+    rewrite Hround. cbn -[equiv straight_path]. now rewrite Hloop, 2 straight_path_same.
+  - intros HlocR. apply Hloc. rewrite Hround in HlocR. cbn -[equiv straight_path] in HlocR |- *.
+    now rewrite Hloop, straight_path_same in HlocR |- *.
+  - rewrite Hround. cbn [get_dest get_start]. assumption.
+Qed.
+
+Lemma phase2_progress c d L : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d ->
+  inv2 c L -> 
+  FirstMoveWeak d c.
+Proof using lt_0n delta_g0.
+intros Hfair Hsim Hinv2. destruct (Hfair) as [Hfair_now _].
+remember (L^-P (L^min (pos_list c))) as r1.
+remember (L^-P (L^max (pos_list c))) as r2.
+remember (L^-P (L^min (L^right r1 (pos_list c)))) as t.
+assert (Ht_InA : InA equiv t (L^right r1 (pos_list c))).
+{ 
+  rewrite Heqt. apply line_iP_min_InA.
+  + cut (InA equiv r2 (L^right r1 (pos_list c))).
+    { intros Hin Hnil. rewrite Hnil in Hin. inv Hin. }
+    rewrite line_right_spec. split.
+    - rewrite Heqr2. apply line_iP_max_InA ; 
+        [rewrite <-length_zero_iff_nil, pos_list_length ; lia | apply Hinv2].
+    - apply Rlt_gt. rewrite Heqr1, Heqr2, line_P_iP_min, line_P_iP_max.
+      case (line_min_le_max L (pos_list c)) as [? | Heq] ; [assumption|].
+      assert (NEr12 : r1 =/= r2).
+      { rewrite Heqr1, Heqr2. apply inv234_endpoints_distinct. auto. }
+      exfalso. apply NEr12. rewrite Heqr1, Heqr2. foldR2. rewrite Heq. reflexivity.
+  + eapply aligned_on_inclA ; [apply line_right_inclA | apply Hinv2].
+}   
+assert (Ht : exists id, t == get_location (c id)).
+{ rewrite <-pos_list_InA. eapply line_right_inclA. apply Ht_InA. }
+destruct Ht as [id Ht].
+revert c Hinv2 Heqr1 Heqr2 Heqt Ht_InA Ht.
+induction (Hfair_now id) as [d Hact | d Hlater IH] ; intros c Hinv2 Heqr1 Heqr2 Heqt Ht_InA Ht.
++ cbn in Hact. apply FirstMoveWeak_Later, FirstMoveWeak_Now with id.
+  destruct (Hsim) as [Hsim_hd _]. destruct (round_simplify_phase2 Hinv2 Hsim_hd) as [r Hround].
+  rewrite (Hround id). clear Hround. foldR2. rewrite Hact. cbn zeta. cbn [get_dest get_start].
+  case ifP_bool.
+  - intros HE. exfalso. rewrite existsb_exists in HE. destruct HE as [x [Hin Hx]].
+    apply (@In_InA _ equiv) in Hin ; autoclass.
+    rewrite andb_true_iff, 2 Rltb_true, <-Ht in Hx.
+    cut (L^P t <= L^P x)%R. { foldR2. lra. }
+    rewrite Heqt, line_P_iP_min. apply line_bounds. rewrite line_right_spec.
+    split ; [assumption|]. rewrite Heqr1, line_P_iP_min. intuition.
+  - intros _. rewrite <-Ht. 
+    cut (L^min (pos_list c) < L^P t)%R. 
+    { intros Hlt Heq. rewrite <-Heq, line_P_iP_min in Hlt. lra. }
+    rewrite line_right_spec in Ht_InA. destruct Ht_InA as [_ Ht_gt].
+    rewrite Heqr1, line_P_iP_min in Ht_gt. foldR2. lra.
++ case (config_stay_or_Nloop c) as [Hstay | [id' HNloop]] ; 
+    [|now apply FirstMoveWeak_Now with id'].
+  destruct (Hsim) as [Hsim_hd _].
+  assert (Hloc := config_stay_location_same Hsim_hd (inv2_valid Hinv2) Hstay).
+  assert (Hps := config_stay_pos_list_same Hsim_hd (inv2_valid Hinv2) Hstay).
+  case (phase2_transitions Hsim_hd Hinv2) as [[w Hinv1R] | Hinv2R]. 
+  { apply FirstMoveWeak_Later, phase1_progress with w.
+    + now destruct Hfair.
+    + now destruct Hsim.
+    + assumption.
+    + unfold gathered_at. setoid_rewrite Hloc. eapply inv234_not_gathered. eauto. }
+  apply FirstMoveWeak_Later, IH.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hfair as [? [? ?]].
+  - assumption.
+  - rewrite Heqr1, Hps. reflexivity.
+  - rewrite Heqr2, Hps. reflexivity.
+  - rewrite Heqt, Hps. reflexivity.
+  - rewrite Hps. assumption.
+  - rewrite Hloc. assumption.
+Qed.
+
+Lemma phase3_progress c d L w1 w2 : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d ->
+  inv3 c L w1 w2 -> 
+  FirstMoveWeak d c.
+Proof using lt_0n delta_g0. 
+intros Hfair Hsim Hinv3. destruct (Hfair) as [Hfair_now _].
+remember (L^-P (L^min (pos_list c))) as r1.
+remember (L^-P (L^max (pos_list c))) as r2.
+assert (Hw12 : exists id, 
+  (get_location (c id) == w1 \/ get_location (c id) == w2) /\
+  get_location (c id) =/= middle r1 r2).
+{ 
+  assert (Hw1 : exists id, w1 == get_location (c id)).
+  { rewrite <-pos_list_InA. apply (@weber_segment_InA _ w1 w2);try apply Hinv3.
+    apply pos_list_non_nil. }
+  assert (Hw2 : exists id, w2 == get_location (c id)).
+  { rewrite <-pos_list_InA. apply (@weber_segment_InA _ w1 w2);try apply  Hinv3.
+    apply pos_list_non_nil. }  
+  assert (Hmiddle : w1 =/= middle r1 r2 \/ w2 =/= middle r1 r2).  
+  { case (w1 =?= middle r1 r2) as [H1 | H1] ; 
+    case (w2 =?= middle r1 r2) as [H2 | H2] ; intuition ; [].
+    exfalso. destruct Hinv3 as [_ [_ [NEw12 _]]]. apply NEw12. etransitivity ; eauto. }
+  case Hmiddle as [H1 | H2].
+  + destruct Hw1 as [id Hw1]. exists id. rewrite <-Hw1. intuition.
+  + destruct Hw2 as [id Hw2]. exists id. rewrite <-Hw2. intuition.
+}
+destruct Hw12 as [id [Hid Hmiddle]].
+revert c w1 w2 Hinv3 Heqr1 Heqr2 Hid Hmiddle.
+induction (Hfair_now id) as [d Hact | d Hlater IH] ; intros c w1 w2 Hinv3 Heqr1 Heqr2 Hid Hmiddle.
++ cbn in Hact. apply FirstMoveWeak_Later, FirstMoveWeak_Now with id.
+  destruct (Hsim) as [Hsim_hd _]. destruct (round_simplify_phase3 Hinv3 Hsim_hd) as [r Hround].
+  rewrite (Hround id). clear Hround. foldR2. rewrite Hact. cbn zeta. cbn [get_dest get_start].
+  rewrite <-Heqr1, <-Heqr2. case ifP_bool.
+  - intros _. intuition.
+  - rewrite orb_false_iff, 2 eqb_false_iff. intuition.
++ case (config_stay_or_Nloop c) as [Hstay | [id' HNloop]] ; 
+    [|now apply FirstMoveWeak_Now with id'].
+  destruct (Hsim) as [Hsim_hd _].
+  assert (Hloc := config_stay_location_same Hsim_hd (inv3_valid Hinv3) Hstay).
+  assert (Hps := config_stay_pos_list_same Hsim_hd (inv3_valid Hinv3) Hstay).
+  case (phase3_transitions Hsim_hd Hinv3) as [[w Hinv1R] | [w1' [w2' Hinv3R]]]. 
+  { apply FirstMoveWeak_Later, phase1_progress with w.
+    + now destruct Hfair.
+    + now destruct Hsim.
+    + assumption.
+    + unfold gathered_at. setoid_rewrite Hloc. eapply inv234_not_gathered. eauto. }
+  assert (Hperm : perm_pairA equiv (w1, w2) (w1', w2')).
+  { rewrite <-segment_eq_iff. intros x. 
+    destruct Hinv3 as [_ [<- _]]. destruct Hinv3R as [_ [<- _]].
+    now rewrite Hps. }
+  apply FirstMoveWeak_Later. apply IH with w1' w2'.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hfair as [? [? ?]].
+  - assumption.
+  - rewrite Heqr1, Hps. reflexivity.
+  - rewrite Heqr2, Hps. reflexivity.
+  - rewrite Hloc. case Hperm as [[-> ->] | [-> ->]] ; intuition.
+  - rewrite Hloc. assumption.
+Qed.
+
+Lemma phase4_progress c d L w1 w2 : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d ->
+  inv4 c L w1 w2 -> 
+  FirstMoveWeak d c.
+Proof using lt_0n delta_g0. 
+intros Hfair Hsim Hinv4. destruct (Hfair) as [Hfair_now _].
+remember (L^-P (L^min (pos_list c))) as r1.
+remember (L^-P (L^max (pos_list c))) as r2.
+assert (Hid : exists id, w2 == get_location (c id)).
+{ rewrite <-pos_list_InA. apply (@weber_segment_InA _ w1 w2);try apply Hinv4.
+  - apply pos_list_non_nil.
+  - apply (inv4_diff Hinv4). }
+destruct Hid as [id Hid].
+revert c w2 Hinv4 Heqr1 Heqr2 Hid.
+induction (Hfair_now id) as [d Hact | d Hlater IH] ; intros c w2 Hinv4 Heqr1 Heqr2 Hid.
++ cbn in Hact. apply FirstMoveWeak_Later, FirstMoveWeak_Now with id.
+  destruct (Hsim) as [Hsim_hd _]. destruct (round_simplify_phase4 Hinv4 Hsim_hd) as [r Hround].
+  rewrite (Hround id). clear Hround. foldR2. rewrite Hact. cbn zeta. cbn [get_dest get_start].
+  case ifP_bool.
+  - intros _. rewrite <-Hid. intros Ew12. destruct Hinv4 as [_ [_ [[_ Hw12] _]]].
+    rewrite Ew12 in Hw12. lra.
+  - rewrite eqb_false_iff. intuition. 
++ case (config_stay_or_Nloop c) as [Hstay | [id' HNloop]] ; 
+    [|now apply FirstMoveWeak_Now with id'].
+  destruct (Hsim) as [Hsim_hd _].
+  assert (Hloc := config_stay_location_same Hsim_hd (inv4_valid Hinv4) Hstay).
+  assert (Hps := config_stay_pos_list_same Hsim_hd (inv4_valid Hinv4) Hstay).
+  case (phase4_transitions Hsim_hd Hinv4) as [[w Hinv1R] | [w2' Hinv4R]]. 
+  { apply FirstMoveWeak_Later, phase1_progress with w.
+    + now destruct Hfair.
+    + now destruct Hsim.
+    + assumption.
+    + unfold gathered_at. setoid_rewrite Hloc. eapply inv234_not_gathered. eauto. }
+  assert (Hw2' : w2' == w2).
+  {
+    assert (Hperm : perm_pairA equiv (w1, w2) (w1, w2')).
+    { rewrite <-segment_eq_iff. intros x. 
+      destruct Hinv4 as [_ [<- _]]. destruct Hinv4R as [_ [<- _]].
+      now rewrite Hps. }
+    case Hperm as [[_ ->] | [_ Ew12]] ; [reflexivity | exfalso].
+    destruct Hinv4 as [_ [_ [[_ Hw12] _]]]. rewrite Ew12 in Hw12. lra.
+  }
+  apply FirstMoveWeak_Later. apply IH with w2'.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hfair as [? [? ?]].
+  - assumption.
+  - rewrite Heqr1, Hps. reflexivity.
+  - rewrite Heqr2, Hps. reflexivity.
+  - rewrite Hloc, Hw2'. assumption.
+Qed.
+
+(* This is the lemma that bridges the gap between [FirstMoveWeak] and [FirstMove]. *)
+Lemma first_move_strengthen d c : 
+  Fair d -> 
+  (Stream.forever (Stream.instant similarity_da_prop)) d ->
+  (exists w, inv1 c w) \/ (exists L, inv2 c L) \/ (exists L w1 w2, inv3 c L w1 w2) \/ (exists L w1 w2, inv4 c L w1 w2) ->
+  FirstMoveWeak d c -> 
+  FirstMove d c.
+Proof using lt_0n delta_g0.
+pose (inv1234 := fun c => 
+  (exists w, inv1 c w) \/ (exists L, inv2 c L) \/ (exists L w1 w2, inv3 c L w1 w2) \/ (exists L w1 w2, inv4 c L w1 w2)).
+fold (inv1234 c).
+assert (Hpreserve : forall c da, 
+  similarity_da_prop da -> inv1234 c -> inv1234 (round gatherW da c)).
+{ intros c' da Hsim [[w Hinv1] | [[L Hinv2] | [[L [w1 [w2 Hinv3]]] | [L [w1 [w2 Hinv4]]]]]].
+  + left. exists w. apply phase1_transitions ; [apply Hsim | assumption].
+  + case (phase2_transitions Hsim Hinv2) ; [left | right ; left] ; eauto.
+  + case (phase3_transitions Hsim Hinv3) ; [left | right ; right ; left] ; eauto.
+  + case (phase4_transitions Hsim Hinv4) ; [left | right ; right ; right] ; eauto. }
+assert (Hvalid : forall c, inv1234 c -> ~invalid c).
+{ intros c' [[w Hinv1] | [[L Hinv2] | [[L [w1 [w2 Hinv3]]] | [L [w1 [w2 Hinv4]]]]]].
+  + eapply inv1_valid ; eauto.
+  + eapply inv2_valid ; eauto.
+  + eapply inv3_valid ; eauto.
+  + eapply inv4_valid ; eauto. }
+intros Hfair Hsim Hinv HFMweak.
+induction HFMweak as [d c id HNloop | d c Hlater IH].
++ destruct Hfair as [Hfair_now Hfair_later].
+  specialize (Hfair_now id). apply eventually_strengthen in Hfair_now ; [|intros ? ; cbn ; apply bool_dec].
+  generalize dependent c.
+  induction Hfair_now as [d Hact | d HNact Hlater IH] ; intros c Hinv HNloop.
+  - apply FirstMove_Now with id ; [exact Hact | exact HNloop].
+  - cbn in HNact. apply not_true_is_false in HNact. destruct Hsim as [Hsim_hd Hsim_tl]. 
+    destruct (round_simplify Hsim_hd (Hvalid c Hinv)) as [r Hround].
+    specialize (Hround id). cbn beta in Hround. foldR2. rewrite HNact in Hround.
+    apply FirstMove_Later, IH.
+    * apply Hfair_later.
+    * apply Hsim_tl.
+    * now apply Hpreserve.
+    * rewrite Hround. cbn -[equiv complement]. 
+      destruct (c id) as [[start dest] ri]. exact HNloop.
++ apply FirstMove_Later, IH.
+  - apply Hfair.
+  - apply Hsim.
+  - apply Hpreserve ; [apply Hsim | assumption].
+Qed.
+
+Lemma phase1_gathered_preserved c da w : 
+  similarity_da_prop da ->
+  inv1 c w ->
+  gathered_at w c -> 
+  gathered_at w (round gatherW da c).
+Proof using lt_0n delta_g0.
+intros Hsim Hinv1 Hgather g.
+destruct (round_simplify_phase1 Hinv1 Hsim) as [r Hround].
+rewrite (Hround (Good g)). clear Hround. destruct_match.
++ rewrite get_location_ratio_0. apply Hgather.
++ cbn -[equiv straight_path]. destruct Hinv1 as [Hstg Hweb]. 
+  specialize (Hstg (Good g)). specialize (Hgather g). destruct (c (Good g)) as [[s d] ri].
+  cbn [get_dest get_start] in Hstg. cbn -[equiv straight_path] in Hgather.
+  case Hstg as [Hstay | Hgo].
+  - rewrite Hstay, straight_path_same in Hgather |- *. assumption.
+  - rewrite Hgo in Hgather |- *. now apply straight_path_end_advance.
+Qed.      
+
+Lemma phase1_gathered_over c d w : 
+  Stream.forever (Stream.instant similarity_da_prop) d ->
+  inv1 c w ->
+  gathered_at w c -> 
+  Gather w (execute gatherW d c).
+Proof using lt_0n delta_g0.
+revert c d. 
+cofix Hind. intros c d Hsim Hinv1 Hgather. constructor.
++ cbn. assumption.
++ cbn. apply Hind.
+  - apply Hsim.
+  - apply phase1_transitions ; [apply Hsim | assumption].
+  - apply phase1_gathered_preserved ; [apply Hsim | assumption ..].
+Qed.
+
+(* This is the well founded relation we will perform induction on. *)
+Definition lt_config x eps c c' := 
+  (0 <= measure x c <= measure x c' - eps)%R. 
+
+Local Instance lt_config_compat : Proper (equiv ==> equiv ==> equiv ==> equiv ==> iff) lt_config.
+Proof using .
+  intros ? ? H1 ? ? H2 ? ? H3 ? ? H4.
+  unfold lt_config.
+  now repeat progress (rewrite H1 at 1; try rewrite H2 at 1; try rewrite H3 at 1; try rewrite H4 at 1).
+Qed.
+
+(* We proove this using the well-foundedness of lt on nat. *)
+Lemma lt_config_wf x eps : (eps > 0)%R -> well_founded (lt_config x eps).
+Proof using . 
+intros Heps. unfold well_founded. intros c.
+pose (f := fun x : R => Z.to_nat (up (x / eps))).
+remember (f (measure x c)) as k. generalize dependent c. 
+pattern k. apply (well_founded_ind lt_wf). clear k.
+intros k IH c Hk. apply Acc_intro. intros c' Hc'. apply IH with (f (measure x c')) ; auto.
+rewrite Hk ; unfold f ; unfold lt_config in Hc'.
+rewrite <-Z2Nat.inj_lt.
++ apply Zup_lt. unfold Rdiv. rewrite <-(Rinv_r eps) by lra. 
+  rewrite <-Rmult_minus_distr_r. apply Rmult_le_compat_r ; intuition.
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition. 
++ apply up_le_0_compat, Rdiv_le_0_compat ; intuition.
+  transitivity eps ; intuition. 
+  apply (Rplus_le_reg_r (- eps)). rewrite Rplus_opp_r. etransitivity ; eauto.
+Qed.
+
+(* The [phaseXXX_correct] lemmas are essentially long and tedious 
+ * inductions on [FirstMove] and on the measure (using [lt_config]). 
+ * They are direct consequences of the [phaseXXX_progress] and [phaseXXX_decrease] lemmas. *)
+Lemma phase1_correct c d w : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d -> 
+  (Stream.forever (Stream.instant flex_da_prop)) d -> 
+  inv1 c w -> 
+  WillGather (execute gatherW d c).
+Proof using lt_0n delta_g0.
+assert (Hdelta1 : (Rmin 1 delta > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+revert d.
+induction c as [c IH] using (well_founded_ind (lt_config_wf w Hdelta1)).
+intros d Hfair Hsim Hflex Hinv1.
+case (gathered_at_dec c w) as [Hgather | HNgather].
+{ apply Stream.Now. exists w. now apply phase1_gathered_over. }
+assert (Hvalid := inv1_valid Hinv1).
+destruct (Hsim) as [Hsim_hd _].
+assert (Hinv1R := phase1_transitions Hsim_hd Hinv1).
+assert (Hprogress := phase1_progress Hfair Hsim Hinv1 HNgather).
+apply first_move_strengthen in Hprogress ; [|assumption | assumption | left ; eauto].
+clear Hsim_hd.
+induction Hprogress as [d c id Hact HNloop | d c Hlater IH2].
++ apply Stream.Later. apply IH ; try assumption.
+  unfold lt_config. split ; [apply measure_nonneg|].
+  apply phase1_decrease ; try assumption.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eauto.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.  
++ apply Stream.Later. 
+  case (gathered_at_dec (round gatherW (Stream.hd d) c) w) as [HRgather | HRNgather].
+  { apply Stream.Now. exists w. apply phase1_gathered_over ; try assumption. now destruct Hsim. }
+  apply IH2 ; try assumption.
+  - intros c' Hc'. apply IH. unfold lt_config in Hc' |- *.
+    split ; [apply measure_nonneg|].
+    etransitivity ; [apply Hc'|]. unfold Rminus. apply Rplus_le_compat_r.
+    unfold measure. apply Rplus_le_compat.
+    * apply measure_dist_decrease.
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv1_valid. eauto.
+      ++apply Hinv1. 
+    * unfold Rminus. apply Rplus_le_compat ; [reflexivity | apply Ropp_le_contravar].
+      apply Rge_le, measure_loops_increase.
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv1_valid. eauto.
+      ++apply Hinv1.
+      ++apply Hinv1R.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eapply inv1_valid. eauto.
+  - apply phase1_transitions ; [now destruct Hsim as [? [? ?]] | assumption].
+Qed.
+
+Lemma phase2_correct c d L : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d -> 
+  (Stream.forever (Stream.instant flex_da_prop)) d -> 
+  inv2 c L -> 
+  WillGather (execute gatherW d c).
+Proof using lt_0n delta_g0. 
+assert (Hdelta1 : (Rmin 1 delta > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+revert d.
+remember (L^-P (L^min (pos_list c))) as r1.
+induction c as [c IH] using (well_founded_ind (lt_config_wf r1 Hdelta1)).
+intros d Hfair Hsim Hflex Hinv2.
+assert (Hvalid := inv2_valid Hinv2).
+destruct (Hsim) as [Hsim_hd Hsim_tl].
+destruct (Hflex) as [Hflex_hd Hflex_tl].
+destruct (Hfair) as [Hfair_hd Hfair_tl].
+case (phase2_transitions Hsim_hd Hinv2) as [[w Hinv1R] | Hinv2R].
+{ apply Stream.Later. apply phase1_correct with w ; assumption. }
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hends := @endpoints_preserved (Stream.hd d) L c r1 Hsim_hd Hvalid).
+feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [apply Hinv2 | now apply Hinv2]. 
+  rewrite Heqr1. apply line_iP_min_InA ; [assumption | apply Hinv2]. }
+{ apply Hinv2. }
+{ rewrite Heqr1. apply Hinv2. }
+{ apply line_bounds. rewrite Heqr1. apply line_iP_min_InA ; [assumption | apply Hinv2]. }
+destruct Hends as [Hends_min Hends_max].
+assert (Hprogress := phase2_progress Hfair Hsim Hinv2).
+apply first_move_strengthen in Hprogress ; [|assumption | assumption | right ; left ; eauto].
+induction Hprogress as [d c id Hact HNloop | d c Hlater IH2].
++ apply Stream.Later. apply IH ; try assumption.
+  - unfold lt_config. split ; [apply measure_nonneg|].
+    rewrite Heqr1. apply phase2_decrease ; try assumption. exists id. intuition.
+  - now rewrite Hends_min, Heqr1.
++ apply Stream.Later. 
+  destruct (Hsim_tl) as [Hsim_tl_hd ?].
+  case (phase2_transitions Hsim_tl_hd Hinv2R) as [[w Hinv1RR] | Hinv2RR].
+  { apply Stream.Later, phase1_correct with w ; try assumption ; [apply Hfair_tl | apply Hflex_tl]. }    
+  assert (HnilR : pos_list (round gatherW (Stream.hd d) c) <> nil).
+  { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+  assert (HendsR := @endpoints_preserved (Stream.hd (Stream.tl d)) L (round gatherW (Stream.hd d) c) r1 Hsim_tl_hd).
+  feed_n 5 HendsR.
+  { eapply inv2_valid. eauto. }
+  { rewrite aligned_on_cons_iff. split ; [apply Hinv2R | now apply Hinv2R]. 
+    rewrite Heqr1, <-Hends_min. apply line_iP_min_InA ; [assumption | apply Hinv2R]. }
+  { apply Hinv2R. }
+  { rewrite Heqr1, <-Hends_min. apply Hinv2R. }
+  { apply line_bounds. rewrite Heqr1, <-Hends_min. apply line_iP_min_InA ; [assumption | apply Hinv2R]. }
+  destruct HendsR as [HendsR_min HendsR_max].
+  apply IH2 ; try assumption.
+  - now rewrite Hends_min, Heqr1. 
+  - intros c' Hc'. apply IH. unfold lt_config in Hc' |- *.
+    split ; [apply measure_nonneg|].
+    etransitivity ; [apply Hc'|]. unfold Rminus. apply Rplus_le_compat_r.
+    unfold measure. apply Rplus_le_compat.
+    * apply measure_dist_decrease ; (try assumption). rewrite Heqr1. apply Hinv2.
+    * unfold Rminus. apply Rplus_le_compat ; [reflexivity | apply Ropp_le_contravar].
+      apply Rge_le, measure_loops_increase ; (try assumption) ; 
+        [rewrite Heqr1 ; apply Hinv2 | rewrite Heqr1, <-Hends_min ; apply Hinv2R].
+  - eapply inv2_valid. eauto.
+  - apply Hflex_tl.
+  - apply Hflex_tl.
+  - apply Hfair_tl.
+  - apply Hfair_tl.
+Qed.
+
+Lemma phase3_correct c d L w1 w2 : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d -> 
+  (Stream.forever (Stream.instant flex_da_prop)) d -> 
+  inv3 c L w1 w2 -> 
+  WillGather (execute gatherW d c).
+Proof using lt_0n delta_g0. 
+assert (Hdelta1 : (Rmin 1 delta > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+intros Hfair Hsim Hflex Hinv3.
+assert (Hinv3' : exists w1 w2, inv3 c L w1 w2). { eauto. }
+clear Hinv3 w1 w2. revert d Hfair Hsim Hflex Hinv3'.
+remember (L^-P (L^min (pos_list c))) as r1.
+remember (L^-P (L^max (pos_list c))) as r2.
+induction c as [c IH] using (well_founded_ind (lt_config_wf (middle r1 r2) Hdelta1)).
+intros d Hfair Hsim Hflex Hinv3'.
+destruct (Hsim) as [Hsim_hd _]. destruct (Hinv3') as [w1 [w2 Hinv3]].
+case (phase3_transitions Hsim_hd Hinv3) as [[w Hinv1R] | Hinv3R'].
+{ apply Stream.Later. eapply phase1_correct.
+  + now destruct Hfair.
+  + now destruct Hsim.
+  + now destruct Hflex.
+  + eassumption. }
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hends := @endpoints_preserved (Stream.hd d) L c (middle r1 r2) Hsim_hd (inv3_valid Hinv3)).
+clear Hsim_hd. feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [apply line_contains_middle | now apply Hinv3]. 
+  + unfold line_contains. now rewrite Heqr1, line_P_iP_min.
+  + unfold line_contains. now rewrite Heqr2, line_P_iP_max. }
+{ apply Hinv3. }
+{ rewrite Heqr1, Heqr2. apply Hinv3. }
+{ rewrite Heqr1, Heqr2. rewrite line_middle. rewrite line_P_iP.
+  + generalize (line_min_le_max L (pos_list c)). lra.
+  + apply (@line_dir_nonnul_sufficient L (pos_list c) w1 w2);try apply Hinv3.
+    apply pos_list_non_nil. }
+destruct Hends as [Hends_min Hends_max].
+assert (Hprogress := phase3_progress Hfair Hsim Hinv3).
+apply first_move_strengthen in Hprogress ; [|assumption | assumption | right ; right ; left ; eauto].
+clear Hinv3 w1 w2.
+induction Hprogress as [d c id Hact HNloop | d c Hlater IH2].
++ destruct Hinv3' as [w1 [w2 Hinv3]]. apply Stream.Later, IH.
+  - unfold lt_config. split ; [apply measure_nonneg|].
+    rewrite Heqr1, Heqr2. eapply phase3_decrease.
+    * now destruct Hsim.
+    * now destruct Hflex. 
+    * eassumption.
+    * assumption.
+    * eauto. 
+  - now rewrite Hends_min, Heqr1.
+  - now rewrite Hends_max, Heqr2.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eauto.   
++ destruct Hinv3' as [w1 [w2 Hinv3]]. destruct Hinv3R' as [w1' [w2' Hinv3R]]. destruct (Hsim) as [_ [Hsim_tl_hd _]].
+  case (phase3_transitions Hsim_tl_hd Hinv3R) as [[w Hinv1RR] | Hinv3RR'].
+  { apply Stream.Later, Stream.Later, phase1_correct with w.
+    + now destruct Hfair as [? [? ?]].
+    + now destruct Hsim as [? [? ?]].
+    + now destruct Hflex as [? [? ?]].
+    + assumption. }  
+  assert (HnilR : pos_list (round gatherW (Stream.hd d) c) <> nil).
+  { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+  assert (HendsR := @endpoints_preserved (Stream.hd (Stream.tl d)) L (round gatherW (Stream.hd d) c) (middle r1 r2) Hsim_tl_hd).
+  feed_n 5 HendsR.
+  { eapply inv3_valid. eauto. }
+  { rewrite aligned_on_cons_iff. split ; [apply line_contains_middle | now apply Hinv3R]. 
+    + unfold line_contains. now rewrite Heqr1, line_P_iP_min.
+    + unfold line_contains. now rewrite Heqr2, line_P_iP_max. }
+  { apply Hinv3R. }
+  { rewrite Heqr1, Heqr2, <-Hends_min, <-Hends_max. apply Hinv3R. }
+  { rewrite Heqr1, Heqr2, Hends_min, Hends_max. rewrite line_middle. rewrite line_P_iP.
+    + generalize (line_min_le_max L (pos_list c)). lra.
+    +  apply (@line_dir_nonnul_sufficient L (pos_list c) w1 w2);try apply Hinv3.
+       apply  pos_list_non_nil. }
+  destruct HendsR as [HendsR_min HendsR_max].
+  apply Stream.Later, IH2 ; try assumption.
+  - now rewrite Hends_min, Heqr1. 
+  - now rewrite Hends_max, Heqr2. 
+  - intros c' Hc'. apply IH. unfold lt_config in Hc' |- *.
+    split ; [apply measure_nonneg|].
+    etransitivity ; [apply Hc'|]. unfold Rminus. apply Rplus_le_compat_r.
+    unfold measure. apply Rplus_le_compat.
+    * apply measure_dist_decrease ; (try assumption).
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv3_valid. eauto.
+      ++rewrite Heqr1, Heqr2. apply Hinv3.
+    * unfold Rminus. apply Rplus_le_compat ; [reflexivity | apply Ropp_le_contravar].
+      apply Rge_le, measure_loops_increase.
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv3_valid. eauto.
+      ++rewrite Heqr1, Heqr2. apply Hinv3.
+      ++rewrite Heqr1, Heqr2, <-Hends_min, <-Hends_max. apply Hinv3R.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eauto.
+Qed.
+
+Lemma phase4_correct c d L w1 w2 : 
+  Fair d ->
+  (Stream.forever (Stream.instant similarity_da_prop)) d -> 
+  (Stream.forever (Stream.instant flex_da_prop)) d -> 
+  inv4 c L w1 w2 -> 
+  WillGather (execute gatherW d c).
+Proof using lt_0n delta_g0. 
+assert (Hdelta1 : (Rmin 1 delta > 0)%R).
+{ unfold Rmin. destruct_match ; lra. }
+intros Hfair Hsim Hflex Hinv4.
+assert (Hinv4' : exists w2, inv4 c L w1 w2). { eauto. }
+clear Hinv4 w2. revert d Hfair Hsim Hflex Hinv4'.
+induction c as [c IH] using (well_founded_ind (lt_config_wf w1 Hdelta1)).
+intros d Hfair Hsim Hflex Hinv4'.
+destruct (Hsim) as [Hsim_hd _]. destruct (Hinv4') as [w2 Hinv4].
+case (phase4_transitions Hsim_hd Hinv4) as [[w Hinv1R] | Hinv4R'].
+{ apply Stream.Later. eapply phase1_correct.
+  + now destruct Hfair.
+  + now destruct Hsim.
+  + now destruct Hflex.
+  + eassumption. }
+assert (Hnil : pos_list c <> nil).
+{ rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+assert (Hends := @endpoints_preserved (Stream.hd d) L c w1 Hsim_hd (inv4_valid Hinv4)).
+clear Hsim_hd. feed_n 4 Hends.
+{ rewrite aligned_on_cons_iff. split ; [|now apply Hinv4]. 
+  apply Hinv4, (@weber_segment_InA _ w1 w2); try apply Hinv4.
+  - apply pos_list_non_nil.
+  - apply (inv4_diff Hinv4). }
+{ apply Hinv4. }
+{ apply Hinv4. }
+{ apply line_bounds. apply (@weber_segment_InA _ w1 w2);try apply Hinv4.
+  - apply pos_list_non_nil.
+  - apply (inv4_diff Hinv4). }
+destruct Hends as [Hends_min Hends_max].
+assert (Hprogress := phase4_progress Hfair Hsim Hinv4).
+apply first_move_strengthen in Hprogress ; [|assumption | assumption | right ; right ; right ; eauto].
+clear Hinv4 w2.
+induction Hprogress as [d c id Hact HNloop | d c Hlater IH2].
++ destruct Hinv4' as [w2 Hinv3]. apply Stream.Later, IH.
+  - unfold lt_config. split ; [apply measure_nonneg|]. eapply phase4_decrease.
+    * now destruct Hsim.
+    * now destruct Hflex. 
+    * eassumption.
+    * assumption.
+    * eauto. 
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eauto.   
++ destruct Hinv4' as [w2 Hinv4]. destruct Hinv4R' as [w2' Hinv4R]. destruct (Hsim) as [_ [Hsim_tl_hd _]].
+  case (phase4_transitions Hsim_tl_hd Hinv4R) as [[w Hinv1RR] | Hinv4RR'].
+  { apply Stream.Later, Stream.Later, phase1_correct with w.
+    + now destruct Hfair as [? [? ?]].
+    + now destruct Hsim as [? [? ?]].
+    + now destruct Hflex as [? [? ?]].
+    + assumption. }  
+  assert (HnilR : pos_list (round gatherW (Stream.hd d) c) <> nil).
+  { rewrite <-length_zero_iff_nil, pos_list_length. lia. }
+  assert (HendsR := @endpoints_preserved (Stream.hd (Stream.tl d)) L (round gatherW (Stream.hd d) c) w1 Hsim_tl_hd).
+  feed_n 5 HendsR.
+  { eapply inv4_valid. eauto. }
+  { rewrite aligned_on_cons_iff. split ; [apply Hinv4R | now apply Hinv4R].
+    apply (@weber_segment_InA _ w1 w2'); try apply Hinv4R.
+    - apply pos_list_non_nil.
+    - apply (inv4_diff Hinv4R). } 
+  { apply Hinv4R. }
+  { apply Hinv4R. }
+  { apply line_bounds, (@weber_segment_InA _ w1 w2'); try apply Hinv4R.
+    - apply pos_list_non_nil.
+    - apply (inv4_diff Hinv4R). }
+  destruct HendsR as [HendsR_min HendsR_max].
+  apply Stream.Later, IH2 ; try assumption.
+  - intros c' Hc'. apply IH. unfold lt_config in Hc' |- *.
+    split ; [apply measure_nonneg|].
+    etransitivity ; [apply Hc'|]. unfold Rminus. apply Rplus_le_compat_r.
+    unfold measure. apply Rplus_le_compat.
+    * apply measure_dist_decrease ; (try assumption).
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv4_valid. eauto.
+      ++apply Hinv4.
+    * unfold Rminus. apply Rplus_le_compat ; [reflexivity | apply Ropp_le_contravar].
+      apply Rge_le, measure_loops_increase.
+      ++now destruct Hsim.
+      ++now destruct Hflex.
+      ++eapply inv4_valid. eauto.
+      ++apply Hinv4.
+      ++apply Hinv4R.
+  - now destruct Hfair.
+  - now destruct Hsim.
+  - now destruct Hflex.
+  - eauto.
+Qed.
+
+(* This is the main theorem. *)
+Theorem gather_correct c d : 
+  (* The demon is fair. *)
+  Fair d -> 
+  (* The frame changes (chosen by the demon) are similarities centered on the observing robot. *)
+  (Stream.forever (Stream.instant similarity_da_prop)) d -> 
+  (* We are in a flexible setting *)
+  (Stream.forever (Stream.instant flex_da_prop)) d -> 
+  (* Initially, the configuration is valid (i.e. not bivalent). *)
+  ~invalid c ->
+  (* Initially, no robot is moving. *)
+  config_stay c -> 
+  (* The robots will eventually gather and stay gathered at the same point forever. *)
+  WillGather (execute gatherW d c).
+Proof using All.
+intros Hfair Hsim Hflex Hvalid Hstay.
+case (inv_initial Hvalid Hstay) as [Hinv1 | [Hinv2 | [Hinv3 | Hinv4]]].
++ destruct Hinv1 as [w Hinv1]. eapply phase1_correct ; eauto.
++ destruct Hinv2 as [L Hinv2]. eapply phase2_correct ; eauto.
++ destruct Hinv3 as [L [w1 [w2 Hinv3]]]. eapply phase3_correct ; eauto.
++ destruct Hinv4 as [L [w1 [w2 Hinv4]]]. eapply phase4_correct ; eauto.
+Qed.  
+End Gathering.
diff --git a/CaseStudies/Gathering/InR2/Weber/Gather_flex_async_Assumptions.v b/CaseStudies/Gathering/InR2/Weber/Gather_flex_async_Assumptions.v
new file mode 100644
index 0000000000000000000000000000000000000000..32e974421439584250106bb7d39e6d3c2f6a8b5b
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Gather_flex_async_Assumptions.v
@@ -0,0 +1,3 @@
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Gather_flex_async.
+
+Print Assumptions gather_correct.
diff --git a/CaseStudies/Gathering/InR2/Weber/Line.v b/CaseStudies/Gathering/InR2/Weber/Line.v
new file mode 100644
index 0000000000000000000000000000000000000000..e0c88fdfc6704d2a4f4e4309ba16b36f4d343589
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Line.v
@@ -0,0 +1,1416 @@
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file defines lines in a euclidean space. Rather than represent lines as 
+ * functions T -> Prop (where T has a euclidean space structure), as we did for segments,
+ * I chose to represent lines as a data type that is used as a function T -> R giving
+ * the coordinate of a point on the line. Membership of a point in a line is handled by a seperate predicate. 
+ * Useful operations on lines are : 
+ * [L^P x] --- [x] is a point in T and [L] is a line 
+ *         This gives the coordinate of the projection of [x] along [L]. 
+ * [L^-P t] --- [t] is a real number and [L] is a line 
+ *         This gives the point on [L] with coordinate [t]. 
+ * [L^min ps] and [L^max ps] --- [ps] is a list of points in T and [L] is a line 
+ *         This gives the minimum / maximum coordinate of the points in [ps] when projected onto [L].
+ * [L^left x ps], [L^on x ps] and [L^right x ps] --- [x] is a point in T, [ps] is a list of points in T and [L] is a line
+ *         This gives the sublist of points in [ps] that have a coordinate 
+ *         strictly smaller / equal / strictly greater than that of [x] when projected on [L].
+ *         Theses three lists form a partition of [ps]. 
+ * Useful definitions are : 
+ * [line_contains L x] --- [x] is a point in T and [L] is a line
+ * [aligned_on L ps] --- [ps] is a list of points in T and [L] is a line
+ * [aligned ps] := [exists L, aligned_on L ps]
+ * The points [L^-P (L^min ps)] and [L^-P (L^max ps)] are commonly refered to as the endpoints of L/ps. *)
+(**************************************************************************)
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+Require Import SetoidDec.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+Require Import Pactole.Spaces.EuclideanSpace.
+Require Import Pactole.Spaces.R2.
+Require Import Pactole.Util.SetoidDefs.
+Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
+Require Import Pactole.Util.Bijection.
+Require Import Pactole.Spaces.Similarity.
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Utils.
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Segment.
+
+
+Reserved Notation "L '^P'" (at level 1).
+Reserved Notation "L '^-P'" (at level 1).
+Reserved Notation "L '^min'" (at level 1).
+Reserved Notation "L '^max'" (at level 1).
+Reserved Notation "L '^left'" (at level 1).
+Reserved Notation "L '^right'" (at level 1).
+Reserved Notation "L '^on'" (at level 1).
+
+Section Lines.
+
+Context {T : Type} `{EuclideanSpace T}.
+Implicit Types (x y z : T) (ps : list T).
+
+(* A line is NOT defined only by its span : 
+ * its origin, direction and scale matter too 
+ * (and changing any of those changes the associated coordinate function). *)
+Record line := { line_org : T ; line_dir : T }.
+Implicit Types (L : line).
+
+(* Line equivalence is thus not equivalence of spans. *)
+Global Instance line_Setoid : Setoid line.
+simple refine {| equiv L1 L2 := line_org L1 == line_org L2 /\ line_dir L1 == line_dir L2 |} ; autoclass ; [].
+Proof.
+split ; auto. intros L1 L2 L3 [Horg12 Hdir12] [Horg23 Hdir23].
+now rewrite Horg12, Horg23, Hdir12, Hdir23.
+Defined.
+
+Global Instance line_org_compat : Proper (equiv ==> equiv) line_org.
+Proof using Type. intros ? ? H1. apply H1. Qed.
+
+Global Instance line_dir_compat : Proper (equiv ==> equiv) line_dir.
+Proof using Type. intros ? ? H1. apply H1. Qed.
+
+Global Instance build_line_compat : Proper (equiv ==> equiv ==> equiv) Build_line.
+Proof using Type. intros ? ? H1 ? ? H2. cbn. intuition. Qed.
+
+Definition line_proj L : T -> R := 
+  fun x => inner_product (x - line_org L) (line_dir L).
+
+Notation "L '^P'" := (line_proj L).
+
+Global Instance line_proj_compat : 
+  Proper (equiv ==> equiv ==> equiv) line_proj.
+Proof using Type. intros ? ? H1 ? ? H2. unfold line_proj. now rewrite H1, H2. Qed.
+
+Lemma line_proj_spec L x : line_proj L x == inner_product (x - line_org L) (line_dir L).
+Proof using Type. reflexivity. Qed. 
+
+Lemma line_proj_combine L (a b : T) t : 
+  L^P (a + t * (b - a))%VS == (L^P (a) + t * (L^P b - L^P a))%R.
+Proof using Type.
+rewrite 3 line_proj_spec. rewrite Rmult_minus_distr_l, <-2 inner_product_mul_l.
+unfold Rminus. rewrite <-inner_product_opp_l, <-2 inner_product_add_l.
+apply inner_product_compat ; auto ; []. repeat rewrite <-add_assoc. apply add_compat ; auto ; [].
+rewrite add_comm. apply add_compat ; auto ; [].
+rewrite <-mul_opp, <-mul_distr_add. apply mul_compat ; auto ; [].
+repeat rewrite <-add_assoc. apply add_compat ; auto ; [].
+now rewrite opp_distr_add, add_sub.
+Qed.
+
+(* The inverse of a line is a function that takes a coordinate
+ * (a real number) and returns the corresponding point on the line (an element of T). *)
+Definition line_proj_inv L : R -> T := 
+  fun t => ((line_org L) + (t / Rsqr (norm (line_dir L))) * (line_dir L))%VS.
+
+Notation "L '^-P'" := (line_proj_inv L).
+
+Global Instance line_proj_inv_compat : Proper (equiv ==> equiv ==> equiv) line_proj_inv.
+Proof using Type. intros ? ? H1 ? ? H2. unfold line_proj_inv. now rewrite H1, H2. Qed.
+
+Lemma line_proj_inv_spec L t : 
+  L^-P t == ((line_org L) + (t / Rsqr (norm (line_dir L))) * (line_dir L))%VS.
+Proof using Type. reflexivity. Qed.
+
+Lemma line_proj_inv_opp L t : L^-P (-t) == 2 * line_org L - L^-P t.
+Proof using Type. 
+rewrite 2 line_proj_inv_spec. rewrite opp_distr_add, add_assoc. apply add_compat.
++ change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1, <-add_assoc, add_opp, add_origin_r.
+  reflexivity.
++ unfold Rdiv. rewrite <-Ropp_mult_distr_l, minus_morph. reflexivity.
+Qed. 
+
+Lemma line_proj_inv_combine L (t t1 t2 : R) : 
+  L^-P (t1 + t * (t2 - t1))%R == (L^-P t1 + t * (L^-P t2 - L^-P t1))%VS.
+Proof using Type.
+rewrite 3 line_proj_inv_spec.
+assert (Hrewrite : forall a b c, (a + b - (a + c) == b - c)%VS).
+{ intros a b c. now rewrite (add_comm a), <-add_assoc, opp_distr_add, (add_assoc a), add_opp, add_origin_l. }
+rewrite Hrewrite. clear Hrewrite. 
+rewrite <-add_assoc. apply add_compat ; [reflexivity|].
+rewrite mul_distr_add, mul_morph, <-minus_morph, mul_morph.
+rewrite 2 add_morph. apply mul_compat ; [|reflexivity].
+lra.
+Qed.
+
+Lemma line_P_iP L t : line_dir L =/= 0%VS -> L^P (L^-P t) == t.
+Proof using Type.
+intros Hdir. 
+rewrite line_proj_spec, line_proj_inv_spec. 
+rewrite 2 inner_product_add_l, inner_product_opp_l.
+rewrite Rplus_comm, <-Rplus_assoc, Rplus_opp_l, Rplus_0_l.
+unfold Rdiv. rewrite inner_product_mul_l, <-squared_norm_product.
+rewrite Rmult_assoc, Rinv_l, Rmult_1_r ; [reflexivity|].
+intros Hnorm. apply Hdir. rewrite <-Rsqr_0, <-RealNormedSpace.square_norm_equiv in Hnorm by lra.
+now rewrite norm_defined in Hnorm.
+Qed.
+
+Definition line_reverse L := Build_line (line_org L) (-line_dir L).
+
+Global Instance line_reverse_compat : Proper (equiv ==> equiv) line_reverse.
+Proof using Type. intros [org1 dir1] [org2 dir2] [H1 H2]. cbn in H1, H2 |- *. now rewrite H1, H2. Qed.
+
+Lemma line_reverse_proj L x : 
+  (line_reverse L)^P x == (-L^P x)%R.
+Proof using Type.
+destruct L as [org dir]. cbv [line_reverse]. cbn. rewrite 2 line_proj_spec. cbn.
+now rewrite inner_product_opp_r.
+Qed.
+
+Lemma line_reverse_proj_inv L t : 
+  (line_reverse L)^-P t == 2 * line_org L - L^-P t.
+Proof using Type. 
+destruct L as [org dir]. rewrite 2 line_proj_inv_spec. cbn.
+rewrite inner_product_opp_l, inner_product_opp_r, Ropp_involutive.
+rewrite opp_distr_add, add_assoc. apply add_compat.
++ change 2%R with (1 + 1)%R. now rewrite <-add_morph, mul_1, <-add_assoc, add_opp, add_origin_r.
++ now rewrite mul_opp.
+Qed. 
+
+Definition line_contains L x : Prop := L^-P (L^P x) == x.
+
+Lemma line_contains_proj_inv: forall L (t:R),
+    line_dir L =/= 0 -> line_contains L (L^-P t).
+Proof using Type.
+  intros L t h. 
+  red.
+  rewrite line_P_iP;auto.
+Qed.
+
+Global Instance line_contains_compat : 
+  Proper (equiv ==> equiv ==> iff) line_contains.
+Proof using Type. intros ? ? H1 ? ? H2. unfold line_contains. now rewrite H1, H2. Qed.
+
+Lemma line_contains_org L : line_contains L (line_org L).
+Proof using Type.
+unfold line_contains. rewrite line_proj_spec, line_proj_inv_spec. 
+destruct L as [org dir]. cbn. unfold Rdiv.
+rewrite add_opp, inner_product_origin_l, Rmult_0_l, mul_0, add_origin_r.
+reflexivity.
+Qed.
+
+Lemma line_reverse_contains L x : 
+  line_contains L x <-> line_contains (line_reverse L) x.
+Proof using Type.
+unfold line_contains. rewrite 2 line_proj_spec, 2 line_proj_inv_spec. 
+destruct L as [org dir]. cbv [line_reverse]. cbn.
+rewrite inner_product_opp_l, 2 inner_product_opp_r, Ropp_involutive.
+unfold Rdiv. rewrite mul_opp, <-minus_morph, <-Ropp_mult_distr_l, Ropp_involutive.
+reflexivity.
+Qed.
+
+Definition aligned_on L ps : Prop := 
+  forall x, InA equiv x ps -> line_contains L x.
+
+(* [aligned_on] doesn't depent on the order of the points. *)
+Global Instance aligned_on_compat : 
+  Proper (equiv ==> PermutationA equiv ==> iff) aligned_on.
+Proof using Type. intros ? ? H1 ? ? H2. unfold aligned_on. setoid_rewrite H1. now setoid_rewrite H2. Qed.
+
+Lemma aligned_on_reverse L ps : 
+  aligned_on L ps <-> aligned_on (line_reverse L) ps.
+Proof using Type. unfold aligned_on. setoid_rewrite <-line_reverse_contains. reflexivity. Qed.
+
+Lemma aligned_on_cons_iff L x ps : 
+  aligned_on L (x :: ps) <-> line_contains L x /\ aligned_on L ps.
+Proof using Type. 
+split.
++ intros Halign. split.
+  - apply Halign. now left.
+  - intros y Hin. apply Halign. now right.
++ intros [Hx Halign]. intros y Hin. inv Hin.
+  - now rewrite H1.
+  - now apply Halign.
+Qed.
+
+Lemma aligned_on_inclA L ps ps' :
+  inclA equiv ps ps' -> aligned_on L ps' -> aligned_on L ps.
+Proof using Type. intros Hincl Halign x Hin. apply Halign. now apply Hincl. Qed.
+
+Lemma aligned_on_dec L ps : {aligned_on L ps} + {~aligned_on L ps}.
+Proof using Type.
+case (Sumbool.sumbool_of_bool (forallb (fun x => L^-P (L^P x) ==b x) ps)).
++ intros HF. left. rewrite forallb_forall in HF. intros x Hin. 
+  rewrite InA_alt in Hin. unfold line_contains. destruct Hin as [y [-> Hin]].
+  rewrite <-eqb_true_iff. now apply HF.
++ intros HF. right. apply (f_equal negb) in HF. rewrite existsb_forallb, existsb_exists in HF.
+  destruct HF as [x [Hin Hx]]. rewrite negb_true_iff, eqb_false_iff in Hx.
+  intros Halign. apply Hx, Halign, In_InA ; autoclass.
+Qed.
+
+Lemma aligned_on_InA_contains L ps t: InA equiv t ps -> aligned_on L ps -> line_contains L t.
+Proof using Type.
+  intros Ht_InA Halign.
+  destruct (PermutationA_split setoid_equiv Ht_InA) as [ t' ht'].
+  rewrite ht' in Halign.
+  apply aligned_on_cons_iff in Halign.
+  destruct Halign as [Halign1 Halign2].
+  assumption.
+Qed.
+
+Definition aligned ps := exists L, aligned_on L ps.
+
+(* [aligned] doesn't depent on the order of the points. *)
+Global Instance aligned_compat : 
+  Proper (PermutationA equiv ==> iff) aligned.
+Proof using Type. intros ? ? Heq. unfold aligned. now setoid_rewrite Heq. Qed.
+
+Lemma aligned_nil : aligned nil.
+Proof using Type. exists (Build_line 0%VS 0%VS). unfold aligned_on. intros x Hin. exfalso. inversion Hin. Qed.
+
+Lemma aligned_single p0 : aligned (p0 :: nil).
+Proof using Type. 
+exists (Build_line p0%VS p0%VS). intros x Hin. inv Hin. 
++ unfold line_contains. rewrite H1, line_proj_spec, line_proj_inv_spec. cbn. unfold Rdiv.
+  now rewrite add_opp, inner_product_origin_l, Rmult_0_l, mul_0, add_origin_r.
++ inv H1.
+Qed.
+
+Lemma aligned_tail p0 ps : aligned (p0 :: ps) -> aligned ps. 
+Proof using Type. 
+intros [L Halign]. exists L. unfold aligned_on in *.
+intros x Hin. apply Halign. now right.
+Qed.
+
+(* An alternate definition of alignment. *)
+Lemma aligned_spec p0 ps : 
+  aligned (p0 :: ps) <-> exists v, forall p, InA equiv p ps -> exists t, (p == p0 + t * v)%VS.
+Proof using Type. 
+split. 
++ intros [L Halign]. exists (line_dir L). intros p Hin. eexists ?[t].
+  rewrite <-(Halign p), <-(Halign p0) ; [| now left | now right].
+  rewrite 2 line_proj_spec, 2 line_proj_inv_spec. rewrite <-add_assoc.
+  apply add_compat ; [reflexivity|].
+  rewrite add_morph. apply mul_compat ; [|reflexivity].
+  assert (Hrewrite : forall a b c, (a = b + c <-> a - b = c)%R).
+  { intros a b c. lra. }
+  rewrite Hrewrite. eauto.
++ intros [v Hv]. exists (Build_line p0 v). intros x Hin.
+  rewrite InA_cons in Hin. destruct Hin as [-> | Hin] ; [now apply line_contains_org|].
+  apply Hv in Hin. destruct Hin as [t ->]. 
+  case (v =?= 0%VS) as [Ev0 | NEv0].
+  - rewrite Ev0, mul_origin, add_origin_r. apply line_contains_org.
+  - unfold line_contains. rewrite line_proj_spec, line_proj_inv_spec. cbn.
+    apply add_compat ; [reflexivity|]. apply mul_compat ; [|reflexivity].
+    rewrite <-add_assoc, add_sub, inner_product_mul_l, <-squared_norm_product.
+    rewrite sqrt_Rsqr by apply norm_nonneg. unfold Rdiv. 
+    rewrite Rmult_assoc, Rinv_r, Rmult_1_r ; [reflexivity|].
+    rewrite <-Rsqr_0. intros Heq. apply R_sqr.Rsqr_inj in Heq ; [| apply norm_nonneg | reflexivity].
+    rewrite norm_defined in Heq. intuition.
+Qed.
+   
+(* This function calculates a line passing through a set of points. *)
+Fixpoint line_calc ps := 
+  match ps with 
+  | nil => Build_line 0%VS 0%VS 
+  | p0 :: tl1 => 
+    match tl1 with 
+    | nil => Build_line p0 0%VS
+    | p1 :: tl2 => if p0 =?= p1 then line_calc tl1 else Build_line p0 (p1 - p0)
+    end 
+  end.
+
+Lemma line_calc_spec ps : 
+  aligned ps <-> aligned_on (line_calc ps) ps.
+Proof using Type. 
+split ; [intros Halign | intros Halign ; eexists ?[L] ; eassumption].
+remember (line_calc ps) as Lcalc.
+induction ps as [|p0 ps IH]. { intros ? Hin. inv Hin. }
+case ps as [|p1 ps'] eqn:E.
++ cbn in HeqLcalc. rewrite HeqLcalc. intros x Hin. rewrite InA_cons in Hin.
+  destruct Hin as [-> | Hin] ; [|inv Hin]. apply line_contains_org.
++ cbn in HeqLcalc. case (p0 =?= p1) as [Heq | HNeq].
+  - cut (aligned_on Lcalc (p1 :: ps')).
+    { intros Hcut x Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin]. 
+      + rewrite Heq. apply Hcut. now left. 
+      + now apply Hcut. }
+    apply IH.
+    * apply aligned_tail in Halign. assumption.
+    * rewrite HeqLcalc. reflexivity.
+  - clear IH. rewrite aligned_spec in Halign. destruct Halign as [v Hv].
+    intros x Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin].
+    * rewrite HeqLcalc. apply line_contains_org.
+    * apply Hv in Hin. destruct Hin as [t ->]. unfold line_contains.
+      rewrite line_proj_spec, line_proj_inv_spec, HeqLcalc. cbn.
+      apply add_compat ; [reflexivity|].
+      rewrite (add_comm p0), <-add_assoc, add_opp, add_origin_r.
+      rewrite inner_product_mul_l. unfold Rdiv. rewrite Rmult_assoc, <-mul_morph.
+      apply mul_compat ; [reflexivity | clear t].
+      rewrite <-squared_norm_product, sqrt_Rsqr by apply norm_nonneg.
+      cut (exists t, v == t * (p1 - p0)).
+      { intros [t ->]. rewrite inner_product_mul_l, Rmult_assoc, <-mul_morph.
+        rewrite <-squared_norm_product, Rinv_r, mul_1 ; [reflexivity|].
+        rewrite <-Rsqr_0. intros Heq. apply R_sqr.Rsqr_inj in Heq ; [| apply norm_nonneg | reflexivity].
+        apply HNeq. rewrite norm_defined in Heq. apply (add_compat _ _ (reflexivity p0)) in Heq.
+        rewrite add_origin_r, add_sub in Heq. intuition. }
+      specialize (Hv p1). feed Hv. { now left. } destruct Hv as [t Hv].
+      assert (Ht0 : t <> 0%R).
+      { intros Ht. rewrite Ht, mul_0, add_origin_r in Hv. intuition. }
+      exists (Rinv t). apply mul_reg_l with t ; [assumption|].
+      rewrite mul_morph, Rinv_r, mul_1, Hv by assumption.
+      rewrite <-add_assoc, add_sub. reflexivity.
+Qed.
+
+(* We can use the function [line_calc] to prove that alignment is decidable. *)
+Corollary aligned_dec ps : {aligned ps} + {~aligned ps}.
+Proof using Type. 
+case (aligned_on_dec (line_calc ps) ps) as [Halign | HNalign].
++ rewrite <-line_calc_spec in Halign. now left.
++ rewrite <-line_calc_spec in HNalign. now right.
+Qed.  
+
+(* If a line contains two distinct points, then its direction isn't null. *)
+Lemma line_dir_nonnul L x y : 
+  x =/= y ->
+  line_contains L x -> 
+  line_contains L y -> 
+  line_dir L =/= 0%VS.
+Proof using Type. 
+intros NExy Lx Ly Hdir. apply NExy. unfold line_contains in Lx, Ly. 
+rewrite line_proj_spec, Hdir, inner_product_origin_r in Lx, Ly.
+now rewrite <-Lx, <-Ly.
+Qed.
+
+(* The maximum coordinate of [ps] on [L], or 0 if [ps] is empty. *)
+Definition line_max L ps := 
+  match ps with 
+  | nil => 0%R 
+  | p :: ps => fold_left Rmax (map (L^P) ps) (L^P p)
+  end.
+
+Notation "L '^max'" := (line_max L).
+
+Global Instance line_max_compat : Proper (equiv ==> PermutationA equiv ==> equiv) line_max.
+Proof using Type. 
+intros L L' HL ps ps' Hps. unfold line_max. 
+case ps as [|p0 ps] ; case ps' as [|p0' ps'].
++ reflexivity.
++ apply PermutationA_nil in Hps ; autoclass.
++ symmetry in Hps. apply PermutationA_nil in Hps ; autoclass.
++ apply fold_left_Rmax_PermutationA. eapply map_extensionalityA_compat_perm in Hps.
+  - exact Hps.
+  - autoclass.
+  - autoclass.
+  - intros ? ? Heq. now rewrite HL, Heq.
+Qed.   
+  
+(* The minimum coordinate of [ps] on [L], or 0 if [ps] is empty. *)
+Definition line_min L ps := 
+  match ps with 
+  | nil => 0%R 
+  | p :: ps => fold_left Rmin (map (L^P) ps) (L^P p)
+  end.
+
+Notation "L '^min'" := (line_min L).
+
+Global Instance line_min_compat : Proper (equiv ==> PermutationA equiv ==> equiv) line_min.
+Proof using Type. 
+intros L L' HL ps ps' Hps. unfold line_min. 
+case ps as [|p0 ps] ; case ps' as [|p0' ps'].
++ reflexivity.
++ apply PermutationA_nil in Hps ; autoclass.
++ symmetry in Hps. apply PermutationA_nil in Hps ; autoclass.
++ apply fold_left_Rmin_PermutationA. eapply map_extensionalityA_compat_perm in Hps.
+  - exact Hps.
+  - autoclass.
+  - autoclass.
+  - intros ? ? Heq. now rewrite HL, Heq.
+Qed.   
+
+Lemma line_min_spec_nonnil L ps : 
+  ps <> nil -> L^min ps == fold_left Rmin (List.map (L^P) ps) (L^P (nth 0 ps 0%VS)).
+Proof using Type. 
+intros Hnil. unfold line_min. case ps as [|p0 ps] ; [intuition | cbn].
+now rewrite Rmin_left by reflexivity.
+Qed. 
+
+Lemma line_max_spec_nonnil L ps : 
+  ps <> nil -> L^max ps == fold_left Rmax (List.map (L^P) ps) (L^P (nth 0 ps 0%VS)).
+Proof using Type. 
+intros Hnil. unfold line_max. case ps as [|p0 ps] ; [intuition | cbn].
+now rewrite Rmax_left by reflexivity.
+Qed. 
+
+Lemma line_min_le_max L ps : (L^min ps <= L^max ps)%R.
+Proof using Type. 
+case (length ps =?= 0%nat) as [Hlen | Hlen] ; cbn in Hlen.
++ rewrite length_zero_iff_nil in Hlen. rewrite Hlen. 
+  unfold line_min, line_max. reflexivity.
++ rewrite length_zero_iff_nil in Hlen. 
+  rewrite line_min_spec_nonnil, line_max_spec_nonnil by assumption.
+  rewrite fold_left_Rmin_le_start, <-fold_left_Rmax_ge_start. reflexivity.
+Qed.
+
+Lemma line_bounds L ps x : 
+  InA equiv x ps -> (L^min ps <= L^P x <= L^max ps)%R.
+Proof using Type.
+intros Hin. 
+assert (Hnil : ps <> nil). { intros ->. inv Hin. }
+rewrite line_min_spec_nonnil, line_max_spec_nonnil by assumption.
+split.
++ apply fold_left_Rmin_le_InA. right. rewrite InA_map_iff ; autoclass.
++ apply fold_left_Rmax_ge_InA. right. rewrite InA_map_iff ; autoclass.
+Qed.
+
+Lemma line_P_iP_min L ps : L^P (L^-P (L^min ps)) == L^min ps.
+Proof using Type.
+case (line_dir L =?= 0%VS).
++ intros Hdir.
+  assert (HP : forall x, L^P x == 0%R).
+  { intros x. now rewrite line_proj_spec, Hdir, inner_product_origin_r. }
+  rewrite HP. unfold line_min.
+  case ps as [|p0 ps] ; [reflexivity|].
+  induction ps as [|p1 ps IH] ; cbn -[equiv] ; auto.
+  rewrite 2 HP, Rmin_left by lra. rewrite HP in IH. exact IH.
++ intros Hdir. rewrite line_P_iP ; intuition.
+Qed. 
+
+Lemma line_P_iP_max L ps : L^P (L^-P (L^max ps)) == L^max ps.
+Proof using Type.
+case (line_dir L =?= 0%VS).
++ intros Hdir.
+  assert (HP : forall x, L^P x == 0%R).
+  { intros x. now rewrite line_proj_spec, Hdir, inner_product_origin_r. }
+  rewrite HP. unfold line_max.
+  case ps as [|p0 ps] ; [reflexivity|].
+  induction ps as [|p1 ps IH] ; cbn -[equiv] ; auto.
+  rewrite 2 HP, Rmax_left by lra. rewrite HP in IH. exact IH.
++ intros Hdir. rewrite line_P_iP ; intuition.
+Qed. 
+
+Lemma line_iP_min_InA L ps : 
+  ps <> nil -> 
+  aligned_on L ps -> 
+  InA equiv (L^-P (L^min ps)) ps.
+Proof using Type. 
+intros Hnil Halign.
+assert (Hps : eqlistA equiv ps (map (L^-P) (map (L^P) ps))).
+{ 
+  clear Hnil. induction ps as [|p0 ps IH]. 
+  + cbn. reflexivity.
+  + cbn. right.
+    - rewrite Halign ; [reflexivity | now left].
+    - apply IH. intros x Hin. apply Halign. now right. 
+}
+rewrite Hps at 2. clear Hps. rewrite InA_map_iff ; autoclass.
+eexists ?[x]. split ; [reflexivity|]. rewrite line_min_spec_nonnil by assumption.
+assert (Hincl : inclA equiv ((L^P (nth 0 ps 0)) :: map (L^P) ps) (map (L^P) ps)).
+{ intros x Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin] ; [|assumption].
+  rewrite InA_map_iff ; autoclass. eexists ?[x]. split ; [reflexivity|]. 
+  apply nth_InA ; autoclass. cut (length ps <> 0%nat). { lia. } now rewrite length_zero_iff_nil. }
+apply Hincl, fold_left_Rmin_InA.
+Qed.
+
+Lemma line_iP_max_InA L ps : 
+  ps <> nil -> 
+  aligned_on L ps -> 
+  InA equiv (L^-P (L^max ps)) ps.
+Proof using Type. 
+intros Hnil Halign.
+assert (Hps : eqlistA equiv ps (map (L^-P) (map (L^P) ps))).
+{ 
+  clear Hnil. induction ps as [|p0 ps IH]. 
+  + cbn. reflexivity.
+  + cbn. right.
+    - rewrite Halign ; [reflexivity | now left].
+    - apply IH. intros x Hin. apply Halign. now right. 
+}
+rewrite Hps at 2. clear Hps. rewrite InA_map_iff ; autoclass.
+eexists ?[x]. split ; [reflexivity|]. rewrite line_max_spec_nonnil by assumption.
+assert (Hincl : inclA equiv ((L^P (nth 0 ps 0)) :: map (L^P) ps) (map (L^P) ps)).
+{ intros x Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin] ; [|assumption].
+  rewrite InA_map_iff ; autoclass. eexists ?[x]. split ; [reflexivity|]. 
+  apply nth_InA ; autoclass. cut (length ps <> 0%nat). { lia. } now rewrite length_zero_iff_nil. }
+apply Hincl, fold_left_Rmax_InA.
+Qed.
+
+Lemma line_reverse_min L ps : (line_reverse L)^min ps == (-L^max ps)%R.
+Proof using Type. 
+case (length ps =?= 0%nat) as [Hnil | HNnil].
++ cbn in Hnil. rewrite length_zero_iff_nil in Hnil. rewrite Hnil. cbn. 
+  now rewrite Ropp_0.
++ cbn in HNnil. rewrite length_zero_iff_nil in HNnil.
+  rewrite line_min_spec_nonnil, line_max_spec_nonnil by assumption.
+  rewrite fold_left_Rmin_Rmax. f_equiv. f_equiv.
+  - rewrite map_map. eapply (@map_extensionalityA_compat _ _ equiv) ; autoclass ; try reflexivity.
+    intros x y Heq. rewrite line_reverse_proj, Ropp_involutive. now rewrite Heq.
+  - rewrite line_reverse_proj, Ropp_involutive. reflexivity.
+Qed. 
+  
+Lemma line_reverse_iP_min L ps : 
+  (line_reverse L)^-P ((line_reverse L)^min ps) == L^-P (L^max ps).
+Proof using Type.
+rewrite line_reverse_min. rewrite line_reverse_proj_inv, <-line_proj_inv_opp, Ropp_involutive. 
+reflexivity.
+Qed.
+
+Lemma line_reverse_max L ps : (line_reverse L)^max ps == (-L^min ps)%R.
+Proof using Type. 
+case (length ps =?= 0%nat) as [Hnil | HNnil].
++ cbn in Hnil. rewrite length_zero_iff_nil in Hnil. rewrite Hnil. cbn. 
+  now rewrite Ropp_0.
++ cbn in HNnil. rewrite length_zero_iff_nil in HNnil.
+  rewrite line_min_spec_nonnil, line_max_spec_nonnil by assumption.
+  rewrite fold_left_Rmin_Rmax, Ropp_involutive. f_equiv.
+  - rewrite map_map. eapply (@map_extensionalityA_compat _ _ equiv) ; autoclass ; try reflexivity.
+    intros x y Heq. rewrite line_reverse_proj. now rewrite Heq.
+  - rewrite line_reverse_proj. reflexivity.
+Qed. 
+  
+Lemma line_reverse_iP_max L ps : 
+  (line_reverse L)^-P ((line_reverse L)^max ps) == L^-P (L^min ps).
+Proof using Type.
+rewrite line_reverse_max. rewrite line_reverse_proj_inv, <-line_proj_inv_opp, Ropp_involutive. 
+reflexivity.
+Qed.
+
+Lemma line_contains_combine L a b t : 
+  line_contains L a ->
+  line_contains L b -> 
+  line_contains L (a + t * (b - a)).
+Proof using Type. 
+intros Ha Hb. unfold line_contains.
+rewrite line_proj_combine, line_proj_inv_combine, Ha, Hb. reflexivity.
+Qed.
+
+Lemma line_contains_middle L a b : 
+  line_contains L a -> 
+  line_contains L b -> 
+  line_contains L (middle a b).
+Proof using Type. 
+intros Ha Hb.
+assert (Hrewrite : middle a b == a + (1 / 2) * (b - a)).
+{ unfold middle. apply mul_reg_l with 2%R ; [lra|].
+  rewrite mul_morph. unfold Rdiv. rewrite Rmult_1_l, Rinv_r, mul_1 by lra.
+  rewrite mul_distr_add, mul_morph, Rinv_r, mul_1 by lra.
+  change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1, <-add_assoc, add_sub. reflexivity. }
+rewrite Hrewrite. now apply line_contains_combine.
+Qed.
+
+Lemma line_middle L ps : 
+  middle (L^-P (L^min ps)) (L^-P (L^max ps)) == L^-P ((L^min ps + L^max ps) / 2).
+Proof using Type. 
+assert (Hrewrite : forall a b, middle a b == a + (1 / 2) * (b - a)).
+{ intros a b. unfold middle. apply mul_reg_l with 2%R ; [lra|].
+  rewrite mul_morph. unfold Rdiv. rewrite Rmult_1_l, Rinv_r, mul_1 by lra.
+  rewrite mul_distr_add, mul_morph, Rinv_r, mul_1 by lra.
+  change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1, <-add_assoc, add_sub. reflexivity. }
+rewrite Hrewrite. clear Hrewrite.
+rewrite <-line_proj_inv_combine. apply line_proj_inv_compat ; [reflexivity|].
+cbn. lra.
+Qed.
+
+Lemma segment_line L a b x : 
+  line_contains L a ->
+  line_contains L b ->
+  line_contains L x ->
+  (segment a b x <-> (L^P a <= L^P x <= L^P b)%R \/ (L^P b <= L^P x <= L^P a)%R).
+Proof using Type.
+intros Ha Hb Hx. split.
++ intros [t [Ht Hx_seg]].
+  assert (HxP : L^P x == L^P (b + t * (a - b))).
+  { apply line_proj_compat ; [reflexivity|]. rewrite Hx_seg.
+    unfold Rminus. rewrite <-add_morph, mul_1, mul_distr_add, mul_opp, minus_morph.
+    rewrite 2 add_assoc. apply add_compat ; [|reflexivity]. now rewrite add_comm. }
+  rewrite line_proj_combine in HxP.
+  case (Rle_dec (L^P a) (L^P b)) as [Cab | Cab].
+  - left. rewrite HxP. nra.
+  - right. rewrite HxP. nra.
++ intros HxP. case (a =?= b) as [Eab | NEab].
+  { rewrite Eab, segment_same. rewrite Eab in HxP.
+    assert (HeqP : L^P x = L^P b). { lra. }
+    now rewrite <-Hx, <-Hb, HeqP. }
+  assert (NEabP : L^P a <> L^P b).
+  { intros EabP. apply NEab. now rewrite <-Ha, <-Hb, EabP. }
+  case HxP as [HxP | HxP].
+  - unfold segment. pose (t := ((L^P b - L^P x) / (L^P b - L^P a))%R). exists t.
+    assert (Ht : (0 <= t <= 1)%R).
+    { unfold t. split ; [apply Rdiv_le_0_compat | apply Rdiv_le_1] ; lra. }
+    split ; [assumption|].
+    assert (Hrewrite : t * a + (1 - t) * b == b + t * (a - b)).
+    { unfold Rminus. rewrite <-add_morph, mul_1, mul_distr_add, mul_opp, minus_morph.
+      rewrite 2 add_assoc. apply add_compat ; [|reflexivity]. now rewrite add_comm. }
+    rewrite Hrewrite. clear Hrewrite.
+    cut (L^P x == L^P (b + t * (a - b))).
+    { intros HeqP. apply (line_proj_inv_compat _ _ (reflexivity L)) in HeqP.
+      rewrite Hx, line_contains_combine in HeqP by assumption. assumption. }
+    rewrite line_proj_combine. unfold t, Rdiv. rewrite Rmult_assoc.
+    setoid_rewrite <-Ropp_minus_distr at 3. rewrite <-Ropp_mult_distr_r, Rinv_l by lra.
+    cbn. lra.
+  - unfold segment. pose (t := ((L^P x - L^P b) / (L^P a - L^P b))%R). exists t.
+    assert (Ht : (0 <= t <= 1)%R).
+    { unfold t. split ; [apply Rdiv_le_0_compat | apply Rdiv_le_1] ; lra. }
+    split ; [assumption|].
+    assert (Hrewrite : t * a + (1 - t) * b == b + t * (a - b)).
+    { unfold Rminus. rewrite <-add_morph, mul_1, mul_distr_add, mul_opp, minus_morph.
+      rewrite 2 add_assoc. apply add_compat ; [|reflexivity]. now rewrite add_comm. }
+    rewrite Hrewrite. clear Hrewrite.
+    cut (L^P x == L^P (b + t * (a - b))).
+    { intros HeqP. apply (line_proj_inv_compat _ _ (reflexivity L)) in HeqP.
+      rewrite Hx, line_contains_combine in HeqP by assumption. assumption. }
+    rewrite line_proj_combine. unfold t, Rdiv. rewrite Rmult_assoc.
+    rewrite Rinv_l by lra. cbn. lra.
+Qed.
+
+Lemma strict_segment_line L a b x : 
+  line_contains L a ->
+  line_contains L b ->
+  line_contains L x -> 
+  (strict_segment a b x <-> (L^P a < L^P x < L^P b)%R \/ (L^P b < L^P x < L^P a)%R).
+Proof using Type. 
+intros Ha Hb Hx. unfold strict_segment. split.
++ intros [Hseg [NExa NExb]].
+  rewrite segment_line in Hseg by eassumption. 
+  assert (NExaP : L^P x <> L^P a).
+  { intros Heq. apply NExa. now rewrite <-Hx, <-Ha, Heq. }
+  assert (NExbP : L^P x <> L^P b).
+  { intros Heq. apply NExb. now rewrite <-Hx, <-Hb, Heq. }
+  lra.
++ intros Hlt. split ; [|split].
+  - rewrite segment_line by eassumption. lra.
+  - intros Exa. rewrite Exa in Hlt. lra.
+  - intros Exb. rewrite Exb in Hlt. lra.
+Qed.
+
+Lemma line_min_max_spec L ps r1 r2 : 
+  ps <> nil ->
+  aligned_on L ps ->
+  perm_pairA equiv (r1, r2) (L^-P (L^min ps), L^-P (L^max ps)) <-> 
+  (InA equiv r1 ps /\ InA equiv r2 ps /\ forall x, InA equiv x ps -> segment r1 r2 x).
+Proof using Type. 
+intros Hnil Halign. 
+assert (Lmin : line_contains L (L^-P (L^min ps))). { now apply Halign, line_iP_min_InA. }
+assert (Lmax : line_contains L (L^-P (L^max ps))). { now apply Halign, line_iP_max_InA. }
+split.
++ intros Hperm. 
+  case Hperm as [[H1 H2] | [H1 H2]] ; setoid_rewrite H1 ; setoid_rewrite H2.
+  - repeat split.
+    * now apply line_iP_min_InA.
+    * now apply line_iP_max_InA.
+    * intros x Hin. rewrite (segment_line L) ; try assumption. 
+      ++left. rewrite line_P_iP_min, line_P_iP_max. now apply line_bounds.
+      ++now apply Halign.
+  - repeat split.
+    * now apply line_iP_max_InA.
+    * now apply line_iP_min_InA.
+    * intros x Hin. rewrite (segment_line L) ; try assumption. 
+      ++right. rewrite line_P_iP_min, line_P_iP_max. now apply line_bounds.
+      ++now apply Halign.
++ intros [Hin1 [Hin2 Hseg]]. apply segment_eq_iff.
+  intros x. split ; apply segment_incl_iff.
+  - rewrite 2 (segment_line L) ; try assumption.
+    * rewrite line_P_iP_min, line_P_iP_max. 
+      split ; [left | left] ; now apply line_bounds.
+    * now apply Halign.
+    * now apply Halign.
+  - split ; apply Hseg.
+    * now apply line_iP_min_InA.
+    * now apply line_iP_max_InA.
+Qed.    
+
+Lemma line_endpoints_invariant L1 L2 ps :
+  ps <> nil -> 
+  aligned_on L1 ps -> 
+  aligned_on L2 ps ->
+  perm_pairA equiv (L1^-P (L1^min ps), L1^-P (L1^max ps)) (L2^-P (L2^min ps), L2^-P (L2^max ps)).
+Proof using Type. 
+intros Hnil Halign1 Halign2. rewrite line_min_max_spec ; [|assumption ..].
+repeat split.
++ now apply line_iP_min_InA. 
++ now apply line_iP_max_InA.
++ intros x Hin. rewrite (segment_line L1).
+  - rewrite line_P_iP_min, line_P_iP_max. left. now apply line_bounds.
+  - apply Halign1. now apply line_iP_min_InA.
+  - apply Halign1. now apply line_iP_max_InA.
+  - now apply Halign1.
+Qed.
+
+(* [L^left x ps] is the sublist of points of [ps] that are
+ * strictly smaller than [x] when projected on [L]. *)
+Definition line_left L x ps := 
+  filter (fun y => Rltb (L^P y) (L^P x)) ps.
+(* [L^right x ps] is the sublist of points of [ps] that are
+ * strictly greater than [x] when projected on [L]. *)
+Definition line_right L x ps := 
+  filter (fun y => Rltb (L^P x) (L^P y)) ps.
+(* [L^on x ps] is the sublist of points of [ps] that are
+ * equal to [x] when projected on [L]. *)
+Definition line_on L x ps := 
+  filter (fun y => L^P x ==b L^P y) ps.
+    
+Global Instance line_left_compat : 
+  Proper (equiv ==> equiv ==> eqlistA equiv ==> eqlistA equiv) line_left.
+Proof using Type. 
+intros ? ? H1 ? ? H2 ? ? H3. unfold line_left. f_equiv ; [|exact H3]. 
+intros ? ? H4. now rewrite H1, H2, H4. 
+Qed.
+
+Global Instance line_right_compat : 
+  Proper (equiv ==> equiv ==> eqlistA equiv ==> eqlistA equiv) line_right.
+Proof using Type. 
+intros ? ? H1 ? ? H2 ? ? H3. unfold line_right. f_equiv ; [|exact H3]. 
+intros ? ? H4. now rewrite H1, H2, H4.
+Qed. 
+
+Global Instance line_on_compat : 
+  Proper (equiv ==> equiv ==> eqlistA equiv ==> eqlistA equiv) line_on.
+Proof using Type. 
+intros ? ? H1 ? ? H2 ? ? H3. unfold line_on. f_equiv ; [|exact H3]. 
+intros ? ? H4. now rewrite H1, H2, H4.
+Qed. 
+
+Global Instance line_on_perm_compat : 
+  Proper (equiv ==> equiv ==> PermutationA equiv ==> PermutationA equiv) line_on.
+Proof using Type. 
+intros ? ? H1 ? ? H2 ? ? H3. unfold line_on.
+rewrite filter_extensionalityA_compat.
+- apply filter_PermutationA_compat.
+  + repeat intro.
+    rewrite H0.
+    reflexivity.
+  + eassumption.
+- repeat intro.
+  rewrite H2,H0,H1.
+  reflexivity.
+- reflexivity.
+Qed. 
+  
+Notation "L '^left'" := (line_left L).
+Notation "L '^right'" := (line_right L).
+Notation "L '^on'" := (line_on L).
+
+Lemma line_left_inclA L x ps : inclA equiv (L^left x ps) ps.
+Proof using Type. unfold line_left. apply filter_inclA. intros ? ? Heq. now rewrite Heq. Qed.
+Lemma line_right_inclA L x ps : inclA equiv (L^right x ps) ps.
+Proof using Type. unfold line_right. apply filter_inclA. intros ? ? Heq. now rewrite Heq. Qed.
+Lemma line_on_inclA L x ps : inclA equiv (L^on x ps) ps.
+Proof using Type. unfold line_on. apply filter_inclA. intros ? ? Heq. now rewrite Heq. Qed.
+
+Lemma line_left_inclA_2 L a ps b:
+  InA equiv b (L^left a ps) -> 
+  inclA equiv (L^left b ps) (L^left a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_left].
+  intros x H1. 
+  apply filter_InA in H1,H0.
+  destruct H1 as [H2 H3].
+  destruct H0 as [H1 H4 ].
+  eapply filter_InA.
+  - repeat intro.
+    now rewrite H0.
+  - split;auto.
+    apply Rltb_true in H3,H4.
+    apply Rltb_true.
+    lra.
+  - repeat intro.
+    now rewrite H3.
+  - repeat intro.
+    now rewrite H3.
+Qed.
+
+Lemma line_right_lt L a ps b:
+  InA equiv b (L^right a ps) -> 
+  inclA equiv (L^right b ps) (L^right a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_right].
+  intros x H1. 
+  apply filter_InA in H1,H0.
+  destruct H1 as [H2 H3].
+  destruct H0 as [H1 H4 ].
+  eapply filter_InA.
+  - repeat intro.
+    now rewrite H0.
+  - split;auto.
+    apply Rltb_true in H3,H4.
+    apply Rltb_true.
+    lra.
+  - repeat intro.
+    now rewrite H3.
+  - repeat intro.
+    now rewrite H3.
+Qed.
+
+Lemma line_on_inclA_2 L a ps b:
+  InA equiv b (L^on a ps) -> 
+  inclA equiv (L^on b ps) (L^on a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_on].
+  intros x H1. 
+  apply filter_InA in H1,H0.
+  destruct H1 as [H2 H3].
+  destruct H0 as [H1 H4 ].
+  eapply filter_InA.
+  - repeat intro.
+    now rewrite H0.
+  - split;auto.
+    rewrite eqb_true_iff.
+    rewrite eqb_true_iff in H3.
+    transitivity ((L ^P) b);auto.
+    now apply eqb_true_iff.
+  - repeat intro.
+    now rewrite H3.
+  - repeat intro.
+    now rewrite H3.
+Qed.
+
+Lemma line_left_inclA_3 L a ps b:
+  ((L^P) b <= (L^P) a)%R -> 
+  inclA equiv (L^left b ps) (L^left a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_left].
+  intros x H1. 
+  apply filter_InA in H1.
+  2:{ repeat intro.
+      now rewrite H3. }
+  destruct H1 as [H2 H3].
+  eapply filter_InA.
+  { repeat intro.
+    now rewrite H1. }
+  split;auto.
+  apply Rltb_true in H3.
+  apply Rltb_true.
+  lra.
+Qed.
+
+Lemma line_right_inclA_3 L a ps b:
+  ((L^P) a <= (L^P) b)%R -> 
+  inclA equiv (L^right b ps) (L^right a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_right].
+  intros x H1. 
+  apply filter_InA in H1.
+  2:{ repeat intro.
+      now rewrite H3. }
+  destruct H1 as [H2 H3].
+  eapply filter_InA.
+  { repeat intro.
+    now rewrite H1. }
+  split;auto.
+  apply Rltb_true in H3.
+  apply Rltb_true.
+  lra.
+Qed.
+
+Lemma line_on_inclA_3 L a ps b:
+  ((L^P) a == (L^P) b)%R -> 
+  inclA equiv (L^on b ps) (L^on a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_on].
+  intros x H1. 
+  apply filter_InA in H1.
+  2:{ repeat intro.
+      now rewrite H3. }
+  destruct H1 as [H2 H3].
+  eapply filter_InA.
+  { repeat intro.
+    now rewrite H1. }
+  split;auto.
+  rewrite eqb_true_iff.
+  rewrite eqb_true_iff in H3.
+  transitivity ((L ^P) b);auto.
+Qed.
+
+
+Lemma line_right_inclA_4 L a ps b:
+  ((L^P) a < (L^P) b)%R -> 
+  inclA equiv (L^on b ps) (L^right a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_right].
+  intros x H1. 
+  apply filter_InA in H1.
+  2:{ repeat intro.
+      now rewrite H3. }
+  destruct H1 as [H2 H3].
+  eapply filter_InA.
+  { repeat intro.
+    now rewrite H1. }
+  split;auto.
+  apply eqb_true_iff in H3.
+  apply Rltb_true.
+  cbn  in H3.
+  lra.
+Qed.
+
+Lemma line_left_inclA_4 L a ps b:
+  ((L^P) b < (L^P) a)%R -> 
+  inclA equiv (L^on b ps) (L^left a ps).
+Proof using Type.
+  intros H0.
+  lazy beta delta [inclA line_left].
+  intros x H1. 
+  apply filter_InA in H1.
+  2:{ repeat intro.
+      now rewrite H3. }
+  destruct H1 as [H2 H3].
+  eapply filter_InA.
+  { repeat intro.
+    now rewrite H1. }
+  split;auto.
+  apply eqb_true_iff in H3.
+  apply Rltb_true.
+  cbn  in H3.
+  lra.
+Qed.
+
+
+Lemma line_on_inclA_ps L a ps1 ps2 : inclA equiv ps1 ps2 -> inclA equiv ((L ^on) a ps1) ((L ^on) a ps2).
+Proof using Type.
+  intros h_incl.
+  unfold inclA , line_on.
+  intros x0 h_InA. 
+  rewrite filter_InA in h_InA;try typeclasses eauto.
+  rewrite filter_InA;try typeclasses eauto.
+  destruct h_InA as [h_InA h_eq].
+  split.
+  - now apply h_incl.
+  - assumption.  
+Qed.
+
+Lemma line_left_inclA_ps L a ps1 ps2 : inclA equiv ps1 ps2 -> inclA equiv ((L ^left) a ps1) ((L ^left) a ps2).
+Proof using Type.
+  intros h_incl.
+  unfold inclA , line_left.
+  intros x0 h_InA. 
+  rewrite filter_InA in h_InA.
+  2:{ repeat red.
+      intros x1 y heq.
+      now rewrite heq. }
+  rewrite filter_InA.
+  2:{ repeat red.
+      intros x1 y heq.
+      now rewrite heq. }
+  destruct h_InA as [h_InA h_eq].
+  split.
+  - now apply h_incl.
+  - assumption.  
+Qed.
+
+Lemma line_right_inclA_ps L a ps1 ps2 : inclA equiv ps1 ps2 -> inclA equiv ((L ^right) a ps1) ((L ^right) a ps2).
+Proof using Type.
+  intros h_incl.
+  unfold inclA , line_right.
+  intros x0 h_InA. 
+  rewrite filter_InA in h_InA.
+  2:{ repeat red.
+      intros x1 y heq.
+      now rewrite heq. }
+  rewrite filter_InA.
+  2:{ repeat red.
+      intros x1 y heq.
+      now rewrite heq. }
+  destruct h_InA as [h_InA h_eq].
+  split.
+  - now apply h_incl.
+  - assumption.  
+Qed.
+
+Lemma line_left_spec L x y ps : 
+  InA equiv y (L^left x ps) <-> (InA equiv y ps /\ (L^P y < L^P x)%R).
+Proof using Type. unfold line_left. now rewrite filter_InA, Rltb_true ; [|intros ? ? ->]. Qed.
+Lemma line_on_spec L x y ps : 
+  InA equiv y (L^on x ps) <-> (InA equiv y ps /\ L^P y == L^P x).
+Proof using Type. unfold line_on. now rewrite filter_InA, eqb_true_iff ; [|intros ? ? ->]. Qed.
+Lemma line_right_spec L x y ps : 
+  InA equiv y (L^right x ps) <-> (InA equiv y ps /\ (L^P y > L^P x)%R).
+Proof using Type. unfold line_right. now rewrite filter_InA, Rltb_true ; [|intros ? ? ->]. Qed.
+
+
+Lemma line_left_on_right_partition L x ps : 
+  PermutationA equiv ps (L^left x ps ++ L^on x ps ++ L^right x ps).
+Proof using Type.
+induction ps as [|y ps IH] ; cbn ; [reflexivity|].
+case (Rtotal_order (L^P x) (L^P y)) as [Hlt | [Heq | Hgt]].
++ case ifP_bool ; [rewrite Rltb_true ; lra | intros _].
+  case ifP_bool ; [rewrite eqb_true_iff ; cbn ; lra | intros _].
+  case ifP_bool ; [intros _ | rewrite Rltb_false ; lra]. 
+  rewrite IH at 1. rewrite 2 PermutationA_middle ; autoclass.
+  unfold line_left, line_on, line_right. reflexivity.
++ case ifP_bool ; [rewrite Rltb_true ; lra | intros _].
+  case ifP_bool ; [intros _ | rewrite eqb_false_iff ; cbn ; lra].
+  case ifP_bool ; [rewrite Rltb_true ; lra | intros _]. 
+  rewrite IH at 1. rewrite PermutationA_middle ; autoclass.
+  unfold line_left, line_on, line_right. reflexivity.
++ case ifP_bool ; [intros _ | rewrite Rltb_false ; lra].
+  case ifP_bool ; [rewrite eqb_true_iff ; cbn ; lra | intros _].
+  case ifP_bool ; [rewrite Rltb_true ; lra | intros _]. 
+  rewrite IH at 1. unfold line_left, line_on, line_right. reflexivity.
+Qed.
+
+Lemma line_left_iP_min L ps : L^left (L^-P (L^min ps)) ps = nil.
+Proof using Type. 
+unfold line_left. rewrite line_P_iP_min, (@filter_nilA _ equiv) ; [| autoclass | now intros ? ? ->].
+intros x Hin. rewrite Rltb_false. apply Rle_not_lt. now apply line_bounds.
+Qed.
+
+Lemma line_right_iP_max L ps : L^right (L^-P (L^max ps)) ps = nil.
+Proof using Type. 
+unfold line_right. rewrite line_P_iP_max, (@filter_nilA _ equiv) ; [| autoclass | now intros ? ? ->].
+intros x Hin. rewrite Rltb_false. apply Rle_not_lt. now apply line_bounds.
+Qed.
+
+Lemma line_reverse_left L w ps:
+  eqlistA equiv
+    ((line_reverse L)^left w ps)
+    ((L^right) w ps).
+Proof using Type.
+  unfold line_left , line_right.
+  apply filter_extensionalityA_compat.
+  2:reflexivity.
+  red.
+  intros x y heqxy. 
+  rewrite heqxy.
+  rewrite 2 line_reverse_proj.
+  destruct (Rltb ((L ^P) w) ((L ^P) y)) eqn:heq.
+  - rewrite Rltb_true in *|-*.
+    lra.
+  - rewrite Rltb_false in *|-*.
+    lra.
+Qed.
+
+Lemma line_reverse_right L w ps:
+  eqlistA equiv
+    ((line_reverse L)^right w ps)
+    ((L^left) w ps).
+Proof using Type.
+  unfold line_left , line_right.
+  apply filter_extensionalityA_compat.
+  2:reflexivity.
+  red.
+  intros x y heqxy.
+  rewrite heqxy.
+  rewrite 2 line_reverse_proj.
+  destruct (Rltb ((L ^P) y) ((L ^P) w)) eqn:heq.
+  - rewrite Rltb_true in *|-*.
+    lra.
+  - rewrite Rltb_false in *|-*.
+    lra.
+Qed.
+
+Lemma line_reverse_on L w ps:
+  eqlistA equiv
+    ((line_reverse L)^on w ps)
+    ((L^on) w ps).
+Proof using Type.
+  unfold line_on.
+  apply filter_extensionalityA_compat.
+  2:reflexivity.
+  red.
+  intros x y heqxy.
+  rewrite heqxy.
+  rewrite 2 line_reverse_proj.
+  destruct (((L ^P) w) ==b ((L ^P) y)) eqn:heq.
+  - rewrite eqb_true_iff in *|-*.
+    rewrite heq.
+    reflexivity.
+  - rewrite eqb_false_iff in *|-*.
+    cbn  .
+    cbn in heq.
+    lra.
+Qed.
+
+Lemma line_left_diff L a ps:
+    ((L ^left) a ps) <> nil ->
+    aligned_on L ps ->
+    (L ^-P) ((L ^max) ((L ^left) a ps)) =/= a.
+Proof using Type.
+  intros h_left h_align.
+  intro abs.
+  assert (aligned_on L ((L ^left) a ps)) as h_align'.
+  { eapply aligned_on_inclA with ps;auto.
+    apply line_left_inclA. }
+  specialize (line_iP_max_InA L _ h_left h_align') as h.
+  rewrite abs in h.
+  rewrite line_left_spec in h.
+  lra.
+Qed.
+
+
+Lemma eqlistA_nil_eq: forall [A : Type] (eqA : A -> A -> Prop) l, eqlistA eqA l nil <-> l = nil.
+Proof using Type.
+  intros A eqA l.
+  split.  
+  - induction l;intros;auto.
+    inversion H0.
+  - intro h.
+    subst.
+    apply eqlistA_nil.
+Qed.
+
+
+Lemma line_right_diff L a ps:
+    ((L ^right) a ps) <> nil ->
+    aligned_on L ps ->
+    (L ^-P) ((L ^min) ((L ^right) a ps)) =/= a.
+Proof using Type.
+  intros h_right h_align.
+  rewrite <- line_reverse_iP_max.
+
+  rewrite <- line_reverse_left.
+  apply line_left_diff;auto.
+  - intro abs.
+    rewrite <- eqlistA_nil_eq with (eqA:=equiv) in abs .
+    rewrite line_reverse_left in abs.
+    rewrite eqlistA_nil_eq in abs.
+    contradiction.
+  - now apply -> aligned_on_reverse.
+Qed.
+  
+Lemma bipartition_min: forall L ps,
+    PermutationA equiv ps
+      ((L ^on) (L^-P (L^min ps)) ps ++ (L ^right) (L^-P (L^min ps)) ps).
+Proof using Type.
+  intros L ps. 
+  rewrite line_left_on_right_partition with (L:=L) (x:=(L^-P (L ^min ps))) at 1.
+  rewrite line_left_iP_min.
+  reflexivity.
+Qed.
+
+Lemma bipartition_max: forall L ps,
+    PermutationA equiv ps
+      ((L ^left) (L^-P (L^max ps)) ps ++(L ^on) (L^-P (L^max ps)) ps).
+Proof using Type.
+  intros L ps. 
+  rewrite line_left_on_right_partition with (L:=L) (x:=(L^-P (L ^max ps))) at 1.
+  rewrite line_right_iP_max.
+  rewrite app_nil_r.
+  reflexivity.
+Qed.
+
+Lemma aggravate_right: forall L ps w w',
+    ((L^P w) < (L^P w'))%R
+    -> L^right w' ps = L^right w' (L^right w ps).
+Proof using Type.
+  symmetry.
+  unfold line_right.
+  rewrite filter_comm.
+  rewrite filter_weakened;auto.
+  intros x H1 H2. 
+  rewrite  Rltb_true in *|- *.
+  etransitivity;eauto.
+Qed.
+
+Lemma aggravate_left: forall L ps w w',
+    ((L^P w') < (L^P w))%R
+    -> L^left w' ps = L^left w' (L^left w ps).
+Proof using Type.
+  symmetry.
+  unfold line_left.
+  rewrite filter_comm.
+  rewrite filter_weakened;auto.
+  intros x H1 H2. 
+  rewrite  Rltb_true in *|- *.
+  etransitivity;eauto.
+Qed.
+
+Lemma aggravate_on_right: forall L ps w w',
+    ((L^P w) < (L^P w'))%R
+    -> L^on w' ps = L^on w' (L^right w ps).
+Proof using Type.
+  intros L ps w w' H0. 
+  unfold line_on, line_right.
+  rewrite <- filter_andb.
+  apply filter_extensionality_compat;auto.
+  repeat intro.
+  subst y.
+  destruct ((L ^P) w' ==b (L ^P) x) eqn:heq;cbn;auto.
+  symmetry.
+  rewrite eqb_true_iff in heq.
+  rewrite Rltb_true.
+  now rewrite <- heq.
+Qed.
+
+Lemma aggravate_on_left: forall L ps w w',
+    ((L^P w') < (L^P w))%R
+    -> L^on w' ps = L^on w' (L^left w ps).
+Proof using Type.
+  intros L ps w' w H0. 
+  unfold line_on, line_left.
+  rewrite <- filter_andb.
+  apply filter_extensionality_compat;auto.
+  repeat intro.
+  subst y.
+  destruct ((L ^P) w ==b (L ^P) x) eqn:heq;cbn;auto;[].
+  symmetry.
+  rewrite eqb_true_iff in heq.
+  rewrite Rltb_true.
+  now rewrite <- heq.
+Qed.
+
+Theorem PermutationA_app_inv_l : forall {A : Type} [eqA : relation A],
+    Equivalence eqA ->
+    forall(l l1 l2:list A),
+    PermutationA eqA (l ++ l1) (l ++ l2) ->
+    PermutationA eqA l1 l2.
+Proof using Type.
+  induction l; simpl; auto.
+  intros.
+  apply IHl.
+  apply PermutationA_cons_inv with a; auto.
+Qed.
+
+Theorem PermutationA_app_inv_r : forall {A : Type} [eqA : relation A],
+    Equivalence eqA ->
+    forall(l l1 l2:list A),
+      PermutationA eqA (l1 ++ l) (l2 ++ l) ->
+      PermutationA eqA l1 l2.
+Proof using Type.
+  intros A eqA H0 l l1 l2 h.
+  rewrite 2 (PermutationA_app_comm _ _ l) in h.
+  eapply PermutationA_app_inv_l ;try typeclasses eauto.
+  eassumption.
+Qed.
+
+Lemma aggravate_right': forall L ps w w',
+    ((L^P w') < (L^P w))%R
+    -> L^right w' (L^left w ps) = nil
+    -> L^on w ps = nil
+    -> PermutationA equiv (L^right w' ps) (L^right w ps).
+Proof using Type.
+  intros L ps w w' H0 H1 H2. 
+  specialize (line_left_on_right_partition L w' ps) as h1.
+  specialize (line_left_on_right_partition L w ps) as h2.
+  rewrite H2 in h2.
+  rewrite app_nil_l in h2.
+  assert (PermutationA equiv ((L ^left) w' ps ++ (L ^on) w' ps) ((L ^left) w ps)).
+  { specialize (line_left_on_right_partition L w' ((L ^left) w ps)) as h3.
+    rewrite H1 in h3.
+    rewrite app_nil_r in h3.
+    rewrite <- aggravate_left in h3;auto.
+    rewrite <- aggravate_on_left in h3;auto.
+    rewrite h3.
+    reflexivity. }
+  rewrite <- H3 in h2.
+  rewrite app_assoc in h1.
+  rewrite h1 in h2 at 1.
+  apply PermutationA_app_inv_l with (2:=h2).
+  typeclasses eauto.
+Qed.
+
+
+Lemma aggravate_left' : forall L ps w w',
+    ((L ^P) w < (L ^P) w')%R ->
+    (L ^left) w' ((L ^right) w ps) = nil ->
+    (L ^on) w ps = nil ->
+    PermutationA equiv ((L ^left) w' ps) ((L ^left) w ps).
+Proof using Type.
+  intros L ps w w' H0 H1 H2. 
+  remember (line_reverse L) as L'.
+  
+  repeat rewrite <- line_reverse_right.
+  apply aggravate_right'.
+  - repeat rewrite line_reverse_proj.
+    lra.
+  - apply PermutationA_nil with equiv;auto.
+    + typeclasses eauto.
+    + setoid_rewrite line_reverse_left.
+      setoid_rewrite line_reverse_right.
+      now rewrite H1.
+  - apply eqlistA_nil_eq with equiv.
+    rewrite <- H2.
+    apply line_reverse_on.
+Qed.
+
+Lemma line_on_length_aligned L x ps : 
+  aligned_on L (x :: ps) ->
+  length (L^on x ps) = countA_occ equiv EQ0 x ps.
+Proof using Type. 
+intros Halign. induction ps as [|p0 ps IH].
++ cbn. reflexivity.
++ cbn. case ifP_bool.
+  - rewrite eqb_true_iff. intros Heq. cbn.
+    apply (line_proj_inv_compat _ _ (reflexivity L)) in Heq.
+    rewrite 2 Halign in Heq ; [| intuition | intuition].
+    case ifP_sumbool ; [intros _ | intuition].
+    rewrite <-IH ; [reflexivity|].
+    intros ? Hin. apply Halign. rewrite InA_cons in Hin. case Hin as [-> | Hin] ; intuition.
+  - rewrite eqb_false_iff. intros HNeq.
+    case ifP_sumbool.
+    * intros Heq. apply (line_proj_compat _ _ (reflexivity L)) in Heq. intuition.
+    * intros _. rewrite <-IH ; [reflexivity|].
+      intros ? Hin. apply Halign. rewrite InA_cons in Hin. case Hin as [-> |] ; intuition.
+Qed.
+
+End Lines.
+
+Notation "L '^P'"   := (line_proj L).
+Notation "L '^-P'"  := (line_proj_inv L).
+Notation "L '^min'" := (line_min L).
+Notation "L '^max'" := (line_max L).
+Notation "L '^left'" := (line_left L).
+Notation "L '^right'" := (line_right L).
+Notation "L '^on'" := (line_on L).
+
+(* Some results about similarities are only available in R2. *)
+Section AlignmentR2.
+Local Existing Instances R2_VS R2_ES.
+Implicit Types (ps : list R2) (sim : similarity R2).
+
+Lemma R2aligned_similarity_weak ps sim :
+  aligned ps -> aligned (map sim ps).
+Proof.
+destruct ps as [|p0 ps].
++ intros _. cbn. apply aligned_nil. 
++ cbn. do 2 (rewrite aligned_spec ; setoid_rewrite InA_Leibniz). 
+  intros [v Halign]. eexists ?[simv]. intros y Hin.
+  rewrite in_map_iff in Hin. destruct Hin as [x [<- Hin]]. destruct (Halign x Hin) as [t Hx]. rewrite Hx.
+  eexists ?[simt]. rewrite sim_add, sim_mul, <-2 add_assoc. f_equiv.
+  instantiate (simt := t). instantiate (simv := (sim v - sim 0)%VS).
+  unfold Rminus. rewrite <-add_morph, mul_distr_add, mul_1. 
+  rewrite <-RealVectorSpace.add_assoc, add_sub. f_equiv.
+  now rewrite mul_opp, minus_morph.
+Qed.
+
+Lemma R2aligned_similarity ps sim :
+  aligned ps <-> aligned (map sim ps).
+Proof.
+split ; [apply R2aligned_similarity_weak|].
+intros H. apply R2aligned_similarity_weak with (sim := inverse sim) in H. revert H.
+apply aligned_compat, eqlistA_PermutationA. rewrite <-List.map_id at 1. rewrite map_map.
+(* f_equiv triggers some dumb inference making FMapInterface.eq_key appear from nowhere... *)
+(* we force the relation here to avoid that. *)
+specialize @map_extensionalityA_compat with (eqA:=@equiv R2 R2_Setoid)(eqB:=@equiv R2 R2_Setoid) as h.
+apply h;try typeclasses eauto;try reflexivity.
+intros x y H. 
+cbn -[equiv].
+rewrite (Bijection.retraction_section sim).
+assumption.
+Qed.
+
+End AlignmentR2.
+
+
+
+
+
+
diff --git a/CaseStudies/Gathering/InR2/Weber/Segment.v b/CaseStudies/Gathering/InR2/Weber/Segment.v
new file mode 100644
index 0000000000000000000000000000000000000000..093edcfffde7ed59968d3c316b654cc46b475994
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Segment.v
@@ -0,0 +1,349 @@
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file gives some properties of segments in a euclidean space.
+ * A vector space is not enough to have the decidability of being in a segment. *)
+(**************************************************************************)
+
+
+Require Import Bool.
+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 FunInd.
+Require Import FMapFacts.
+Require Import SetoidDec.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+Require Import Pactole.Spaces.EuclideanSpace.
+Require Import Pactole.Util.SetoidDefs.
+Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
+Require Import Pactole.Util.Bijection.
+Require Import Pactole.Spaces.Similarity.
+Require Import Pactole.CaseStudies.Gathering.InR2.Weber.Utils.
+
+
+Section Segment.
+Context {T : Type} `{EuclideanSpace T}.
+Implicit Types (a b c x y z : T).
+
+(* [segment a b] represents the set of points that are in the segment 
+ * [a, b], endpoints INCLUDED. *)
+Definition segment a b : T -> Prop := fun x =>
+  exists t : R, (0 <= t <= 1)%R /\ (x == t * a + (1 - t) * b)%VS.
+
+Global Instance segment_compat : Proper (equiv ==> equiv ==> equiv ==> iff) segment.
+Proof using .
+intros ? ? H1 ? ? H2 ? ? H3. unfold segment.
+setoid_rewrite H1. setoid_rewrite H2. setoid_rewrite H3. reflexivity.
+Qed.
+
+Lemma segment_sym a b x : segment a b x <-> segment b a x.
+Proof using Type.
+revert a b x. 
+cut (forall a b x, segment a b x -> segment b a x).
+{ intros H1. split ; apply H1. }
+intros a b x. unfold segment. intros [t [Ht Hx]].
+exists (1 - t)%R. split ; [lra|].
+assert (H1 : ((1 - (1 - t)) == t)%R).
+{ unfold Rminus. now rewrite Ropp_plus_distr, Ropp_involutive, <-Rplus_assoc, Rplus_opp_r, Rplus_0_l. } 
+  rewrite H1, RealVectorSpace.add_comm. exact Hx.
+Qed. 
+
+Lemma segment_start a b : segment a b a.
+Proof using Type. 
+unfold segment. exists 1%R. split ; [lra|]. 
+now rewrite mul_1, Rminus_eq_0, mul_0, add_origin_r.
+Qed.
+
+Lemma segment_end a b : segment a b b.
+Proof using Type. 
+unfold segment. exists 0%R. split ; [lra|]. 
+now rewrite mul_0, Rminus_0_r, mul_1, add_origin_l.
+Qed.
+
+Lemma segment_same a x : segment a a x <-> x == a.
+Proof using Type.
+split ; [|intros -> ; apply segment_start].
+intros [t [Ht Hx]]. revert Hx. unfold Rminus. 
+now rewrite <-add_morph, mul_1, minus_morph, add_sub.
+Qed.
+
+Lemma segment_straight_path a b (r : ratio) : segment a b (straight_path a b r).
+Proof using Type.
+rewrite segment_sym. unfold segment. exists r. split.
++ apply ratio_bounds.
++ cbn -[equiv RealVectorSpace.add mul opp].
+  rewrite mul_distr_add, 2 (RealVectorSpace.add_comm (r * b)), RealVectorSpace.add_assoc.
+  apply add_compat ; auto. unfold Rminus. rewrite <-add_morph, mul_1. 
+  apply add_compat ; auto. now rewrite minus_morph, mul_opp.
+Qed.
+
+(* This lemma is the reason why we need a euclidean space (a vector space is not enough). *)
+Lemma segment_dec a b x : {segment a b x} + {~segment a b x}.
+Proof using H. 
+case (a =?= b) as [Hab | Hab].
++ case (x =?= a) as [Hxa | Hxa]. 
+  - left. rewrite Hxa, Hab. apply segment_start.
+  - right. rewrite Hab in *. rewrite segment_same. intuition.
++ assert (H1 : forall t, x == t * b + (1 - t) * a <-> x - a == t * (b - a)).
+  { 
+    intros t. unfold Rminus. rewrite <-add_morph, mul_1, minus_morph.
+    rewrite mul_distr_add, mul_opp, (add_comm a), add_assoc.
+    split ; [intros -> | intros H1].
+    + now rewrite <-add_assoc, add_opp, add_origin_r. 
+    + apply (add_reg_r (-a)). now rewrite <-add_assoc, add_opp, add_origin_r.
+  }
+  pose (t := (inner_product (x - a) (b - a) / (inner_product (b - a) (b - a)))%R).
+  case (x =?= t * b + (1 - t) * a) as [Hx | Hx].
+  - case (Rle_dec 0%R t) as [Ht0 | Ht0].
+    * case (Rle_dec t 1%R) as [Ht1 | Ht1].
+      ++left. rewrite segment_sym. exists t. split ; intuition.
+      ++right. rewrite segment_sym. intros [s [Hs Hx']].
+        rewrite H1 in Hx, Hx'. 
+        assert (Hst : t == s).
+        { rewrite Hx' in Hx. apply mul_reg_r in Hx ; auto.
+          unfold complement. rewrite (sub_origin b a). intuition. }
+        rewrite <-Hst in Hs. intuition.
+    * right. rewrite segment_sym. intros [s [Hs Hx']].
+      rewrite H1 in Hx, Hx'. 
+      assert (Hst : t == s).
+      { rewrite Hx' in Hx. apply mul_reg_r in Hx ; auto.
+        unfold complement. rewrite (sub_origin b a). intuition. }
+      rewrite <-Hst in Hs. intuition.
+  - right. rewrite segment_sym. intros [s [Hs Hx']].
+    assert (Hst : s == t).
+    {  
+      rewrite H1 in Hx'.
+      apply inner_product_compat in Hx'. specialize (Hx' (b - a) (b - a) ltac:(auto)).
+      rewrite inner_product_mul_l in Hx'.
+      apply (Rmult_eq_compat_r (Rinv (inner_product (b - a) (b - a)))) in Hx'.
+      rewrite Rmult_assoc, Rinv_r, Rmult_1_r in Hx'.
+      + rewrite <-Hx'. reflexivity.
+      + rewrite <-squared_norm_product. 
+        rewrite <-Rsqr_0, <-square_norm_equiv ; try lra. 
+        rewrite norm_defined, sub_origin. intuition.
+    }
+    rewrite Hst in Hx'. intuition.
+Qed.
+
+Definition segment_decb a b x := if segment_dec a b x then true else false.
+
+Lemma segment_decb_true a b x : 
+  segment_decb a b x = true <-> segment a b x.
+Proof using Type. unfold segment_decb. case ifP_sumbool ; intuition. Qed. 
+
+Lemma segment_decb_false a b x : 
+  segment_decb a b x = false <-> ~segment a b x.
+Proof using Type. unfold segment_decb. case ifP_sumbool ; intuition. Qed.
+
+Lemma segment_incl_iff a b a' b' : 
+  (forall x, segment a b x -> segment a' b' x) <-> segment a' b' a /\ segment a' b' b.
+Proof using Type.
+split.
++ intros Hincl. split ; apply Hincl ; [apply segment_start | apply segment_end].
++ intros [[ta [Hta Ha]] [tb [Htb Hb]]] x [t [Ht Hx]].
+  exists (t * ta + (1 - t) * tb)%R. split ; [nra|].
+  rewrite Hx, Ha, Hb. unfold Rminus.
+  repeat rewrite mul_distr_add. repeat rewrite mul_morph.
+  assert (Hrewrite : forall x1 x2 x3 x4 : T, x1 + x2 + (x3 + x4) == (x1 + x3) + (x2 + x4)).
+  { intros x1 x2 x3 x4. rewrite <-add_assoc, (add_assoc x2).
+    rewrite (add_comm x2). now repeat rewrite add_assoc. }
+  rewrite Hrewrite. clear Hrewrite. apply add_compat. 
+  - rewrite add_morph. reflexivity.
+  - rewrite add_morph. apply mul_compat ; auto ; []. nra.
+Qed.
+
+Lemma segment_eq_weak a b a' b' : 
+  (forall x, segment a b x <-> segment a' b' x) -> a == a' \/ a == b'.
+Proof using Type. 
+intros Heq. 
+assert (Hincl : forall x, segment a b x -> segment a' b' x). { apply Heq. }
+assert (Hincl' : forall x, segment a' b' x -> segment a b x). { apply Heq. }
+apply segment_incl_iff in Hincl. apply segment_incl_iff in Hincl'.
+destruct Hincl as [[ta [Hta Ha]] [tb [Htb Hb]]].
+destruct Hincl' as [[ta' [Hta' Ha']] [tb' [Htb' Hb']]].
+assert (Ha2 := Ha). rewrite Ha', Hb' in Ha. repeat rewrite mul_distr_add in Ha. repeat rewrite mul_morph in Ha.
+assert (Hrewrite : forall x1 x2 x3 x4, x1 + x2 + (x3 + x4) == (x1 + x3) + (x2 + x4)).
+{ intros x1 x2 x3 x4. rewrite <-add_assoc, (add_assoc x2).
+  rewrite (add_comm x2). now repeat rewrite add_assoc. }
+rewrite Hrewrite in Ha. clear Hrewrite. rewrite 2 add_morph in Ha.
+pose (t := (ta * ta' + (1 - ta) * tb')%R).
+assert (Ht1 : (ta * (1 - ta') + (1 - ta) * (1 - tb') = 1 - t)%R).
+{ unfold t. nra. }
+fold t in Ha. rewrite Ht1 in Ha. apply (add_compat _ _ (reflexivity (-a))) in Ha.
+rewrite (add_comm (-a)), add_opp, add_assoc, <-(mul_1 (-a)), mul_opp, <-minus_morph, add_morph in Ha.
+unfold Rminus in Ha. rewrite <-(Ropp_involutive t) in Ha at 1.
+rewrite <-Ropp_plus_distr, minus_morph, <-mul_opp, <-mul_distr_add, add_comm in Ha.
+symmetry in Ha. rewrite mul_eq0_iff, sub_origin in Ha.  
+case Ha as [Ht | Eab].
++ assert (Ht' : t = 1%R). { cbn in Ht. lra. } clear Ht.
+  unfold t in Ht'.
+  case (ta =?= 0%R) as [Hta0 | Hta0].
+  - right. now rewrite Ha2, Hta0, mul_0, add_origin_l, Rminus_0_r, mul_1.
+  - case (ta =?= 1%R) as [Hta1 | Hta1].
+    * left. now rewrite Ha2, Hta1, mul_1, Rminus_eq_0, mul_0, add_origin_r.
+    * cbn in Hta0, Hta1. assert (Hta0' : ta' = 1%R). { nra. }
+      left. now rewrite Ha', Hta0', mul_1, Rminus_eq_0, mul_0, add_origin_r.
++ left. setoid_rewrite Eab in Heq. setoid_rewrite segment_same in Heq.
+  symmetry. rewrite Heq. apply segment_start.
+Qed.
+
+Lemma segment_eq_iff a b a' b' : 
+  (forall x, segment a b x <-> segment a' b' x) <-> perm_pairA equiv (a, b) (a', b').
+Proof using Type. 
+split.
++ intros Heq. case (a =?= b) as [Eab | NEab].
+  - setoid_rewrite Eab in Heq. setoid_rewrite segment_same in Heq.
+    left. split.
+    * rewrite Eab. symmetry. rewrite Heq. apply segment_start.
+    * symmetry. rewrite Heq. apply segment_end.
+  - destruct (segment_eq_weak _ _ _ _ Heq) as [Haa' | Hab'] ; [left | right].
+    * split ; [now auto|].
+      setoid_rewrite segment_sym in Heq. 
+      destruct (segment_eq_weak _ _ _ _ Heq) as [Hbb' | Hba'] ; [auto | exfalso].
+      apply NEab. now rewrite Haa', Hba'.
+    * split ; [now auto|].
+      setoid_rewrite segment_sym in Heq. 
+      destruct (segment_eq_weak _ _ _ _ Heq) as [Hbb' | Hba'] ; [exfalso | auto].
+      apply NEab. now rewrite Hab', Hbb'.
++ intros [[Ha Hb] | [Ha Hb]] x ; rewrite Ha, Hb ; [|rewrite segment_sym] ; reflexivity.
+Qed.      
+  
+Lemma segment_decb_eq_iff a b x a' b' x' : 
+  segment_decb a b x = segment_decb a' b' x' <-> 
+  (segment a b x <-> segment a' b' x').
+Proof using Type. unfold segment_decb. case ifP_sumbool ; case ifP_sumbool ; intuition. Qed. 
+
+Lemma segment_decb_sym a b x : 
+  segment_decb a b x = segment_decb b a x.
+Proof using Type. unfold segment_decb. case ifP_sumbool ; case ifP_sumbool ; rewrite segment_sym ; intuition. Qed.
+
+End Segment.
+
+
+Section StrictSegment.
+Context {T : Type} `{EuclideanSpace T}.
+Implicit Types (a b c x y z : T).
+
+(* [strict_segment a b] represents the set of points that are in the segment 
+ * [a, b], endpoints EXCLUDED. *)
+Definition strict_segment a b : T -> Prop := fun x =>
+  segment a b x /\ x =/= a /\ x =/= b.
+
+Global Instance strict_segment_compat : Proper (equiv ==> equiv ==> equiv ==> iff) strict_segment.
+Proof using .
+intros ? ? H1 ? ? H2 ? ? H3. unfold strict_segment. now rewrite H1, H2, H3. 
+Qed.
+
+Lemma strict_segment_sym a b x : strict_segment a b x <-> strict_segment b a x.
+Proof using Type. unfold strict_segment. rewrite segment_sym. intuition. Qed.
+
+Lemma strict_segment_incl a b x :
+  strict_segment a b x -> segment a b x.
+Proof using Type. unfold strict_segment. intuition. Qed.
+
+Lemma segment_spec_strict a b x : 
+  segment a b x <-> strict_segment a b x \/ x == a \/ x == b.  
+Proof using Type.
+split.
++ case (x =?= a) ; [tauto|]. case (x =?= b) ; [tauto|].
+  unfold strict_segment. left. intuition.
++ intros [? |[-> | ->]].
+  - now apply strict_segment_incl.
+  - now apply segment_start.
+  - now apply segment_end.
+Qed.
+
+Lemma strict_segment_dec a b x : 
+    {strict_segment a b x} + {~strict_segment a b x}.
+Proof using H. 
+unfold strict_segment. 
+case (x =?= a) as [Hxa | Hxa] ; [intuition|]. 
+case (x =?= b) as [Hxb | Hxb] ; [intuition|].
+case (segment_dec a b x) as [Hseg | Hseg] ; intuition.
+Qed.
+
+Definition strict_segment_decb a b x := if strict_segment_dec a b x then true else false.
+
+Lemma strict_segment_decb_eq_iff a b x a' b' x' : 
+  strict_segment_decb a b x = strict_segment_decb a' b' x' <-> 
+  (strict_segment a b x <-> strict_segment a' b' x').
+Proof using Type. unfold strict_segment_decb. case ifP_sumbool ; case ifP_sumbool ; intuition. Qed.
+
+Lemma strict_segment_decb_sym a b x : 
+  strict_segment_decb a b x = strict_segment_decb b a x.
+Proof using Type. unfold strict_segment_decb. case ifP_sumbool ; case ifP_sumbool ; rewrite strict_segment_sym ; intuition. Qed.
+
+Lemma strict_segment_decb_true a b x : 
+  strict_segment_decb a b x = true <-> strict_segment a b x.
+Proof using Type. unfold strict_segment_decb. case ifP_sumbool ; intuition. Qed.
+
+Lemma strict_segment_decb_false a b x : 
+  strict_segment_decb a b x = false <-> ~strict_segment a b x. 
+Proof using Type. unfold strict_segment_decb. case ifP_sumbool ; intuition. Qed. 
+
+End StrictSegment.
+
+Require Import Pactole.Spaces.R2.
+
+(* Some results for similarities are only available in R². *)
+Section SegmentsInR2.
+Local Existing Instances R2_VS R2_ES.
+Implicit Types (sim : similarity R2) (a b x : R2).
+
+Lemma R2segment_similarity_weak sim a b x : 
+  segment a b x -> segment (sim a) (sim b) (sim x).
+Proof. 
+intros [t [Ht Hx]]. exists t. split ; auto ; []. rewrite Hx.
+rewrite sim_add, 2 sim_mul. repeat rewrite <-add_assoc. f_equiv.
+assert (Hrewrite : (1 - (1 - t))%R = t). { lra. } 
+rewrite Hrewrite. clear Hrewrite.
+rewrite add_comm, <-add_assoc. setoid_rewrite <-(add_origin_r ((1 - t) * sim b)%VS) at 2.
+f_equiv. unfold Rminus. rewrite <-add_morph, mul_1, minus_morph.
+rewrite <-add_assoc, (add_assoc (-sim 0)%VS), (add_comm (-sim 0)%VS).
+rewrite add_opp, add_origin_l, add_opp. reflexivity.
+Qed.
+
+Lemma R2segment_similarity sim a b x : 
+  segment a b x <-> segment (sim a) (sim b) (sim x).
+Proof.
+split.
++ now apply R2segment_similarity_weak.
++ intros Hseg. apply (R2segment_similarity_weak (inverse sim)) in Hseg. 
+  cbn in Hseg. repeat rewrite retraction_section in Hseg. exact Hseg. 
+Qed.
+
+Lemma R2strict_segment_similarity sim a b x :
+  strict_segment a b x <-> strict_segment (sim a) (sim b) (sim x).
+Proof.
+unfold strict_segment. rewrite (R2segment_similarity sim).
+assert (Hsim : forall x1 x2, x1 =/= x2 <-> sim x1 =/= sim x2).
+{ intros x1 x2. cbn. split.
+  + intros NEx Esim. apply (f_equal (retraction sim)) in Esim.
+    rewrite 2 retraction_section in Esim. intuition.
+  + intros NEsim Ex. apply NEsim. now f_equal. }
+rewrite (Hsim x a), (Hsim x b). reflexivity.
+Qed. 
+
+End SegmentsInR2.
diff --git a/CaseStudies/Gathering/InR2/Weber/Utils.v b/CaseStudies/Gathering/InR2/Weber/Utils.v
new file mode 100644
index 0000000000000000000000000000000000000000..994c7e816332c2275c1f6090a7bd5c1a23d32168
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Utils.v
@@ -0,0 +1,818 @@
+
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022). *)
+(**************************************************************************)
+
+
+Require Import Bool.
+Require Import Lia Field.
+Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def Ranalysis.
+Require Import List.
+Require Import SetoidList.
+Require Import Relations.
+Require Import RelationPairs.
+Require Import Morphisms.
+Require Import Psatz.
+Require Import Inverse_Image.
+Require Import FunInd.
+Require Import FMapFacts.
+Require Import SetoidDec.
+Require Import Permutation.
+Require Import Program.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
+Require Import Pactole.Util.Stream.
+Require Import Pactole.Spaces.R2.
+Require Import Pactole.Observations.MultisetObservation.
+Require Import Pactole.Core.Identifiers.
+Require Import Pactole.Spaces.Similarity.
+
+Set Implicit Arguments.
+
+Close Scope R_scope.
+
+Local Existing Instance R2_VS.
+Local Existing Instance R2_ES.
+
+
+(* Change the left hand side of a setoid-equality with a convertible term. *)
+Ltac change_LHS E := 
+  match goal with 
+  | [ |- ?LHS == ?RHS ] => change (E == RHS)
+  end.
+
+(* Change the right hand side of a setoid-equality with a convertible term. *)
+Ltac change_RHS E := 
+  match goal with 
+  | [ |- ?LHS == ?RHS ] => change (LHS == E)
+  end.
+
+(* This tactic feeds the precondition of an implication in order to derive the conclusion
+  (taken from http://comments.gmane.org/gmane.science.mathematics.logic.coq.club/7013).
+  Usage: feed H.
+  H: P -> Q  ==becomes==>  H: P
+                          ____
+                          Q
+  After completing this proof, Q becomes a hypothesis in the context. *)
+Ltac feed H :=
+  match type of H with
+  | ?foo -> _ =>
+    let FOO := fresh in
+    assert foo as FOO; [|specialize (H FOO); clear FOO]
+  end.
+
+(* Generalization of feed for multiple hypotheses.
+  feed_n is useful for accessing conclusions of long implications.
+  Usage: feed_n 3 H.
+    H: P1 -> P2 -> P3 -> Q.
+  We'll be asked to prove P1, P2 and P3, so that Q can be inferred. *)
+Ltac feed_n n H := match constr:(n) with
+  | O => idtac
+  | (S ?m) => feed H ; [| feed_n m H]
+  end.
+
+(* Simplify a goal involving calculations in R2 by expanding everything. 
+ * This is rarely useful. *)
+Ltac simplifyR2 :=
+unfold Rminus ; 
+repeat (try rewrite mul_distr_add ;
+        try rewrite <-add_morph ;
+        try rewrite mul_0 ;
+        try rewrite mul_1 ; 
+        try rewrite add_origin_l ; 
+        try rewrite add_origin_r ; 
+        try rewrite mul_opp ; 
+        try rewrite minus_morph ;
+        try rewrite opp_opp ; 
+        try rewrite opp_origin ;
+        try rewrite R2_opp_dist ; 
+        try rewrite add_opp).
+
+Lemma contra (P Q : Prop) : (Q -> P) -> (~P -> ~Q).
+Proof. intuition. Qed.
+
+Lemma contra_bool b1 b2 : (b2 = true -> b1 = true) -> (negb b1 = true -> negb b2 = true).
+Proof using . case b1 ; case b2 ; intuition. Qed.
+
+(* This inductive type inspired by mathcomp allows elegant destruction 
+ * of if-statements (more so than the 'destruct_match' tactic).
+ * Usage : [case ifP_sumbool] will destuct the first if-statement 
+ * with a sumbool condition in the goal. *)
+Section IfSpecSumbool.
+Variables (P Q : Prop) (T : Type).
+Implicit Types (u v : T).
+
+Inductive if_spec_sumbool u v : {P} + {Q} -> T -> Prop := 
+  | if_spec_sumbool_true (p : P) : if_spec_sumbool u v (left p) u 
+  | if_spec_sumbool_false (q : Q) : if_spec_sumbool u v (right q) v.
+
+Lemma ifP_sumbool (c : {P} + {Q}) u v : if_spec_sumbool u v c (if c then u else v).
+Proof using . case c ; constructor. Qed.
+
+End IfSpecSumbool.
+
+(* Usage : [case ifP_bool] will destuct the first if-statement 
+ * with a boolean condition in the goal. *)
+Section IfSpecBool.
+Variables (T : Type) (u v : T).
+
+Inductive if_spec_bool (B NB : Prop) : bool -> T -> Prop := 
+  | if_spec_bool_true : B -> if_spec_bool B NB true u 
+  | if_spec_bool_false : NB -> if_spec_bool B NB false v.
+
+Lemma ifP_bool b : if_spec_bool (b = true) (b = false) b (if b then u else v).
+Proof using . case b ; constructor ; reflexivity. Qed.
+
+End IfSpecBool.
+
+
+(* Lemma ifP_bool_test b : b = negb b -> (if b then b || b else false) = false.
+Proof. case ifP_bool. ... *)
+
+Section ForallTriplets.
+Variable (A B C : Type).
+Implicit Types (R : A -> B -> C -> Prop).
+
+(* The proposition R holds for every triplet in the cartesian product (l1 * l2 * l3). *)
+Definition ForallTriplets R l1 l2 l3 : Prop :=
+  Forall (fun x => Forall (fun y => Forall (fun z => R x y z) l3) l2) l1.
+
+Local Instance Forall_PermutationA_compat_strong {T : Type} `{Setoid T} : 
+  Proper ((equiv ==> iff) ==> PermutationA equiv ==> iff) (@Forall T).
+Proof using . 
+intros P P' HP l l' Hl. elim Hl.
++ intuition.
++ intros x x' t t' Hx Ht IH. repeat rewrite Forall_cons_iff. now f_equiv ; [apply HP|].
++ intros x y t. repeat rewrite Forall_cons_iff. repeat rewrite <-and_assoc. f_equiv.
+  - rewrite and_comm. now f_equiv ; auto.
+  - f_equiv. intros ? ? ->. now apply HP.
++ intros t1 t2 t3 _ IH1 _ IH2. rewrite IH1, <-IH2.
+  f_equiv. intros ? ? ->. symmetry. now apply HP.
+Qed.
+
+Local Instance ForallTriplets_PermutationA_compat  `{Setoid A} `{Setoid B} `{Setoid C} : 
+  Proper ((equiv ==> equiv ==> equiv ==> iff) ==> 
+    PermutationA equiv ==> PermutationA equiv ==> PermutationA equiv ==> iff) ForallTriplets.
+Proof using . 
+intros R R' HR l1 l1' Hl1 l2 l2' Hl2 l3 l3' Hl3. unfold ForallTriplets.
+repeat (f_equiv ; auto ; intros ? ? ?). now apply HR.
+Qed.
+
+(* As with Forall, ForallTriplets is decidable (provided R is decidable). *)
+Lemma ForallTriplets_dec R l1 l2 l3 : 
+  (forall x y z, {R x y z} + {~ R x y z}) ->
+  {ForallTriplets R l1 l2 l3} + {~ ForallTriplets R l1 l2 l3}.
+Proof using . intros Rdec. unfold ForallTriplets. repeat (apply Forall_dec ; intros ?). now apply Rdec. Qed.
+
+(* The equivalence between ForallTriplets and regular forall. *)
+Lemma ForallTriplets_forall R l1 l2 l3 : 
+  (ForallTriplets R l1 l2 l3) <-> forall x y z, List.In x l1 -> List.In y l2 -> List.In z l3 -> R x y z.
+Proof using . 
+unfold ForallTriplets. split.
++ intros H x y z Hinx Hiny Hinz.
+  rewrite Forall_forall in H. specialize (H x Hinx).
+  rewrite Forall_forall in H. specialize (H y Hiny).
+  rewrite Forall_forall in H. specialize (H z Hinz).
+  exact H.
++ intros H. 
+  rewrite Forall_forall. intros x Hinx.
+  rewrite Forall_forall. intros y Hiny.
+  rewrite Forall_forall. intros z Hinz.
+  auto.
+Qed.
+
+End ForallTriplets.
+
+
+Global Instance countA_occ_compat_setoid {A : Type} `{eq_dec : EqDec A} : 
+  Proper (equiv ==> PermutationA equiv ==> eq) (countA_occ equiv eq_dec).
+Proof using . repeat intros ?. apply countA_occ_compat ; autoclass. Qed.
+
+Lemma countA_occ_le {A : Type} `{eq_dec : EqDec A} w ps ps' :
+  Forall2 (fun x x' => x' == x \/ x' == w) ps ps' -> 
+    countA_occ equiv eq_dec w ps <= countA_occ equiv eq_dec w ps'.
+Proof using . 
+intros HF. induction HF as [| x x' l l' Hxx' Hll' IH] ; [auto|].
+cbn -[equiv]. repeat destruct_match ; intuition.
+rewrite H, H1 in *. intuition.
+Qed.
+
+Lemma countA_occ_lt {A : Type} `{eq_dec : EqDec A} w ps ps' : 
+  Forall2 (fun x x' => x' == x \/ x' == w) ps ps' -> 
+  List.Exists (fun '(x, x') => x' =/= x) (combine ps ps') ->
+    countA_occ equiv eq_dec w ps < countA_occ equiv eq_dec w ps'.
+Proof using .
+intros HF HE. induction HF as [| x x' l l' Hxx' Hll' IH].
++ rewrite Exists_exists in HE. destruct HE as [x [In_nil _]]. now apply in_nil in In_nil.
++ cbn -[complement equiv] in HE. rewrite Exists_cons in HE.
+  destruct HE as [Dxx' | HE].
+  - destruct Hxx' as [Exx' | Ex'w] ; intuition.
+    rewrite Ex'w in Dxx' |- *. cbn -[equiv].
+    repeat destruct_match ; intuition. rewrite Nat.lt_succ_r. now apply countA_occ_le.
+  - destruct Hxx' as [Exx' | Ex'w] ; cbn -[equiv] ; repeat destruct_match ; intuition.
+    rewrite Exx', H in *. intuition.
+Qed.
+
+Lemma list_all_eq_or_perm {A : Type} `{Setoid A} `{EqDec A} (x0 : A) l : 
+  (forall x, InA equiv x l -> x == x0) \/ (exists x1 l1, PermutationA equiv l (x1 :: l1) /\ x1 =/= x0).
+Proof.
+case (Forall_dec (equiv x0) (equiv_dec x0) l) as [Hall_eq | HNall_eq].
++ left. intros x Hin. rewrite Forall_forall in Hall_eq.
+  rewrite InA_alt in Hin. destruct Hin as [y [Hxy Hin]].
+  rewrite Hxy. symmetry. now apply Hall_eq.
++ right. apply neg_Forall_Exists_neg in HNall_eq ; [|apply equiv_dec].
+  rewrite Exists_exists in HNall_eq. destruct HNall_eq as [x1 [Hin Hx1]].
+  apply (@In_InA _ equiv) in Hin ; autoclass.
+  apply PermutationA_split in Hin ; autoclass.
+  destruct Hin as [l1 Hperm].
+  exists x1. exists l1. split ; [exact Hperm | symmetry ; exact Hx1].
+Qed.
+
+Lemma add_sub {A : Type} `{R : RealVectorSpace A} (x y : A) :
+  (x + (y - x) == y)%VS.
+Proof. now rewrite (RealVectorSpace.add_comm y), RealVectorSpace.add_assoc, add_opp, add_origin_l. Qed. 
+
+Lemma colinear_exists_mul u v : 
+  ~ u == 0%VS -> colinear u v -> exists t, v == (t * u)%VS. 
+Proof.
+intros Hu_n0 Hcol. destruct (colinear_decompose Hu_n0 Hcol) as [Hdecomp | Hdecomp].
++ exists (norm v / norm u)%R. rewrite Hdecomp at 1. unfold Rdiv, unitary.
+  now rewrite mul_morph.
++ exists ((- norm v) / norm u)%R. rewrite Hdecomp at 1. unfold Rdiv, unitary.
+  now rewrite mul_morph.
+Qed.   
+
+Lemma sub_origin {T : Type} `{RealVectorSpace T} (x y : T) : 
+  (x - y == 0)%VS <-> x == y.
+Proof.
+split.
++ intros Hxy. apply (add_reg_r (-y)). now rewrite add_opp.
++ intros ->. now rewrite add_opp. 
+Qed.
+
+Lemma eqb_true_iff {A} `{EqDec A} (x y : A) :
+  (x ==b y) = true <-> x == y.
+Proof. unfold equiv_decb. case ifP_sumbool ; intuition. Qed.
+
+Lemma eqb_false_iff {A} `{EqDec A} (x y : A) :
+  (x ==b y) = false <-> ~(x == y).
+Proof. unfold equiv_decb. case ifP_sumbool ; intuition. Qed.
+
+Lemma eqb_sym {A} `{EqDec A} (x y : A) : (x ==b y) = (y ==b x).
+Proof. unfold equiv_decb. case ifP_sumbool ; case ifP_sumbool ; intuition. Qed.
+
+Lemma eqb_same {A} `{EqDec A} (x : A) : (x ==b x) = true.
+Proof. rewrite eqb_true_iff. reflexivity. Qed.
+  
+Definition perm_pairA {A : Type} (eqA : relation A) : relation (A * A) := 
+  fun '(x1, x2) '(y1, y2) => (eqA x1 y1 /\ eqA x2 y2) \/ (eqA x1 y2 /\ eqA x2 y1). 
+
+Global Instance perm_pairA_subrelation {A : Type} (eqA : relation A) :
+  subrelation (eqA * eqA) (perm_pairA eqA).
+Proof using . intros [x1 x2] [y1 y2] [H1 H2]. cbn in H1, H2. unfold perm_pairA. now left. Qed.
+
+Global Instance perm_pairA_equiv A eqA : 
+  Equivalence eqA -> Equivalence (@perm_pairA A eqA).
+Proof. 
+intros Hequiv. split.
++ intros [x1 x2]. left. intuition.
++ intros [x1 x2] [y1 y2] Hperm. case Hperm as [[H1 H2] | [H1 H2]].
+  - left. intuition.
+  - right. intuition.
++ intros [x1 x2] [y1 y2] [z1 z2] [[H1 H2] | [H1 H2]] [[H3 H4] | [H3 H4]] ; 
+    solve [left ; split ; etransitivity ; eauto | right ; split ; etransitivity ; eauto].
+Qed.
+
+Lemma option_map_map {A B C} (f : A -> B) (g : B -> C) x :
+  option_map g (option_map f x) = option_map (fun x => g (f x)) x.
+Proof. case x as [x|] ; cbn ; reflexivity. Qed.
+
+Lemma option_map_injective {A B} `{Setoid A} `{Setoid B} (f : A -> B) : 
+  Preliminary.injective equiv equiv f -> 
+  Preliminary.injective equiv equiv (option_map f).
+Proof.
+intros Hinj [?|] [?|] Heq ; cbn in Heq ; intuition.
+apply Hinj in Heq. intuition.
+Qed.
+
+Lemma option_map_id {A} `{Setoid A} (f : A -> A) x : 
+  (forall x, f x == x) -> option_map f x == x.
+Proof. intros Hid. case x as [x|] ; cbn ; auto. Qed.
+
+Lemma existsb_compat_eqlist {A} (eqA : relation A) : 
+  Proper ((eqA ==> eq) ==> eqlistA eqA ==> eq) (@existsb A).
+Proof.
+intros f1 f2 Hf l1 l2 Hl. elim Hl ; cbn ; clear l1 l2 Hl.
++ reflexivity.
++ intros x1 x2 l1 l2 Hx Hl IH. rewrite IH. rewrite (Hf _ _ Hx). reflexivity.
+Qed.
+
+Global Instance existsb_compat_perm {A} (eqA : relation A) :
+  Equivalence eqA ->
+  Proper ((eqA ==> eq) ==> PermutationA eqA ==> eq) (@existsb A).
+Proof. 
+intros Hequiv f1 f2 Hf l1 l2 Hl. elim Hl ; cbn ; clear l1 l2 Hl.
++ reflexivity.
++ intros x1 x2 l1 l2 Hx Hl IH. rewrite IH. rewrite (Hf _ _ Hx). reflexivity.
++ intros x y l. rewrite 2 orb_assoc, (orb_comm (f1 y)). f_equal ; [f_equal|].
+  - now apply Hf.
+  - now apply Hf.
+  - eapply existsb_compat_eqlist ; [exact Hf | reflexivity].
++ intros l1 l2 l3 _ H1 _ H2. rewrite H1, <-H2. symmetry.
+  eapply existsb_compat_eqlist ; [exact Hf | reflexivity].
+Qed.
+
+Lemma existsb_map {A B} (f : B -> bool) (g : A -> B) l : 
+  existsb f (List.map g l) = existsb (fun x => f (g x)) l.
+Proof.
+induction l as [|x l IH] ; cbn.
++ reflexivity. 
++ rewrite IH. reflexivity.
+Qed.
+
+Lemma In_InA_is_leibniz {A : Type} (eqA : relation A) x l : 
+  (forall x y, eqA x y <-> x = y) -> (InA eqA x l <-> List.In x l).
+Proof using . 
+intros H. induction l as [|y l IH].
++ cbn. split ; [intros Hnil ; inv Hnil | intuition].
++ cbn. split.
+  - intros HinA. rewrite InA_cons in HinA. destruct HinA as [Heq | HinA].
+    * left. symmetry. now rewrite <-H.
+    * right. now rewrite <-IH.
+  - intros [Heq | Hin].
+    * apply InA_cons_hd. now rewrite H.
+    * apply InA_cons_tl. now rewrite IH.
+Qed.
+
+Lemma split_filter_perm {A} (P : A -> bool) l : 
+  Permutation l (List.filter P l ++ List.filter (fun x => negb (P x)) l).
+Proof. 
+induction l as [|x l IH] ; cbn ; [reflexivity|].
+case (P x) ; cbn [negb].
++ now rewrite <-app_comm_cons, <-IH.
++ change (x :: List.filter (fun x0 => negb (P x0)) l) with ((x :: nil) ++ List.filter (fun x0 => negb (P x0)) l).
+  rewrite app_assoc. setoid_rewrite Permutation_app_comm at 2.
+  rewrite <-app_assoc, <-IH. cbn. reflexivity.
+Qed.
+
+Lemma gt_lt_iff x y : x > y <-> y < x.
+Proof. lia. Qed.
+
+Lemma ge_le_iff x y : x >= y <-> y <= x.
+Proof. lia. Qed.
+
+Lemma countA_occ_length_eq {A} (eqA : relation A) 
+  (eq_dec : forall x y : A, {eqA x y} + {~ eqA x y}) x l : 
+  Equivalence eqA ->
+    countA_occ eqA eq_dec x l = length l <-> 
+    (forall y, InA eqA y l -> eqA x y).
+Proof.
+intros Hequiv. induction l as [|y l IH] ; cbn.
++ split ; [intros _ | now intuition]. intros ? Hin. inv Hin.
++ case ifP_sumbool. 
+  - intros Exy. split. 
+    * intros Hlen z Hin. inversion Hlen as [H]. clear Hlen. rewrite IH in H.
+      rewrite InA_cons in Hin. case Hin as [Eyz | Hin].
+      ++transitivity y ; intuition.
+      ++apply H. auto.
+    * intros Heq. f_equal. rewrite IH. intros z Hin. apply Heq. right. exact Hin.
+  - intros NExy. split.
+    * intros Hcount. exfalso. generalize (countA_occ_length_le eqA eq_dec l x). lia.
+    * intros Hin. exfalso. apply NExy. symmetry. apply Hin. now left.
+Qed.
+
+Lemma countA_occ_length_filter {A} `{eq_dec : EqDec A} x l :
+  countA_occ equiv eq_dec x l = length (List.filter (equiv_decb x) l).
+Proof. 
+induction l as [|y l IH] ; cbn.
++ reflexivity. 
++ case ifP_sumbool ; [intros E | intros NE].
+  - rewrite <-eqb_true_iff, eqb_sym in E. now rewrite E, IH.
+  - rewrite <-eqb_false_iff, eqb_sym in NE. now rewrite NE, IH.
+Qed.
+
+#[export]
+Instance map_extensionalityA_compat_perm {A B} (eqA : relation A) (eqB : relation B) :
+  Equivalence eqA ->
+  Equivalence eqB ->
+  Proper ((eqA ==> eqB) ==> PermutationA eqA ==> PermutationA eqB) (@List.map A B).
+Proof. 
+intros HeqA HeqB f1 f2 Hf l1 l2 Hl. 
+induction Hl as [| x1 x2 l1 l2 Heq Hperm IH | x y l | l1 l2 l3 H1 IH1 H2 IH2] ; cbn.
++ apply permA_nil.
++ apply permA_skip ; [now apply Hf | apply IH].
++ rewrite permA_swap. apply eqlistA_PermutationA.
+  change (f1 x :: f1 y :: List.map f1 l) with (List.map f1 (x :: y :: l)).
+  change (f2 x :: f2 y :: List.map f2 l) with (List.map f2 (x :: y :: l)).
+  eapply map_extensionalityA_compat ; autoclass. reflexivity.
++ rewrite IH1, <-IH2. apply eqlistA_PermutationA. symmetry. 
+  eapply map_extensionalityA_compat ; autoclass. reflexivity.       
+Qed.
+
+Lemma straight_path_end_advance (x y:R2) r1 r2 : 
+  straight_path x y r1 == y -> straight_path x y (add_ratio r1 r2) == y.
+Proof. 
+intros Hend. rewrite straight_path_end in *.
+case Hend as [? | Hr] ; [tauto|].
+right. generalize (add_ratio_ge_left r1 r2), (ratio_bounds (add_ratio r1 r2)).
+rewrite Hr. cbn. lra.
+Qed.
+
+Lemma straight_path_eq_0 r (x y:R2) : 
+  proj_ratio r == 0%R -> straight_path x y r == x.
+Proof. cbn -[mul opp RealVectorSpace.add]. intros ->. now rewrite mul_0, add_origin_r. Qed.
+
+Lemma straight_path_eq_1 (r : ratio) (x y:R2) : 
+  proj_ratio r == 1%R -> straight_path x y r == y.
+Proof. cbn -[mul opp RealVectorSpace.add]. intros ->. now rewrite mul_1, add_sub. Qed.
+
+Lemma odd_div2 k : Nat.Odd k -> (Nat.div2 k + Nat.div2 k = k - 1)%nat.
+Proof using .
+intros Hodd. rewrite <-Nat.odd_spec in Hodd. rewrite (Nat.div2_odd k) at 3.
+rewrite Hodd. cbn -[Nat.div2]. set (div2k := Nat.div2 k). lia.
+Qed.
+
+Lemma add_ratio_1_left (r1 r2 : ratio) : 
+  proj_ratio r1 == 1%R -> proj_ratio (add_ratio r1 r2) == 1%R.
+Proof. 
+unfold add_ratio. case (Rle_dec R1 (r1 + r2)) as [H | H].
++ reflexivity.
++ intros Hr. exfalso. apply H. cbn in *. generalize (ratio_bounds r2). lra.
+Qed.
+
+
+Definition BtoR : bool -> R := fun b => if b then 1%R else 0%R.
+
+Lemma BtoR_le b1 b2 : (BtoR b1 <= BtoR b2)%R <-> (b1 = true -> b2 = true).
+Proof using . unfold BtoR. repeat destruct_match ; intuition. Qed.
+
+Lemma BtoR_1_le b1 b2 : (BtoR b1 + 1 <= BtoR b2)%R <-> (b1 = false /\ b2 = true).
+Proof using . unfold BtoR. repeat destruct_match ; intuition ; lra. Qed.
+
+
+Definition Rltb x y := if Rlt_dec x y then true else false.
+
+Lemma Rltb_true x y : Rltb x y = true <-> (x < y)%R.
+Proof. unfold Rltb. case ifP_sumbool ; intuition. Qed.
+
+Lemma Rltb_false x y : Rltb x y = false <-> ~(x < y)%R.
+Proof. unfold Rltb. case ifP_sumbool ; intuition. Qed.
+
+Lemma even_div2_p1 (n : nat) : 
+  Nat.Even n -> Nat.div2 (n+1) = Nat.div2 n.
+Proof. 
+intros [k ->]. induction k as [|k IH].
++ rewrite Nat.mul_0_r, Nat.add_0_l. now cbn.
++ assert (H1 : (2 * S k + 1 = S (S (2 * k + 1)))%nat). { lia. }
+  assert (H2 : (2 * S k = S (S (2 * k)))%nat). { lia. }
+  rewrite H1, H2. cbn -[Nat.mul]. f_equal. exact IH.
+Qed.
+ 
+Lemma fold_left_Rmin_le_start l d : 
+  (fold_left Rmin l d <= d)%R.
+Proof. 
+revert d. induction l as [|x l IH] ; cbn.
++ reflexivity.
++ intros d. now rewrite IH, Rmin_l.
+Qed.
+
+Lemma fold_left_Rmin_le l t d : 
+  (exists x, InA equiv x l /\ (x <= t)%R) -> 
+  (fold_left Rmin l d <= t)%R.
+Proof.
+intros [x [Hin Hx]]. revert d. induction l as [|y l IH] ; cbn.
++ inv Hin.
++ intros d. rewrite InA_cons in Hin. case Hin as [-> | Hin].
+  - now rewrite <-Hx, fold_left_Rmin_le_start, Rmin_r.
+  - rewrite IH ; [reflexivity | apply Hin].
+Qed.  
+
+Lemma fold_left_Rmin_ge l t d :
+  (forall x, InA equiv x (d :: l) -> (t <= x)%R) -> 
+  (t <= fold_left Rmin l d)%R.
+Proof. 
+revert d. induction l as [|y l IH] ; cbn.
++ intuition.
++ intros d H. apply IH.
+  intros x Hin. apply H. rewrite InA_cons in Hin.
+  case Hin as [-> | Hin].
+  - unfold Rmin. case ifP_sumbool ; intros _ ; intuition.
+  - intuition.
+Qed.
+
+Lemma fold_left_Rmin_le_InA l d :
+  forall x, InA equiv x (d :: l) -> (fold_left Rmin l d <= x)%R. 
+Proof.
+intros x Hin. rewrite InA_cons in Hin. case Hin as [-> | Hin] ; [apply fold_left_Rmin_le_start|]. 
+revert d. induction Hin as [y l Heq | y l Hin IH] ; intros d.
++ cbn. rewrite fold_left_Rmin_le_start, Rmin_r, Heq. reflexivity.
++ cbn. rewrite IH. reflexivity.
+Qed. 
+
+Lemma fold_left_Rmin_InA l d : 
+  InA equiv (fold_left Rmin l d) (d :: l).
+Proof.
+revert d. induction l as [|x l IH].
++ cbn -[equiv]. intros d. now left.
++ cbn -[equiv]. intros d. unfold Rmin at 2. case ifP_sumbool ; intros _.
+  - assert (Hincl : inclA equiv (d :: l) (d :: x :: l)).
+    { intros ? Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin] ; intuition. }
+    apply Hincl, IH.
+  - right. apply IH.
+Qed.
+
+Lemma fold_left_Rmin_PermutationA l1 l2 d1 d2 : 
+  PermutationA equiv (d1 :: l1) (d2 :: l2) -> 
+  fold_left Rmin l1 d1 == fold_left Rmin l2 d2.
+Proof. 
+intros Hperm. apply Rle_antisym.
++ apply fold_left_Rmin_le_InA. rewrite Hperm. apply fold_left_Rmin_InA.
++ apply fold_left_Rmin_le_InA. rewrite <-Hperm. apply fold_left_Rmin_InA.
+Qed.
+
+Lemma fold_left_Rmax_ge_start l d : 
+  (d <= fold_left Rmax l d)%R.
+Proof. 
+revert d. induction l as [|x l IH] ; cbn.
++ reflexivity.
++ intros d. now rewrite <-IH, <-Rmax_l.
+Qed.
+
+Lemma fold_left_Rmax_ge l t d : 
+  (exists x, InA equiv x l /\ (t <= x)%R) -> 
+  (t <= fold_left Rmax l d)%R.
+Proof. 
+intros [x [Hin Hx]]. revert d. induction l as [|y l IH] ; cbn.
++ inv Hin.
++ intros d. rewrite InA_cons in Hin.
+  case Hin as [ Heq | Hin].
+  - (* TODO: understand why we need this. *)
+    unfold R_Setoid,eq_setoid in Heq. cbn in Heq. rewrite Heq in Hx. now rewrite Hx, <-fold_left_Rmax_ge_start, <-Rmax_r. 
+  - rewrite IH ; [reflexivity | apply Hin].
+Qed.
+
+Lemma fold_left_Rmax_le l t d :
+  (forall x, InA equiv x (d :: l) -> (x <= t)%R) -> 
+  (fold_left Rmax l d <= t)%R.
+Proof.
+revert d. induction l as [|y l IH] ; cbn.
++ intuition.
++ intros d H. apply IH.
+  intros x Hin. apply H. rewrite InA_cons in Hin.
+  case Hin as [-> | Hin].
+  - unfold Rmax. case ifP_sumbool ; intros _ ; intuition.
+  - intuition.
+Qed.
+
+Lemma fold_left_Rmax_ge_InA l d :
+  forall x, InA equiv x (d :: l) -> (x <= fold_left Rmax l d)%R. 
+Proof. 
+intros x Hin. rewrite InA_cons in Hin. case Hin as [-> | Hin] ; [apply fold_left_Rmax_ge_start|]. 
+revert d. induction Hin as [y l Heq | y l Hin IH] ; intros d.
++ cbn. rewrite <-fold_left_Rmax_ge_start, <-Rmax_r, Heq. reflexivity.
++ cbn. rewrite <-IH. reflexivity.
+Qed.  
+
+Lemma fold_left_Rmax_InA l d : 
+  InA equiv (fold_left Rmax l d) (d :: l).
+Proof.
+revert d. induction l as [|x l IH].
++ cbn -[equiv]. intros d. now left.
++ cbn -[equiv]. intros d. unfold Rmax at 2. case ifP_sumbool ; intros _.
+  - right. apply IH.
+  - assert (Hincl : inclA equiv (d :: l) (d :: x :: l)).
+    { intros ? Hin. rewrite InA_cons in Hin. destruct Hin as [-> | Hin] ; intuition. }
+    apply Hincl, IH.
+Qed.
+
+Lemma fold_left_Rmax_PermutationA l1 l2 d1 d2 : 
+  PermutationA equiv (d1 :: l1) (d2 :: l2) -> 
+  fold_left Rmax l1 d1 == fold_left Rmax l2 d2.
+Proof. 
+intros Hperm. apply Rle_antisym.
++ apply fold_left_Rmax_ge_InA. rewrite <-Hperm. apply fold_left_Rmax_InA.
++ apply fold_left_Rmax_ge_InA. rewrite Hperm. apply fold_left_Rmax_InA.
+Qed.
+
+Lemma fold_left_Rmin_Rmax l d : 
+  (fold_left Rmin l d == - fold_left Rmax (List.map Ropp l) (-d))%R.
+Proof. 
+revert d. induction l as [|x l IH] ; intros d ; cbn.
++ now rewrite Ropp_involutive.
++ rewrite IH, Ropp_Rmin. reflexivity.
+Qed.
+
+Lemma filter_nilA {A} (eqA : relation A) (f : A -> bool) l : 
+  Reflexive eqA ->
+  Proper (eqA ==> eq) f ->  
+  List.filter f l = nil <-> (forall x, InA eqA x l -> f x = false).
+Proof. 
+intros Hcompat. induction l as [|x l IH]. 
++ cbn. setoid_rewrite InA_nil. intuition.
++ split.
+  - cbn. case ifP_bool ; [discriminate|]. intros Hf Hnil y Hin.
+    rewrite InA_cons in Hin. case Hin as [-> | Hin].
+    * assumption.
+    * rewrite IH in Hnil by assumption. now apply Hnil.
+  - intros Hf. cbn.
+    assert (Hcond : f x = false). { apply Hf. now left. }
+    rewrite Hcond, IH by assumption. intros y Hin. apply Hf. now right.
+Qed.     
+
+(* This is a slightly stronger version of Stream.eventually,
+ * and is often more convenient to use in proofs. *)
+Inductive eventually_strong {A} (P : Stream.t A -> Prop) s : Prop := 
+ | ES_Now : P s -> eventually_strong P s
+ | ES_Later: ~P s -> eventually_strong P (Stream.tl s) -> eventually_strong P s.
+
+(* Whenever [P] is decidable, the strong and weak versions of eventually
+ * are equivalent *)
+Lemma eventually_strengthen {A} (P : Stream.t A -> Prop) s : 
+ (forall x, {P x} + {~P x}) ->
+ Stream.eventually P s -> eventually_strong P s. 
+Proof.
+intros Hdec HE. induction HE as [s HP | s Hlater IH].
++ now left.
++ case (Hdec s) as [HP | HNP].
+ - now left.
+ - now right.
+Qed. 
+
+Lemma div2_sum_p1_ge n : n <= Nat.div2 (n+1) + Nat.div2 (n+1).
+Proof.
+case (Nat.Even_or_Odd n) as [[k ->] | [k ->]].
++ rewrite Nat.add_1_r, Nat.div2_succ_double. lia.
++ rewrite 2 Nat.add_1_r. cbn -[Nat.mul]. rewrite Nat.div2_double. lia.
+Qed.
+
+Lemma Rmin_nonneg (x y : R) : (0 <= x)%R -> (0 <= y)%R -> (0 <= Rmin x y)%R.
+Proof. intros. unfold Rmin. destruct (Rle_dec x y) ; assumption. Qed.
+
+Lemma Rmax_nonneg_left (x y : R) : (0 <= x)%R -> (0 <= Rmax x y)%R.
+Proof. intros. transitivity x ; [assumption | apply Rmax_l]. Qed.
+
+Lemma Rmax_nonneg_right (x y : R) : (0 <= y)%R -> (0 <= Rmax x y)%R.
+Proof. intros. transitivity y ; [assumption | apply Rmax_r]. Qed.
+
+Lemma Rmax_pos_left (x y : R) : (0 < x)%R -> (0 < Rmax x y)%R.
+Proof. intros. apply (Rlt_le_trans _ x) ; [assumption | apply Rmax_l]. Qed.
+
+Lemma Rmax_pos_right (x y : R) : (0 < y)%R -> (0 < Rmax x y)%R.
+Proof. intros. apply (Rlt_le_trans _ y) ; [assumption | apply Rmax_r]. Qed.
+
+Ltac pos_R_close := solve [ easy | lra | nra | lia | nia ].
+
+(* Attempt to prove a goal of the form 0%R < ?X or 0%R <= ?X. *)
+(* This either fails or succeeds completely (i.e. it does not leave any subgoal open). *)
+(* If you have inequalities you would want pos_R to use, consider using generalize before pos_R. *)
+Ltac pos_R :=
+  match goal with 
+  (* Using hypotheses. *)
+  | [ |- _ -> _] => intros ; pos_R
+  | [ H : ?P |- ?P ] => exact H
+  (* Change Rgt/Rge to Rlt/Rle. *)
+  | [ |- (_ > _)%R ] => apply Rlt_gt ; pos_R
+  | [ |- (_ >= _)%R ] => apply Rle_ge ; pos_R
+  (* Addition. *)
+  | [ |- (0 < ?X + ?Y)%R ] => 
+    solve [ apply Rplus_lt_0_compat ; pos_R
+          | apply Rplus_lt_le_0_compat ; pos_R
+          | apply Rplus_le_lt_0_compat ; pos_R ]
+  | [ |- (0 <= ?X + ?Y)%R ] => apply Rplus_le_le_0_compat ; pos_R
+  (* Multiplication. *)
+  | [ |- (0 < ?X * ?Y)%R ] => apply Rmult_lt_0_compat ; pos_R
+  | [ |- (0 <= ?X * ?Y)%R ] => apply Rmult_le_pos ; pos_R
+  (* Inverse. *)
+  | [ |- (0 < / _)%R ] => apply Rinv_0_lt_compat ; pos_R
+  (* Division. *)
+  | [ |- (0 < _ / _)%R ] => apply Rdiv_lt_0_compat ; pos_R
+  | [ |- (0 <= _ / _)%R ] => apply Rdiv_le_0_compat ; pos_R
+  (* Square. *)
+  | [ |- (0 < Rsqr _)%R ] => apply Rlt_0_sqr ; pos_R_close
+  | [ |- (0 <= Rsqr _)%R ] => apply Rle_0_sqr ; pos_R_close
+  (* Absolute value. *)
+  | [ |- (0 < Rabs _)%R ] => apply Rabs_pos_lt ; pos_R_close
+  | [ |- (0 <= Rabs _)%R ] => apply Rabs_pos ; pos_R_close
+  (* Norm. *)
+  | [ |- (0 <= norm _)%R ] => apply norm_nonneg ; pos_R_close
+  (* Distance. *)
+  | [ |- (0 <= dist _ _)%R ] => apply dist_nonneg ; pos_R_close
+  (* Exponential. *)
+  | [ |- (0 < exp _)%R ] => apply exp_pos ; pos_R_close
+  | [ |- (0 <= exp _)%R ] => apply Rlt_le, exp_pos ; pos_R_close
+  (* INR. *)
+  | [ |- (0 < INR _)%R ] => apply lt_0_INR ; pos_R_close
+  | [ |- (0 <= INR _)%R ] => apply pos_INR ; pos_R_close
+  (* Minimum. *)
+  | [ |- (0 < Rmin _ _)%R ] => apply Rmin_pos ; pos_R
+  | [ |- (0 <= Rmin _ _)%R ] => apply Rmin_nonneg ; pos_R
+  (* Maximum. *)
+  | [ |- (0 < Rmax _ _)%R ] => 
+      solve [ apply Rmax_pos_left ; pos_R
+            | apply Rmax_pos_right ; pos_R ]
+  | [ |- (0 <= Rmax _ _)%R ] => 
+      solve [ apply Rmax_nonneg_left ; pos_R
+            | apply Rmax_nonneg_right ; pos_R ]
+  (* Strengthen Rle to Rlt. *)
+  | [ |- (_ <= _)%R ] => apply Rlt_le ; pos_R
+  (* Try to close the goal with lra/nra. *)
+  | _ => pos_R_close
+  end.
+
+Ltac simpl_R := repeat first 
+  [ rewrite Rmult_0_l
+  | rewrite Rmult_0_r
+  | rewrite Rmult_1_l
+  | rewrite Rmult_1_r
+  | rewrite Rplus_0_l
+  | rewrite Rplus_0_r
+  | rewrite Rminus_0_l
+  | rewrite Rminus_0_r
+  | rewrite Ropp_0 
+  | rewrite sqrt_0
+  | rewrite Rsqr_0
+  ].
+
+Ltac simpl_fct := 
+  cbv [fct_cte Ranalysis1.id Ranalysis1.comp plus_fct mult_fct inv_fct div_fct].
+
+Local Instance derivable_pt_lim_compat : Proper ((equiv ==> equiv) ==> equiv ==> equiv ==> iff) derivable_pt_lim.
+Proof.
+intros f1 f2 Hf x1 x2 Hx l1 l2 Hl. unfold derivable_pt_lim. split.
++ intros H eps Heps. destruct (H eps Heps) as [delta H']. exists delta.
+  intros h. rewrite <-Hl, <-Hx. rewrite <-(Hf _ _ (eq_refl x1)), <-(Hf _ _ (eq_refl (x1 + h)%R)).
+  now apply H'.
++ intros H eps Heps. destruct (H eps Heps) as [delta H']. exists delta.
+  intros h. rewrite Hl, Hx. rewrite (Hf _ _ (eq_refl x2)), (Hf _ _ (eq_refl (x2 + h)%R)).
+  now apply H'.
+Qed.
+
+Lemma derivable_pt_lim_inv :
+  forall (f : R -> R) (x l : R),
+    derivable_pt_lim f x l ->
+    f x <> 0%R ->
+    derivable_pt_lim (/ f) x ((- l) / Rsqr (f x))%R.
+Proof.
+intros f x l Hderiv Hf. 
+assert (Hdiv := derivable_pt_lim_div (fct_cte 1%R) f x 0%R l). feed_n 3 Hdiv.
+{ apply derivable_pt_lim_const. }
+{ assumption. }
+{ lra. }
+revert Hdiv. apply derivable_pt_lim_compat.
++ intros x1 x2 Hx. simpl_fct. simpl. rewrite Hx. lra.
++ reflexivity.
++ cbv [fct_cte]. now simpl_R.
+Qed.  
+
+(* Simplify a goal of the form [derivable_pt_lim f x d], assuming 
+ * f is not a [fun x => ...], i.e. is something like [plus_fct (comp _ _) (fct_cte _)] *)
+Ltac simp_derivable_rec :=
+  match goal with 
+  | [ |- derivable_pt_lim Ranalysis1.id _ _ ] => apply derivable_pt_lim_id
+  | [ |- derivable_pt_lim (Ranalysis1.fct_cte _) _ _ ] => apply derivable_pt_lim_const
+  | [ |- derivable_pt_lim sqrt _ _ ] => apply derivable_pt_lim_sqrt ; simp_derivable_rec
+  | [ |- derivable_pt_lim (_ + _)%F _ _ ] => apply derivable_pt_lim_plus ; simp_derivable_rec
+  | [ |- derivable_pt_lim (_ * _)%F _ _ ] => apply derivable_pt_lim_mult ; simp_derivable_rec
+  | [ |- derivable_pt_lim (_ / _)%F _ _ ] => apply derivable_pt_lim_div ; simp_derivable_rec
+  | [ |- derivable_pt_lim (/ _)%F _ _ ] => apply derivable_pt_lim_inv ; simp_derivable_rec
+  | [ |- derivable_pt_lim (Ranalysis1.comp _ _)%F _ _ ] => apply derivable_pt_lim_comp ; simp_derivable_rec
+  | _ => idtac
+  end.
+
+(* Recursively break up a goal of the form [derivable_pt_lim f x d]. 
+ * To have any hope this works, you better use an existential variable for [d],
+ * i.e. use [evar (e : R). replace d with e. simp_derivable. unfold e.] and then prove e = d. *)
+Ltac simp_derivable :=
+  lazymatch goal with 
+  | [ |- derivable_pt_lim ?f _ _ ] => 
+      let f_1 := eval cbv beta in (f Ranalysis_reg.AppVar) in 
+      let f_2 := Ranalysis_reg.rew_term f_1 in 
+      change f with f_2 ;
+      simp_derivable_rec
+  end.
+
+(* Try to prove a goal of the form [derivable_pt_lim f x d], where [d] can be any expression. *)
+Ltac prove_derivable_pt_lim := 
+  lazymatch goal with 
+  | [ |- derivable_pt_lim _ _ ?l ] => 
+      let l_evar := fresh "l" in 
+      evar (l_evar : R) ;
+      replace l with l_evar ; 
+      [simp_derivable | unfold l_evar] ;
+      simpl_fct
+  end.
diff --git a/CaseStudies/Gathering/InR2/Weber/Weber_point.v b/CaseStudies/Gathering/InR2/Weber/Weber_point.v
new file mode 100644
index 0000000000000000000000000000000000000000..ca8fa40c86db23433d691e600071e4d1e1a7afff
--- /dev/null
+++ b/CaseStudies/Gathering/InR2/Weber/Weber_point.v
@@ -0,0 +1,4039 @@
+
+(**************************************************************************)
+(**  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     *)
+(**************************************************************************)
+
+
+(**************************************************************************)
+(* Author : Mathis Bouverot-Dupuis (June 2022).
+
+ * This file gives the definition and some properties of the Weber point
+ * in the real plane (R²). The weber point (also called geometric median) 
+ * of a finite collection of points is unique if the points are not aligned,
+ * and has the property that moving any point towards the weber point 
+ * in a straight line doesn't change the weber point. *) 
+(**************************************************************************)
+
+
+Require Import Bool.
+Require Import Lia Field.
+Require Import Rbase Rbasic_fun R_sqrt Rtrigo_def Rpower RList Ranalysis.
+Require Import List.
+Require Import SetoidList.
+Require Import Relations.
+Require Import RelationPairs.
+Require Import Morphisms.
+Require Import Psatz.
+Require Import Inverse_Image.
+Require Import FunInd.
+Require Import FMapFacts.
+Require Import SetoidDec.
+
+(* Helping typeclass resolution avoid infinite loops. *)
+(* Typeclasses eauto := (bfs). *)
+
+Require Import Pactole.Spaces.R2.
+Require Import Pactole.Util.SetoidDefs.
+Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
+Require Import Pactole.Spaces.Similarity.
+Require Import Pactole.Util.Bijection.
+Require Export Pactole.CaseStudies.Gathering.InR2.Weber.Utils.
+(* This makes type class resolution go wrong. *)
+(* Remove Hints FMapFacts.eq_key_elt_Setoid FMapFacts.eq_key_Setoid FMapFacts.eq_key_elt_Equivalence FMapFacts.eq_key_Equivalence FMapFacts.Equal_ST : typeclass_instances. *)
+Require Export Pactole.CaseStudies.Gathering.InR2.Weber.Segment.
+Require Export Pactole.CaseStudies.Gathering.InR2.Weber.Line.
+
+
+Set Implicit Arguments.
+
+Global Arguments inclA_nil [A] {eqA} {_} l _.
+
+
+Close Scope R_scope.
+Close Scope VectorSpace_scope.
+
+
+Section WeberPoint.
+Implicit Types (points : list R2).
+
+Local Existing Instances R2_VS R2_ES ForallTriplets_PermutationA_compat.
+
+
+(* A Weber point of a finite collection P of points 
+ * is a point that minimizes the sum of the distances to elements of P.
+ * In general, there is no unique weber point for a collection.
+ * However it is unique if the points in P are not colinear. *)
+
+(* Compute the sum of the element of a list of real numbers [l] *)
+Fixpoint list_sum l :=
+  match l with 
+  | nil => 0%R 
+  | hd :: tl => (hd + list_sum tl)%R 
+  end.
+
+(* Typeclasses eauto := (dfs). *)
+Local Instance list_sum_compat : 
+  Proper (PermutationA equiv ==> equiv) list_sum.
+Proof.
+intros l l' Hll'. elim Hll'.
++ now reflexivity.
++ intros x x' t t' Hxx' Htt' IH. cbn -[equiv]. now rewrite Hxx', IH.
++ intros x y t. cbn -[equiv].
+  repeat rewrite <-Rplus_assoc. f_equiv. now rewrite Rplus_comm.
++ intros t t' t'' _ IH1 _ IH2. now rewrite IH1, IH2.
+Qed.
+
+Lemma list_sum_app xs ys : list_sum (xs ++ ys) = (list_sum xs + list_sum ys)%R.
+Proof. 
+induction xs as [| x xs IH] ; simpl.
++ lra.
++ rewrite IH. lra.
+Qed.
+
+Lemma list_sum_le l l' : Forall2 Rle l l' -> (list_sum l <= list_sum l')%R.
+Proof. 
+intros HF. induction HF as [| x x' l l' Hx Hl IH] ; try now auto.
+cbn -[equiv]. now apply Rplus_le_compat.
+Qed.
+
+Lemma Rminus_eq0_iff x y : (x - y)%R = 0%R <-> x = y. 
+Proof. lra. Qed.
+
+Lemma Rle_minus_sim r1 r2 : (r1 <= r2)%R -> (0 <= r2 - r1)%R.
+Proof. lra. Qed.
+
+Lemma eq_sym_iff {A : Type} (x y : A) : x = y <-> y = x.
+Proof. intuition. Qed.
+
+Lemma list_sum_le_eq l l' : 
+  Forall2 Rle l l' -> list_sum l == list_sum l' -> Forall2 equiv l l'.
+Proof. 
+intros HF. induction HF as [|x x' l l' Hx HF IH] ; cbn.
++ intros _. constructor.
++ intros Hsum.
+  assert (H : x = x' /\ list_sum l = list_sum l').
+  {
+    rewrite (eq_sym_iff x), (eq_sym_iff (list_sum l)). 
+    rewrite <-(Rminus_eq0_iff x'), <-(Rminus_eq0_iff (list_sum l')). apply Rplus_eq_R0.
+    - lra.
+    - apply list_sum_le in HF. lra.
+    - lra. 
+  }
+  constructor ; intuition.
+Qed.
+
+Lemma list_sum_le_eps l l' eps : Forall2 Rle l l' -> 
+  List.Exists (fun '(x, x') => (x <= x' + eps)%R) (combine l l') -> 
+  (list_sum l <= list_sum l' + eps)%R.
+Proof. 
+intros HF HE. induction HF as [| x x' l l' Hx Hl IH].
++ exfalso. rewrite Exists_exists in HE. destruct HE as [x [Hin_nil _]]. now apply in_nil in Hin_nil.
++ cbn -[equiv] in *. rewrite Exists_cons in HE. 
+  destruct HE as [Hx_eps | HE].
+  - apply list_sum_le in Hl. lra.
+  - apply IH in HE. lra.
+Qed.
+
+Lemma list_sum_ge_0 l : 
+  Forall (Rle 0%R) l -> (0 <= list_sum l)%R. 
+Proof.
+intros H. transitivity (list_sum (alls 0%R (length l))).
++ apply Req_le_sym. induction l as [|x l IH] ; auto.
+  cbn. rewrite IH ; try lra. rewrite Forall_cons_iff in H. apply H.
++ apply list_sum_le. induction l as [|x l IH] ; cbn ; auto.
+  rewrite Forall_cons_iff in H. apply Forall2_cons ; intuition.
+Qed. 
+
+Lemma list_sum_map_opp {A : Type} (f : A -> R) l : 
+  (-(list_sum (map f l)) == list_sum (map (fun x => -(f x)) l))%R.
+Proof. 
+induction l as [|x l IH]. 
++ cbn. now rewrite Ropp_0.
++ cbn. now rewrite Ropp_plus_distr, <-IH.
+Qed. 
+
+Lemma list_sum_map_mult_l {A : Type} (f : A -> R) l t :
+  (t * list_sum (map f l) == list_sum (map (fun x => t * f x) l))%R.
+Proof. 
+induction l as [|x l IH] ; cbn. 
++ lra.
++ rewrite <-IH. lra.
+Qed.
+
+Lemma list_sum_map_mult_r {A : Type} (f : A -> R) l t :
+  (list_sum (map f l) * t == list_sum (map (fun x => f x * t) l))%R.
+Proof. 
+induction l as [|x l IH] ; cbn. 
++ lra.
++ rewrite <-IH. lra.
+Qed. 
+
+Lemma list_sum_map_add {A : Type} (f g : A -> R) l :
+  (list_sum (map f l) + list_sum (map g l) == list_sum (map (fun x => f x + g x) l))%R.
+Proof.
+induction l as [|x l IH] ; cbn.
++ lra.
++ rewrite <-IH. lra.
+Qed. 
+
+Lemma list_sum_map_add_combine {A : Type} (f1 f2 : A -> R) l1 l2 : 
+  length l1 = length l2 -> 
+  (list_sum (map f1 l1) + list_sum (map f2 l2) == 
+    list_sum (map (fun '(x1, x2) => f1 x1 + f2 x2) (combine l1 l2)))%R.
+Proof.
+revert l2. induction l1 as [|x1 l1 IH] ; intros l2 Hlen ; cbn.
++ rewrite Rplus_0_l. case l2 as [|x2 l2] ; cbn in * ; [reflexivity | discriminate].
++ case l2 as [|x2 l2] ; cbn in * ; [discriminate|]. inv Hlen.
+ rewrite <-IH by auto. lra.
+Qed.
+
+(* This is the function that a weber point minimizes. *)
+Definition dist_sum points (x : R2) := 
+  list_sum (List.map (dist x) points).
+
+Local Instance dist_sum_compat : Proper (PermutationA equiv ==> equiv ==> equiv) dist_sum.
+Proof using .
+intros p p' Hpp' x y Hxy. unfold dist_sum. f_equiv. now rewrite Hpp', Hxy.
+Qed.
+
+Lemma dist_sum_nonneg points x : (0 <= dist_sum points x)%R.
+Proof.
+apply list_sum_ge_0. rewrite Forall_forall. intros d Hin. rewrite in_map_iff in Hin.
+destruct Hin as [y [Hd Hin]]. rewrite <-Hd. apply dist_nonneg.
+Qed.
+
+(* [argmin f] is the set of elements that minimize [f],
+ * when sets are represented by membership predicates. *)
+Definition argmin {A : Type} (f : A -> R) : A -> Prop := 
+  fun a => forall b, (f a <= f b)%R.
+  
+Local Instance argmin_compat {A : Type} : Proper (equiv ==> @eq A ==> iff) argmin.
+Proof using . 
+intros f g Hfg a a' Haa'. unfold argmin.
+rewrite <-Haa'. setoid_rewrite <-Hfg.
+reflexivity.
+Qed.
+
+(* The set of Weber points of a finite collection of points *)
+Definition Weber points : R2 -> Prop := argmin (dist_sum points).
+
+(* [Weber] doesn't depend on the order of the points. *)
+Global Instance weber_compat : Proper (PermutationA equiv ==> equiv ==> iff) Weber.
+Proof using .
+  intros p p' Hpp' x y Hxy. unfold Weber. f_equiv ; try auto. intros z. now f_equiv.
+Qed.
+
+(* [OnlyWeber ps w] means that [w] is the unique weber point of [ps]. *)
+Definition OnlyWeber ps w : Prop := Weber ps w /\ (forall x, Weber ps x -> x == w).
+
+Global Instance only_weber_compat : Proper (PermutationA equiv ==> equiv ==> iff) OnlyWeber.
+Proof. intros ? ? H1 ? ? H2. unfold OnlyWeber. setoid_rewrite H1. setoid_rewrite H2. reflexivity. Qed.
+
+Lemma weber_gathered ps x0 :
+  Forall (equiv x0) (x0 :: ps) -> OnlyWeber (x0 :: ps) x0.
+Proof.
+intros H. rewrite Forall_cons_iff in H. destruct H as [_ H].
+assert (Hpos : forall x ps, (0 <= dist_sum ps x)%R).
+{ intros x ps'. apply list_sum_ge_0. rewrite Forall_map, Forall_forall.
+  intros y _. apply dist_nonneg. }
+assert (H0 : dist_sum (x0 :: ps) x0 == 0%R).
+{ cbn -[dist]. rewrite dist_same, Rplus_0_l.
+  induction H as [|x l Hx Hl IH].
+  + cbn. reflexivity.
+  + cbn -[dist]. rewrite Hx at 1. rewrite dist_same, Rplus_0_l. now apply IH. }
+split.
++ intros y. rewrite H0. apply Hpos.
++ intros y Hweb.
+  assert (Hy : dist_sum (x0 :: ps) y == 0%R).  
+  { apply Rle_antisym.
+    - specialize (Hweb x0). rewrite H0 in Hweb. exact Hweb.
+    - apply Hpos. }
+  cbn -[dist] in Hy. apply Rplus_eq_R0 in Hy.
+  - apply dist_defined, Hy.
+  - apply dist_nonneg.
+  - apply Hpos.   
+Qed.
+
+(* This lemma isn't needed. *)
+(* We can show that a weber point can equivalently be defined as
+ * an argmin on a compact set of points (instead of an argmin on the whole plane),
+ * and a continuous function always has a minimum on a compact set. *)
+(* RMK : this is also true if [points] is empty. *)
+(*Lemma weber_exists points : 
+  exists x, Weber points x.
+Proof. Admitted.*)
+
+Lemma dist_middle_leq (a b x : R2) : 
+  (dist x (middle a b) <= (dist x a + dist x b) / 2)%R.
+Proof. 
+unfold middle, Rdiv. repeat rewrite norm_dist. apply Rmult_le_reg_l with 2%R ; try lra. 
+rewrite <-Rmult_assoc, Rinv_r_simpl_m by lra. rewrite <-(Rabs_pos_eq 2%R) at 1 by lra. 
+rewrite <-norm_mul, mul_distr_add by lra. rewrite <-mul_opp, mul_morph, Rmult_1_l, Rinv_r, mul_1 by lra.
+change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1.
+rewrite <-RealVectorSpace.add_assoc, opp_distr_add, (RealVectorSpace.add_assoc x (-a)).
+rewrite (RealVectorSpace.add_comm _ (-b)), RealVectorSpace.add_assoc.
+rewrite Rplus_comm. apply triang_ineq.
+Qed.
+
+Lemma Rmult_reg_iff_l r r1 r2 : 
+  ~ r == 0%R -> (r * r1 == r * r2 <-> r1 == r2)%R.
+Proof. cbn. nra. Qed.
+
+Lemma dist_middle_eq_iff (a b x : R2) : 
+  ~ a == b ->
+  (dist x (middle a b) == (dist x a + dist x b) / 2)%R <-> 
+  (exists t, (0 <= t)%R /\ (x - a == t * (a - b) \/ x - b == t * (b - a))%VS).
+Proof.
+intros Hab. unfold middle, Rdiv. 
+repeat rewrite norm_dist. rewrite <-(@Rmult_reg_iff_l 2%R) by (cbn ; lra). 
+rewrite <-Rmult_assoc, Rinv_r_simpl_m by lra. rewrite <-(Rabs_pos_eq 2%R) at 1 by lra. 
+rewrite <-norm_mul, mul_distr_add by lra. rewrite <-mul_opp, mul_morph, Rmult_1_l, Rinv_r, mul_1 by lra.
+change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1.
+rewrite <-RealVectorSpace.add_assoc, opp_distr_add, (RealVectorSpace.add_assoc x (-a)).
+rewrite (RealVectorSpace.add_comm _ (-b)), RealVectorSpace.add_assoc, Rplus_comm.
+symmetry ; split.
++ intros [t [Ht [Hx1 | Hx1]]] ; rewrite Hx1.
+  - assert (Hx2 : (x - b == (t + 1) * (a - b))%VS).
+    { rewrite <-add_morph, mul_1, <-Hx1, RealVectorSpace.add_assoc. f_equiv. 
+      now rewrite RealVectorSpace.add_comm, add_sub. }
+    rewrite Hx2, add_morph, 3 norm_mul, 3 Rabs_pos_eq by lra. 
+    cbn -[mul opp RealVectorSpace.add norm]. lra.
+  - assert (Hx2 : (x - a == (t + 1) * (b - a))%VS).
+    { rewrite <-add_morph, mul_1, <-Hx1, RealVectorSpace.add_assoc. f_equiv. 
+      now rewrite RealVectorSpace.add_comm, add_sub. }
+    rewrite Hx2, add_morph, 3 norm_mul, 3 Rabs_pos_eq by lra. 
+    cbn -[mul opp RealVectorSpace.add norm]. lra.
++ intros H.
+  assert (Hcol : colinear (x - a) (x - b)).
+  { 
+    apply colinear_inner_product_spec.
+    apply (f_equal Rsqr) in H. rewrite squared_norm_product in H.
+    rewrite inner_product_add_l in H. setoid_rewrite inner_product_add_r in H.
+    rewrite <- 2 squared_norm_product in H.
+    rewrite R_sqr.Rsqr_plus in H. rewrite inner_product_sym in H.
+    assert (Hprod : inner_product (x - a) (x - b) = (norm (x - a) * norm (x - b))%R) by nra.
+    apply (f_equal Rabs) in Hprod. setoid_rewrite Rabs_pos_eq in Hprod at 2.
+    - exact Hprod.
+    - apply Rmult_le_pos ; apply norm_nonneg.
+  }
+  rewrite <-colinear_opp_compat_r, <-colinear_add in Hcol.
+  assert (H1 : (x - a - (x - b) == b - a)%VS).
+  { rewrite opp_distr_add, opp_opp, RealVectorSpace.add_assoc, RealVectorSpace.add_comm.
+    f_equiv. now rewrite (RealVectorSpace.add_comm x), <-RealVectorSpace.add_assoc, add_opp, add_origin_r. }
+  rewrite H1 in Hcol. clear H1. symmetry in Hcol.
+  apply colinear_decompose in Hcol. 2: rewrite R2sub_origin ; intuition.
+  pose (t := (norm (x - a) * / (norm (b - a)))%R).
+  unfold unitary in Hcol. rewrite 2 mul_morph, <-Ropp_mult_distr_l in Hcol. fold t in Hcol.
+  assert (Ht1 : (0 <= t)%R).
+  { unfold t. apply Rle_mult_inv_pos ; [apply norm_nonneg |].
+    apply Rle_neq_lt ; [apply norm_nonneg |]. rewrite <-norm_dist. intros Hab'.
+    apply Hab, dist_defined. now rewrite dist_sym. }
+  destruct Hcol as [Hxa | Hxa].
+  - assert (Hxb : (x - b == (t - 1) * (b - a))%VS).
+    { unfold Rminus. rewrite <-add_morph, minus_morph, mul_1, <-Hxa, opp_distr_add.
+      rewrite <-RealVectorSpace.add_assoc. now rewrite add_sub. } 
+    rewrite Hxa, Hxb, add_morph, 3 norm_mul, <-Rmult_plus_distr_r, (Rabs_pos_eq t) in H by auto.
+    apply Rmult_eq_reg_r in H. 
+    2: rewrite <-norm_dist, dist_sym ; intros Hdist ; now apply Hab, dist_defined.
+    case (Rle_dec 1 t) as [Ht2 | Ht2].
+    * exists (t - 1)%R. split ; [lra | now right].
+    * case (Rle_dec (1 / 2)%R t) as [Ht3 | Ht3]. 
+      ++rewrite Rabs_pos_eq, Rabs_left1 in H by lra. assert (Ht : t = 0%R) by lra. lra.
+      ++rewrite 2 Rabs_left1 in H by lra. assert (Ht : t = 0%R) by lra.
+        exists 0%R. split ; [lra | left]. now rewrite Hxa, Ht, 2 mul_0.
+  - exists t. split ; [auto | left]. rewrite Hxa, minus_morph, <-mul_opp, opp_distr_add, opp_opp.
+    now rewrite RealVectorSpace.add_comm.
+Qed.
+  
+(* If the points aren't aligned, then the weber point is unique. *)
+Lemma weber_Naligned_unique points w : 
+  ~aligned points -> Weber points w -> OnlyWeber points w.
+Proof.
+intros HNalign Hweb. split ; [auto|]. intros w2 Hweb2.
+case (w =?= w2) as [Heq | HNeq] ; [intuition | exfalso].
+apply HNalign. apply aligned_tail with w. rewrite aligned_spec.
+exists (w - w2)%VS. intros x Hin.
+cut (dist x (middle w w2) == (dist x w + dist x w2) / 2)%R. 
+{ 
+  rewrite dist_middle_eq_iff by intuition. intros [t [Ht [Hx | Hx]]].
+  + exists t. now rewrite <-Hx, add_sub.
+  + exists (-(1 + t))%R. rewrite minus_morph, <-mul_opp, opp_distr_add, opp_opp, (RealVectorSpace.add_comm _ w2).
+    rewrite <-add_morph, mul_1, <-Hx. now rewrite RealVectorSpace.add_assoc, 2 add_sub. 
+} 
+pose (l1 := map (dist (middle w w2)) points).
+pose (l2 := map (fun x => (dist w x + dist w2 x) / 2)%R points).
+assert (H := @list_sum_le_eq l1 l2). feed_n 2 H.
++ unfold l1, l2. rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length.
+  intros [x1 x2]. rewrite in_combine_id. intros [-> _]. rewrite 3 (dist_sym x2). apply dist_middle_leq.
++ unfold l1, l2. apply Rle_antisym.
+  - apply list_sum_le. rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length.
+    intros [x1 x2]. rewrite in_combine_id. intros [-> _]. rewrite 3 (dist_sym x2). apply dist_middle_leq.
+  - unfold Rdiv. rewrite <-list_sum_map_mult_r.
+    apply Rmult_le_reg_l with 2%R ; try lra.
+    rewrite <-Rmult_assoc, Rinv_r_simpl_m, <-list_sum_map_add by lra.
+    change 2%R with (1 + 1)%R. rewrite Rmult_plus_distr_r, Rmult_1_l.
+    apply Rplus_le_compat ; [apply Hweb | apply Hweb2].
++ revert H. unfold l1, l2. rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length.
+  intros H. rewrite 3 (dist_sym _ x). specialize (H (x, x)). apply H. 
+  rewrite in_combine_id. split ; [reflexivity|]. rewrite <-InA_Leibniz. exact Hin.
+Qed.
+
+Lemma dist_sum_similarity (f : similarity R2) points x : 
+  (dist_sum (List.map f points) (f x) == f.(zoom) * dist_sum points x)%R.
+Proof. 
+unfold dist_sum. rewrite map_map, list_sum_map_mult_l. f_equiv.
+apply eqlistA_PermutationA.
+(* f_equiv triggers some dumb inference making FMapInterface.eq_key appear from nowhere... *)
+(* we force the relation here to avoid that. *)
+specialize @map_extensionalityA_compat with (eqA:=@equiv R2 R2_Setoid)(eqB:=@equiv R R_Setoid) as h.
+unfold Proper,respectful in h.
+eapply h;try typeclasses eauto;try reflexivity.
+intros x0 y H. 
+now rewrite H, dist_prop.
+Qed.
+
+Lemma weber_similarity_weak points w (f : similarity R2) : 
+  Weber points w -> Weber (List.map f points) (f w).
+Proof.
+unfold Weber, argmin. intros H b. 
+rewrite <-(Bijection.section_retraction f b). 
+remember (Bijection.retraction f b) as b'.
+repeat rewrite dist_sum_similarity. apply Rmult_le_compat_l.
++ now apply Rlt_le, zoom_pos.
++ now apply H.
+Qed.
+
+(* A weber point is preserved by similarities. 
+ * This is important because it means that all robots will calculate the same correct weber point, 
+ * even though they view the configuration up to a change of frame (i.e. a similarity). *)
+Lemma weber_similarity points w (f : similarity R2) : 
+  Weber points w <-> Weber (List.map f points) (f w).
+Proof.
+split.
++ now apply weber_similarity_weak.
++ intros H. apply (@weber_similarity_weak _ _ (f ⁻¹)) in H.
+  revert H. rewrite map_map.
+  pose (f1f := fun x => (f ⁻¹) (f x)). 
+  fold f1f. change ((f ⁻¹) (f w)) with (f1f w).
+  assert (forall x, f1f x == x).
+  { intros x. cbn -[equiv]. now rewrite Bijection.retraction_section. }
+  apply weber_compat.
+  - apply eqlistA_PermutationA. rewrite <-List.map_id at 1.
+    specialize @map_extensionalityA_compat with (eqA:=@equiv R2 R2_Setoid)(eqB:=@equiv R2 R2_Setoid) as h.
+    unfold Proper,respectful in h.
+    eapply h;try typeclasses eauto;try reflexivity.
+    intros x y Hxy. now rewrite H.
+  - now rewrite H.
+Qed.  
+
+Lemma list_sum_lt l l' : 
+  l <> nil -> Forall2 Rlt l l' -> (list_sum l < list_sum l')%R.
+Proof. 
+intros Hnil HF. induction HF as [|x x' l l' Hx Hl IH] ; cbn.
++ intuition.
++ apply Rplus_lt_le_compat ; auto. case l as [|y l].
+  - inv Hl. reflexivity.
+  - apply Rlt_le, IH. discriminate.
+Qed.
+
+(* A weber point of aligned points is on the same line. *)
+Lemma weber_aligned L ps w : 
+  ps <> nil ->
+  aligned_on L ps -> 
+  Weber ps w -> 
+  line_contains L w.
+Proof. 
+intros Hnil Halign Hweb. 
+case (line_dir L =?= 0%VS) as [Hdir | Hdir].
++ cut (w == line_org L). { intros ->. apply line_contains_org. }
+  assert (H : forall x, InA equiv x ps -> x == line_org L).
+  { intros x Hin. rewrite <-Halign, line_proj_spec, line_proj_inv_spec by assumption.
+    rewrite Hdir, inner_product_origin_r. now rewrite mul_origin, add_origin_r. }
+  case ps as [|x0 ps] ; [tauto|]. apply H. left.
+  apply (@weber_gathered ps x0).
+  - rewrite Forall_forall. intros x Hin. apply (@In_InA _ equiv) in Hin ; autoclass.
+    rewrite (H x), (H x0) ; [reflexivity | now left | assumption]. 
+  - assumption.
++ (* [p] is the orthogonal projection of [w] on the [L]. *)
+  pose (p := L^-P (L^P w)).
+  assert (Hline_p : line_contains L p).
+  { unfold p, line_contains. now rewrite line_P_iP by assumption. }
+  pose (v := line_dir L).
+  case (w =?= p) as [Hwp | Hwp].
+  (* [w] is on the line. *)
+  - rewrite Hwp. unfold p, line_contains. now rewrite line_P_iP.
+  (* [w] is not on the line : contradiction. *)
+  - exfalso. 
+    cut (dist_sum ps p < dist_sum ps w)%R. { generalize (Hweb p). lra. }
+    apply list_sum_lt. { now rewrite <-length_zero_iff_nil, map_length, length_zero_iff_nil. }
+    rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now repeat rewrite map_length.
+    intros [x x2]. rewrite in_combine_id. intros [<- Hin].
+    assert (Hp : (p - x = (inner_product (w - x) v / (norm v)²) * v)%VS).
+    {
+      unfold p. rewrite <-(Halign x) at 1 by (apply In_InA ; autoclass).
+      rewrite 2 line_proj_inv_spec.
+      (* TODO: why do we need to make the type of a b c explicit? *)
+      assert (Hrewrite : forall a b c:R2, (a + b - (a + c) == b - c)%VS).
+      { intros a b c. rewrite (RealVectorSpace.add_comm a b), <-RealVectorSpace.add_assoc.
+        apply add_compat ; [reflexivity|]. now rewrite opp_distr_add, RealVectorSpace.add_assoc, add_opp, add_origin_l. }
+      rewrite Hrewrite. clear Hrewrite. fold v. rewrite <-minus_morph, add_morph. f_equal.
+      unfold Rdiv. rewrite Ropp_mult_distr_l, <-Rmult_plus_distr_r. f_equal.
+      rewrite 2 line_proj_spec. rewrite <-inner_product_opp_l, <-inner_product_add_l.
+      assert (Hrewrite : forall a b c:R2, (b - a - (c - a) == b - c)%VS). 
+      { intros a b c. rewrite opp_distr_add, opp_opp.
+        rewrite (RealVectorSpace.add_comm _ a), <-RealVectorSpace.add_assoc.
+        apply add_compat ; [reflexivity|]. now rewrite RealVectorSpace.add_assoc, (RealVectorSpace.add_comm _ a), add_opp, add_origin_l. }
+      rewrite Hrewrite. clear Hrewrite. reflexivity.
+    }
+    assert (Hwx : (w - x == w - p + (p - x))%VS).
+    { rewrite <-RealVectorSpace.add_assoc. f_equiv. rewrite RealVectorSpace.add_assoc.
+      now rewrite (RealVectorSpace.add_comm _ p), add_opp, add_origin_l. }
+    rewrite 2 norm_dist, <-pos_Rsqr_lt, Hwx by apply norm_nonneg.
+    assert (Hperp : perpendicular (w - p) (p - x)).
+    {
+      unfold perpendicular.
+      assert (Hrewrite : (w - p == w - x - (p - x))%VS).
+      { now rewrite <-RealVectorSpace.add_assoc, opp_distr_add, add_sub. }
+      rewrite Hrewrite, inner_product_add_l, inner_product_opp_l, Hp, 2 inner_product_mul_r, inner_product_mul_l.
+      unfold Rdiv. repeat rewrite Rmult_assoc. rewrite squared_norm_product, Rinv_l, Rmult_1_r ; [lra|].
+      now rewrite <-squared_norm_product, <-Rsqr_0, <-square_norm_equiv, norm_defined by lra. 
+    }
+    apply Pythagoras in Hperp. rewrite Hperp.
+    cut (0 < (norm (w - p))²)%R. { lra. }
+    rewrite <-Rsqr_0. apply R_sqr.Rsqr_incrst_1 ; try solve [lra | apply norm_nonneg].
+    apply Rle_neq_lt ; [apply norm_nonneg|].
+    rewrite eq_sym_iff, norm_defined, R2sub_origin. intuition.
+Qed.
+
+(* [contract ps ps' x] means that [ps'] is obtained from [ps] by moving
+ * each point towards [x].*)
+Definition contract ps ps' (x:R2) : Prop := 
+  Forall2 (fun p p' => segment x p p') ps ps'.
+
+Lemma segment_dist_weak (x a b : R2) : 
+  segment a b x -> dist a b = (dist a x + dist x b)%R.
+Proof. 
+unfold segment. intros [t [Ht Hx]].
+assert (H1t : (0 <= 1 - t <= 1)%R). { lra. }
+pose (rt := exist (fun t => 0 <= t <= 1)%R t Ht).
+pose (r1t := exist (fun t => 0 <= t <= 1)%R (1 - t)%R H1t).
+assert (Hx1 : x == straight_path a b r1t).
+{ 
+cbn -[equiv mul RealVectorSpace.add opp]. unfold Rminus. 
+rewrite <-add_morph, mul_1, RealVectorSpace.add_assoc, add_sub.
+rewrite minus_morph, mul_distr_add, opp_distr_add, mul_opp, opp_opp.
+rewrite RealVectorSpace.add_assoc, RealVectorSpace.add_comm, Hx.
+f_equiv. unfold Rminus. rewrite <-add_morph, mul_1, minus_morph. reflexivity.
+}
+assert (Hx2 : x == straight_path b a rt).
+{
+cbn -[equiv mul RealVectorSpace.add opp]. rewrite Hx. 
+unfold Rminus. rewrite <-add_morph, mul_distr_add, mul_1, minus_morph, mul_opp.
+rewrite 2 RealVectorSpace.add_assoc. f_equiv. rewrite RealVectorSpace.add_comm.
+reflexivity.
+}
+rewrite Hx1 at 1. rewrite straight_path_dist_start.
+rewrite (dist_sym _ x), Hx2, straight_path_dist_start.
+cbn -[dist]. rewrite (dist_sym a b). lra.
+Qed.
+
+Lemma segment_dist_sum (a b x : R2) :
+  segment a b x <-> dist a b = (dist a x + dist x b)%R.
+Proof.
+split ; [apply segment_dist_weak|]. 
+case (a =?= b) as [Eab | NEab].
++ rewrite Eab, dist_same. rewrite dist_sym. intros Hdist. 
+  assert (H : dist x b = 0%R). { lra. }
+  rewrite dist_defined in H. rewrite H. apply segment_start.
++ intros Hdist. destruct (triang_ineq_eq _ _ _ Hdist) as [Hcol _].
+  apply colinear_exists_mul in Hcol. destruct Hcol as [t Hxa].
+  2: intros H ; apply NEab ; rewrite R2sub_origin in H ; now symmetry.
+  assert (Ht : (Rabs t + Rabs (t - 1) == 1)%R).
+  {
+    assert (Hxb : (x - b)%VS == ((t - 1) * (b - a))%VS).
+    { unfold Rminus. rewrite <-add_morph, minus_morph, mul_1, <-Hxa.
+      rewrite <-RealVectorSpace.add_assoc. f_equiv. 
+      now rewrite opp_distr_add, add_sub. }
+    rewrite (dist_sym x a), (dist_sym b a) in Hdist. repeat rewrite norm_dist in Hdist.
+    rewrite Hxa, Hxb in Hdist. repeat rewrite norm_mul in Hdist.
+    rewrite <-Rmult_plus_distr_r in Hdist.
+    rewrite <-(Rmult_1_l (norm (b - a)%VS)) in Hdist at 1.
+    apply Rmult_eq_reg_r in Hdist ; [now symmetry|].
+    rewrite norm_defined. intros H. apply NEab. rewrite R2sub_origin in H. now symmetry.
+  }
+  rewrite segment_sym. exists t. split.
+  - case (Rle_dec 0%R t) as [Rle_0t | NRle_0t] ; case (Rle_dec t 1%R) as [Rle_t1 | NRle_t1] ; [lra | exfalso ..]. 
+    * rewrite Rabs_pos_eq in Ht by auto. rewrite Rabs_pos_eq in Ht by lra.
+      assert (Ht1 : t == 1%R). { cbn in Ht |- *. lra. }
+      intuition.
+    * rewrite Rabs_left1 in Ht by lra. rewrite Rabs_left1 in Ht by lra.
+      assert (Ht0 : t == 0%R). { cbn in Ht |- *. lra. }
+      intuition.
+    * rewrite Rabs_left1 in Ht by lra. rewrite Rabs_pos_eq in Ht by lra. lra.
+  - unfold Rminus. rewrite <-add_morph, mul_1, minus_morph, <-mul_opp.
+    rewrite (RealVectorSpace.add_comm a), RealVectorSpace.add_assoc, <-mul_distr_add.
+    rewrite <-Hxa, <-RealVectorSpace.add_assoc, (RealVectorSpace.add_comm _ a), add_opp, add_origin_r.
+    reflexivity.
+Qed.
+
+Lemma segment_argmin (x a b : R2) : 
+  segment x a b <-> argmin (fun y => (dist y b - dist y a)%R) x.
+Proof.
+split.
++ intros Hsegment y. rewrite segment_dist_sum in Hsegment. rewrite Hsegment.
+  unfold Rminus at 1. rewrite Ropp_plus_distr, <-Rplus_assoc, Rplus_opp_r, Rplus_0_l.
+  cut (dist y a <= dist y b + dist b a)%R. { lra. }
+  apply RealMetricSpace.triang_ineq.
++ unfold argmin. intros Hdist. rewrite segment_dist_sum.
+  apply Rle_antisym ; [apply RealMetricSpace.triang_ineq |]. 
+  specialize (Hdist b). rewrite dist_same, Rminus_0_l in Hdist.
+  lra.
+Qed.
+
+Lemma argmin_sum2 {A : Type} (f g : A -> R) x0 x :
+  argmin f x0 -> argmin g x0 -> 
+  (argmin (fun x => (f x + g x)%R)x <-> argmin f x /\ argmin g x).
+Proof. 
+intros Hf0 Hg0. split ; unfold argmin in *.
++ intros Hfg.
+  cut (f x = f x0 /\ g x = g x0). { intros [-> ->]. intuition. }
+  rewrite <-(Rminus_eq0_iff (f x)), <-(Rminus_eq0_iff (g x)).
+  apply Rplus_eq_R0.
+  - apply Rle_minus_sim, Hf0.
+  - apply Rle_minus_sim, Hg0.
+  - cut ((f x + g x)%R = (f x0 + g x0)%R). { lra. }
+    apply Rle_antisym ; [apply Hfg|].
+    apply Rplus_le_compat ; auto.    
++ intros [Hf Hg] y. apply Rplus_le_compat.
+  - apply Hf.
+  - apply Hg.
+Qed. 
+
+(* When a collection of functions have a common argmin, 
+ * minimizing their sum is equivalent to minimizing each function. *)
+Lemma argmin_sum {I A : Type} (F : I -> A -> R) (x0 x : A) (li : list I) :
+  (Forall (fun i => argmin (F i) x0) li) -> 
+    (argmin (fun x => list_sum (map (fun i => F i x) li))) x <-> 
+    Forall (fun i => argmin (F i) x) li.
+Proof. 
+intros HF0. revert x. induction li as [|i li IH] ; intros x. 
++ cbn. split ; intros _ ; [constructor|].
+  unfold argmin. reflexivity.
++ rewrite Forall_cons_iff in HF0 |- *. destruct HF0 as [Hf0 HF0]. 
+  specialize (IH HF0). cbn. rewrite <-IH. apply argmin_sum2 with x0 ; auto.
+  rewrite IH. exact HF0.
+Qed.
+
+Lemma argmin_dist_sum_diff ps ps' x : 
+  length ps' = length ps -> 
+  argmin (fun y => (dist_sum ps' y - dist_sum ps y)%R) x <-> 
+  argmin (fun y => list_sum (map (fun '(p', p) => (dist y p' - dist y p)%R) (combine ps' ps))) x.
+Proof.
+intros Hlen. f_equiv. intros y. unfold Rminus, dist_sum. 
+rewrite list_sum_map_opp, list_sum_map_add_combine by now symmetry.
+reflexivity.
+Qed. 
+
+(* This is quite an amusing property. *)
+Lemma contract_argmin ps ps' x0 x : 
+  contract ps ps' x0 ->  
+  (contract ps ps' x <-> argmin (fun y => (dist_sum ps' y - dist_sum ps y)%R) x).
+Proof. 
+intros Hcontr0. split. 
+(* This direction is always true. *)
++ clear Hcontr0. intros Hcontr. rewrite argmin_dist_sum_diff by now symmetry ; eapply Forall2_length, Hcontr.
+  intros y. apply list_sum_le. 
+  rewrite Forall2_Forall, combine_map, Forall_map, Forall_forall by now rewrite 2 map_length.
+  intros [a b] Hin. rewrite in_combine_id in Hin. destruct Hin as [<- Hin].
+  destruct a as [p' p]. apply segment_argmin. assert (H := Hcontr). 
+  revert H. unfold contract. rewrite Forall2_Forall, Forall_forall by eapply Forall2_length, Hcontr.
+  intros H. specialize (H (p, p')). apply H. now rewrite in_combine_sym.
+(* This direction holds because the lines [pi, pi'[ have an intersection point [x0]. *)
++ assert (Hlen : length ps' = length ps). { symmetry. eapply Forall2_length, Hcontr0. }
+  rewrite argmin_dist_sum_diff by auto. rewrite argmin_sum.
+  - unfold contract. rewrite Forall2_Forall, 2 Forall_forall by now symmetry. 
+    intros Hargmin [p p'] Hin. specialize (Hargmin (p', p)). 
+    feed Hargmin. { now rewrite in_combine_sym. } rewrite segment_argmin. exact Hargmin.
+  - unfold contract in Hcontr0. rewrite Forall2_Forall, Forall_forall in Hcontr0 by now symmetry.
+    rewrite Forall_forall. intros [p' p] Hin. specialize (Hcontr0 (p, p')).
+    rewrite <-segment_argmin. apply Hcontr0. now rewrite in_combine_sym.
+Qed.
+
+(* See the thesis of Zohir Bouzid, lemma 3.1.5. *)
+Theorem weber_contract_strong ps ps' w0 : 
+  contract ps ps' w0 -> Weber ps w0 ->
+  (forall w, Weber ps' w <-> (Weber ps w /\ contract ps ps' w)).
+Proof. 
+intros Hcontract0 Hweb0 w. rewrite contract_argmin in * ; eauto. unfold Weber in *.
+rewrite <-(argmin_sum2 _ Hweb0 Hcontract0). f_equiv. 
+intros x. rewrite Rplus_minus. reflexivity.
+Qed.
+
+(* We can even move points onto the weber point, it will still be preserved. *)
+Corollary weber_contract ps ps' w : 
+  contract ps ps' w -> Weber ps w -> Weber ps' w.
+Proof. 
+intros Hcontract Hweb. rewrite (weber_contract_strong Hcontract Hweb). intuition.
+Qed.
+
+(* See the thesis of Zohir Bouzid, corollary 3.1.1. *)
+Corollary weber_contract_unique ps ps' w : 
+  contract ps ps' w -> OnlyWeber ps w -> OnlyWeber ps' w. 
+Proof.
+intros Hcontract [Hweb HwebU]. split.
++ rewrite (weber_contract_strong Hcontract Hweb). intuition.
++ intros w' Hweb'. apply HwebU. 
+  rewrite (weber_contract_strong Hcontract Hweb) in Hweb'.
+  intuition.
+Qed.
+
+
+(* Compute the sum of the element of a list of vectors [l] *)
+Fixpoint list_sumVS {T} `{RealVectorSpace T} (l : list T) :=
+  match l with 
+  | nil => 0%VS 
+  | hd :: tl => (hd + list_sumVS tl)%VS 
+  end.
+
+Local Instance list_sumVS_compat {T} `{RealVectorSpace T} : 
+  Proper (PermutationA equiv ==> equiv) list_sumVS.
+Proof.
+intros l l' Hll'. elim Hll'.
++ now reflexivity.
++ intros x x' t t' Hxx' Htt' IH. cbn -[equiv]. now rewrite Hxx', IH.
++ intros x y t. cbn -[equiv].
+  now rewrite 2 RealVectorSpace.add_assoc, (RealVectorSpace.add_comm y x).
++ intros t t' t'' _ IH1 _ IH2. now rewrite IH1, IH2.
+Qed.
+
+Lemma list_sumVS_nil {T} `{RealVectorSpace T} : 
+  list_sumVS (nil : list T) = 0%VS.
+Proof. now unfold list_sumVS. Qed.
+
+Lemma list_sumVS_app {T} `{RealVectorSpace T} (xs ys : list T) : 
+  list_sumVS (xs ++ ys) == (list_sumVS xs + list_sumVS ys)%VS.
+Proof. 
+induction xs as [| x xs IH].
++ simpl. now rewrite add_origin_l.
++ simpl. rewrite IH, add_assoc. reflexivity.
+Qed.
+
+Lemma dist_convex (x a b : R2) t :
+  (0 <= t <= 1)%R -> 
+  (dist x (a + t * (b - a)) <= dist x a + t * (dist x b - dist x a))%R.
+Proof. 
+intros Ht. repeat rewrite norm_dist.
+assert (Hrw : (x - (a + t * (b - a)) = (1 - t) * (x - a) + t * (x - b))%VS).
+{ 
+  repeat rewrite mul_distr_add.
+  repeat rewrite mul_opp.
+  unfold Rminus. repeat rewrite <-add_morph.
+  repeat rewrite mul_1. repeat rewrite <-add_assoc. f_equal.
+  repeat rewrite opp_distr_add. repeat rewrite opp_opp.
+  repeat rewrite add_assoc. rewrite (add_comm _ (-a)). 
+  repeat rewrite <-add_assoc. f_equal.
+  rewrite (add_comm _ (t * a)%VS). repeat rewrite add_assoc. f_equal.
+  rewrite <-add_origin_r. rewrite (add_comm (- t * x)%VS), <-add_assoc.
+  f_equal.
+  + now rewrite minus_morph, opp_opp. 
+  + now rewrite add_comm, minus_morph, add_opp.
+}
+rewrite Hrw ; clear Hrw. rewrite triang_ineq.
+rewrite norm_mul, Rabs_right by pos_R.
+rewrite norm_mul, Rabs_right by pos_R.
+apply Req_le.
+
+unfold Rminus. rewrite Rmult_plus_distr_l, Rmult_plus_distr_r, Rmult_1_l.
+repeat rewrite Rplus_assoc. f_equal.
+rewrite Rplus_comm. f_equal.
+rewrite <-Ropp_mult_distr_r, <-Ropp_mult_distr_l.
+reflexivity.
+Qed.
+
+Lemma dist_sum_convex ps (a b : R2) t :
+  (0 <= t <= 1)%R -> 
+  (dist_sum ps (a + t * (b - a)) <= dist_sum ps a + t * (dist_sum ps b - dist_sum ps a))%R.
+Proof. 
+intros Ht. induction ps as [| p ps IH].
++ unfold dist_sum. simpl. lra.
++ unfold dist_sum. cbn [map list_sum].
+  transitivity ((dist a p + t * (dist b p - dist a p)) + 
+    (list_sum (map (dist a) ps) + t * (list_sum (map (dist b) ps) - list_sum (map (dist a) ps))))%R.
+  - apply Rplus_le_compat.
+    * repeat rewrite (dist_sym p). apply dist_convex. assumption.
+    * apply IH.
+  - unfold Rminus. repeat rewrite Rplus_assoc. apply Rplus_le_compat_l.
+    repeat rewrite Rmult_plus_distr_l.
+    repeat rewrite <-Rplus_assoc. rewrite (Rplus_comm _ (t * dist b p)%R).
+    repeat rewrite Rplus_assoc. apply Rplus_le_compat_l.
+    rewrite Ropp_plus_distr, Rmult_plus_distr_l.
+    repeat rewrite <-Rplus_assoc. apply Rplus_le_compat_r.
+    rewrite (Rplus_comm (t * - dist a p)%R). repeat rewrite Rplus_assoc.
+    apply Rplus_le_compat_l. rewrite Rplus_comm. reflexivity.
+Qed.
+
+Section WeberFirstOrder.
+
+Variables (w : R2) (ps : list R2).
+
+(* The points in ps that are not equal to w. *)
+Definition qs : list R2 := List.filter (fun x => if x =?= w then false else true) ps.
+
+(* A unitary vector from w to x. If x == w, then u x == 0%VS (lemma unitary_origin). *)
+Definition u (x : R2) : R2 := unitary (x - w).
+
+(* The multiplicity of a point in ps. *)
+Definition mult (x : R2) : nat := countA_occ equiv R2_EqDec x ps.
+
+(* The minimum distance from w to x \in qs. *)
+Definition dmin : R := MinRlist (List.map (dist w) qs).
+
+(* The maximum distance from w to x \in qs. *)
+Definition dmax : R := MaxRlist (List.map (dist w) qs).
+
+Lemma qs_not_w (x : R2) :
+  InA equiv x qs -> x <> w.
+Proof using Type.
+unfold qs. rewrite filter_InA.
++ destruct_match ; intuition.
++ intros a b H. now rewrite H.
+Qed.
+
+Lemma qs_dist_pos (x : R2) : 
+  InA equiv x qs -> (0 < dist w x)%R.
+Proof using Type. 
+unfold qs. rewrite filter_InA.
++ destruct_match.
+  - intuition.
+  - case (dist_nonneg w x).
+    * intuition.
+    * intros Hdist. symmetry in Hdist. rewrite dist_defined in Hdist. intuition.
++ intros a b H. now rewrite H.
+Qed.
+
+Lemma dmin_def (x : R2) :
+  InA equiv x qs -> (dmin <= dist w x)%R.
+Proof using Type. 
+intros Hin. unfold dmin. apply MinRlist_P1. apply in_map. now rewrite <-InA_Leibniz.
+Qed.
+
+Lemma dmin_pos : (0 < dmin)%R.
+Proof using Type.
+unfold dmin. apply MinRlist_P2. 
+intros y Hin. rewrite in_map_iff in Hin. destruct Hin as [x [Hx Hin]].
+rewrite <-Hx. apply qs_dist_pos. now rewrite InA_Leibniz.
+Qed.
+
+Lemma dmax_def (x : R2) :
+  InA equiv x qs -> (dist w x <= dmax)%R.
+Proof using Type. 
+intros Hin. unfold dmax. apply MaxRlist_P1. apply in_map. now rewrite <-InA_Leibniz.
+Qed.
+
+Lemma dmax_nonneg : (0 <= dmax)%R.
+Proof using Type. 
+unfold dmax, qs. induction ps as [ | p ps' IH].
++ simpl. reflexivity.
++ cbn -[dist]. case (p =?= w) as [Hp | Hp].
+  - apply IH.
+  - cbn -[dist]. destruct_match ; pos_R.
+Qed.  
+
+(* This is true even if x == w. *)
+Lemma ux_formula (x : R2) : (x - w = norm (x - w) * u x)%VS.
+Proof using Type. 
+case (x =?= w).
++ intro Heq. rewrite Heq, add_opp, norm_origin, mul_0. reflexivity.
++ intro Hneq. unfold u, unitary. rewrite mul_morph. rewrite Rinv_r.
+  - rewrite mul_1. reflexivity.
+  - rewrite norm_defined. rewrite R2sub_origin. tauto.
+Qed.
+
+
+Lemma one_plus_sqrt_deriv (c : R) : 
+  (-1 < c)%R -> derivable_pt_lim (fun x => sqrt (1 + x)%R) c (/ (2 * sqrt (1 + c)))%R.
+Proof using Type. 
+intros Hc. prove_derivable_pt_lim.
++ lra.
++ lra.
+Qed.
+
+Lemma one_plus_sqrt_deriv_twice (c : R) : 
+  (-1 < c)%R -> derivable_pt_lim (fun x => (/ (2 * sqrt (1 + x)))%R) c (- / (4 * Rpower (1 + c) (3/2)))%R.
+Proof using Type. 
+intros Hc. prove_derivable_pt_lim.
++ lra.
++ enough (0 < sqrt (1 + c))%R by lra.
+  apply sqrt_lt_R0. lra.
++ rewrite Rmult_0_l, Rplus_0_l, Rplus_0_l, Rmult_1_r.
+  rewrite Ropp_div. f_equal. 
+  unfold Rdiv, Rsqr. repeat rewrite Rinv_mult. repeat rewrite <-Rmult_assoc.
+  rewrite Rinv_r, Rmult_1_l by lra.
+  enough ((/ sqrt (1 + c) * / sqrt (1 + c) * / sqrt (1 + c))%R =
+    (/ Rpower (1 + c) (3 * / 2))%R) by lra.
+  repeat rewrite <-Rinv_mult. f_equal.
+  rewrite <-Rpower_sqrt by lra. repeat rewrite <-Rpower_plus. f_equal. lra.
+Qed.
+
+Lemma Rabs_le_iff (a b : R) : (-b <= a <= b)%R <-> (Rabs a <= b)%R.
+Proof using Type. unfold Rabs. case (Rcase_abs a) ; lra. Qed.
+
+(* This is an application of the mean value theorem. *)
+Lemma sqrt_bound_helper (x : R) :
+  (Rabs x <= 1 / 2 -> 
+    exists y z, 
+      Rabs y <= Rabs x /\
+      Rabs z <= Rabs y /\
+      sqrt (1 + x) = 1 + x / 2 - x * y / (4 * Rpower (1 + z) (3/2)))%R.
+Proof using Type.
+intros Hx_abs. rewrite <-Rabs_le_iff in Hx_abs.
+remember (fun x => sqrt (1 + x)) as f.
+remember (fun x => / (2 * sqrt (1 + x)))%R as f'.
+remember (fun x => - / (4 * Rpower (1 + x) (3/2)))%R as f''.
+case (Rtotal_order 0 x) as [Hx | [Hx | Hx]].
++ assert (Hmvt := MVT_cor2 f f' 0 x Hx). feed Hmvt.
+  { intros y Hy. rewrite Heqf, Heqf'. apply one_plus_sqrt_deriv. lra. }
+  destruct Hmvt as [y [Hmvt [Hy1 Hy2]]]. exists y.
+  rewrite Rminus_0_r, Heqf, Rplus_0_r, sqrt_1 in Hmvt. 
+  unfold Rminus in Hmvt. rewrite eq_sym_iff in Hmvt. apply Rplus_reg_r in Hmvt.
+
+  assert (Hmvt' := MVT_cor2 f' f'' 0 y Hy1). feed Hmvt'.
+  { intros z Hz. rewrite Heqf', Heqf''. apply one_plus_sqrt_deriv_twice. lra. } 
+  destruct Hmvt' as [z [Hmvt' [Hz1 Hz2]]]. exists z.
+  rewrite Rminus_0_r in Hmvt'.
+  unfold Rminus in Hmvt'. rewrite eq_sym_iff in Hmvt'. apply Rplus_reg_r in Hmvt'.
+
+  assert (Hx' : Rabs x = x).
+  { apply Rabs_right. lra. }
+  assert (Hy' : Rabs y = y).
+  { apply Rabs_right. lra. }
+  assert (Hz' : Rabs z = z).
+  { apply Rabs_right. lra. }
+  rewrite Hx', Hy', Hz'.
+
+  split ; [| split] ; [lra | lra | ].
+  rewrite <-Hmvt, Rplus_comm. unfold Rminus. rewrite Rplus_assoc. f_equal.
+  rewrite <-Hmvt', Rplus_comm, Rmult_plus_distr_r. f_equal.
+  - rewrite Heqf', Rplus_0_r, sqrt_1, Rmult_1_r. unfold Rdiv. rewrite Rmult_comm. reflexivity.
+  - unfold Rdiv. rewrite Ropp_mult_distr_r, Rmult_assoc. rewrite Rmult_comm. f_equal.
+    * nra.
+    * rewrite Heqf''. reflexivity.
++ rewrite <-Hx. exists 0%R. exists 0%R. split ; [| split].
+  - rewrite Rabs_R0. lra.
+  - rewrite Rabs_R0. lra.
+  - rewrite Rplus_0_r, sqrt_1. lra.
++ assert (Hmvt := MVT_cor2 f f' x 0 Hx). feed Hmvt.
+  { intros y Hy. rewrite Heqf, Heqf'. apply one_plus_sqrt_deriv. lra. }
+  destruct Hmvt as [y [Hmvt [Hy1 Hy2]]]. exists y.
+  rewrite Rminus_0_l in Hmvt. apply Ropp_eq_compat in Hmvt.
+  rewrite Ropp_minus_distr, Ropp_mult_distr_r, Ropp_involutive in Hmvt.
+  unfold Rminus in Hmvt. rewrite eq_sym_iff in Hmvt. apply Rplus_reg_r in Hmvt.
+  rewrite Heqf, Rplus_0_r, sqrt_1 in Hmvt.
+
+  assert (Hmvt' := MVT_cor2 f' f'' y 0 Hy2). feed Hmvt'.
+  { intros z Hz. rewrite Heqf', Heqf''. apply one_plus_sqrt_deriv_twice. lra. } 
+  destruct Hmvt' as [z [Hmvt' [Hz1 Hz2]]]. exists z.
+  rewrite Rminus_0_l in Hmvt'. apply Ropp_eq_compat in Hmvt'.
+  rewrite Ropp_minus_distr, Ropp_mult_distr_r, Ropp_involutive in Hmvt'.
+  unfold Rminus in Hmvt'. rewrite eq_sym_iff in Hmvt'. apply Rplus_reg_r in Hmvt'.
+
+  assert (Hx' : Rabs x = (-x)%R).
+  { apply Rabs_left. lra. }
+  assert (Hy' : Rabs y = (-y)%R).
+  { apply Rabs_left. lra. }
+  assert (Hz' : Rabs z = (-z)%R).
+  { apply Rabs_left. lra. }
+  rewrite Hx', Hy', Hz'.
+
+  split ; [| split] ; [lra | lra | ].
+  rewrite <-Hmvt, Rplus_comm. unfold Rminus.
+  repeat rewrite Rplus_assoc. f_equal.
+  rewrite <-Hmvt', Rplus_comm, Rmult_plus_distr_r. f_equal.
+  - rewrite Heqf', Rplus_0_r, sqrt_1, Rmult_1_r. unfold Rdiv. rewrite Rmult_comm. reflexivity.
+  - unfold Rdiv. rewrite Ropp_mult_distr_r. 
+    rewrite (Rmult_comm _ x), Rmult_assoc. f_equal.
+    rewrite (Rmult_comm _ y). f_equal.
+    rewrite Heqf''. reflexivity.
+Qed.
+ 
+Lemma sqrt_bound :
+  exists c, 
+    (0 < c /\ 
+    forall x, Rabs x <= 1 / 2 -> 
+      sqrt (1 + x) <= 1 + x / 2 + c * Rsqr x /\
+      sqrt (1 + x) >= 1 + x / 2 - c * Rsqr x)%R.
+Proof using Type.
+enough (H : exists c, (0 < c)%R /\ forall y, (Rabs y <= 1 / 2 -> Rabs (/ (4 * Rpower (1 + y) (3/2))) < c)%R).
+{ 
+  destruct H as [c [Hc H]]. exists c. split.
+  + exact Hc. 
+  + intros x Hx. 
+    destruct (sqrt_bound_helper _ Hx) as [y [z [Hy [Hz Hsqrt]]]].
+    assert (Hbound : (Rabs (x * y / (4 * Rpower (1 + z) (3 / 2))) <= c * x²)%R).
+    { 
+      unfold Rdiv. rewrite Rabs_mult, Rmult_comm. 
+      apply Rmult_le_compat ; [pos_R | pos_R | | ].
+      * apply Rlt_le. apply H. lra.
+      * rewrite Rabs_mult, R_sqr.Rsqr_abs. unfold Rsqr. 
+        apply Rmult_le_compat_l ; [pos_R | assumption].
+    }
+
+    split.
+    - rewrite Hsqrt. apply Rplus_le_compat_l.
+      transitivity (Rabs (- (x * y / (4 * Rpower (1 + z) (3 / 2)))))%R ; [apply Rle_abs |]. 
+      rewrite Rabs_Ropp. exact Hbound.
+    - rewrite Hsqrt. apply Rle_ge, Rplus_le_compat_l, Ropp_le_contravar.
+      transitivity (Rabs (x * y / (4 * Rpower (1 + z) (3 / 2))))%R ; [apply Rle_abs |].
+      exact Hbound.
+} 
+enough (H : exists c, (0 < c)%R /\ forall y, (Rabs y <= 1 / 2 -> Rpower (1 + y) (3/2) > c)%R).
+{
+  destruct H as [c [Hc H]]. exists (/ (4 * c))%R. 
+  split.
+  + pos_R.
+  + intros y Hy. specialize (H y Hy).
+    rewrite Rabs_inv, Rabs_mult, (Rabs_right 4) by lra.
+    apply Rinv_lt_contravar.
+    - pos_R.
+    - apply Rmult_lt_compat_l ; [lra |].
+      rewrite Rabs_right ; [exact H |].
+      pos_R.
+}
+unfold Rpower. remember (exp (3 / 2 * ln (1 / 4)))%R as c. exists c.
+split.
++ rewrite Heqc. pos_R.
++ intros y Hy. apply Rlt_gt. rewrite Heqc. 
+  apply exp_increasing, Rmult_lt_compat_l ; [lra |].
+  apply ln_increasing ; [lra |].
+  rewrite <-Rabs_le_iff in Hy. lra.
+Qed.
+
+Lemma dist_formula_1 (x eps : R2) :
+  (Rsqr (dist (w + eps) x) = 
+    Rsqr (dist w x) - 2 * dist w x * inner_product eps (u x) + Rsqr (norm eps))%R.
+Proof using Type.
+repeat rewrite norm_dist. rewrite R2plus_opp_assoc_comm, norm_add. f_equal.
+unfold Rminus. rewrite Ropp_mult_distr_l. f_equal.
+rewrite Ropp_mult_distr_r. rewrite Rmult_assoc. f_equal.
+rewrite inner_product_sym. rewrite <-inner_product_mul_r. f_equal.
+change ((w - x)%VS == (- norm (w - x) * u x)%VS). apply opp_reg.
+rewrite opp_distr_add, opp_opp, add_comm, ux_formula.
+rewrite minus_morph, opp_opp. f_equiv.
+rewrite <-norm_opp, opp_distr_add, opp_opp, add_comm.
+reflexivity.
+Qed.
+
+(* A consequence of dist_formula_1. *)
+Lemma dist_formula_2 (x eps : R2) : 
+  x =/= w ->
+  (dist (w + eps) x = dist w x * 
+    sqrt (1 - 2 * inner_product eps (u x) / dist w x + Rsqr (norm eps / dist w x)))%R.
+Proof using Type.
+intros Hin.
+assert (H : forall p, dist p x = sqrt (Rsqr (dist p x))).
+{ intros p. rewrite sqrt_Rsqr ; [reflexivity | apply dist_nonneg]. }
+rewrite (H w) at 1. rewrite <-sqrt_mult_alt ; [| now apply Rle_0_sqr].
+rewrite (H (w + eps)%VS). f_equal.
+rewrite Rmult_plus_distr_l, Rmult_minus_distr_l. rewrite dist_formula_1. f_equal.
++ f_equal.
+  - rewrite Rmult_1_r. reflexivity.
+  - unfold Rdiv. rewrite (Rmult_comm _ (Rinv _)), <-(Rmult_assoc (Rsqr _)).
+    unfold Rsqr. rewrite (Rmult_assoc (dist w x)), Rinv_r.
+    * lra.
+    * rewrite dist_defined. now auto.
++ unfold Rdiv. rewrite R_sqr.Rsqr_mult, R_sqr.Rsqr_inv'.
+  rewrite <-Rmult_assoc. rewrite Rinv_r_simpl_m ; [reflexivity |].
+  rewrite <-Rsqr_0, <-square_dist_equiv ; [| reflexivity].
+  rewrite dist_defined. now auto.
+Qed.
+
+Lemma dist_bound_helper : 
+  exists c, 
+    (0 < c)%R /\ 
+    forall x eps, 
+      InA equiv x qs -> 
+      (norm eps < c)%R -> 
+      (Rabs (-2 * inner_product eps (u x) / dist w x + Rsqr (norm eps / dist w x)) <= 1 / 2)%R.
+Proof using Type.
+exists (dmin / 8)%R. split.
++ generalize dmin_pos. pos_R.
++ intros x eps Hin Heps. rewrite Rabs_triang. 
+  assert (Hrw : (1 / 2 = 1 / 4 + 1 / 4)%R) by lra. rewrite Hrw ; clear Hrw.
+  apply Rplus_le_compat.
+  - unfold Rdiv at 1. rewrite Rmult_assoc, Rabs_mult, (Rabs_left (-2)%R) by lra.
+    assert (Hrw : Ropp (-2)%R = 2%R) by lra. rewrite Hrw ; clear Hrw.
+    rewrite Rabs_mult, Rabs_inv, (Rabs_right (dist w x)) ; [| now apply Rle_ge, dist_nonneg].
+    enough (Rabs〈eps, u x〉%VS * / dist w x <= 1 / 8)%R by lra.
+    transitivity ((dist w x / 8) * / dist w x)%R.
+    * apply Rmult_le_compat_r. 
+      ++generalize (qs_dist_pos Hin). pos_R.
+      ++rewrite Cauchy_Schwarz. unfold u. rewrite norm_unitary, Rmult_1_r. 
+        --transitivity (dmin / 8)%R ; [apply Rlt_le ; exact Heps |].
+          unfold Rdiv. apply Rmult_le_compat_r ; [lra |].
+          now apply dmin_def.
+        --intros Heq. rewrite R2sub_origin in Heq. revert Heq. now apply qs_not_w.
+      * unfold Rdiv. rewrite Rmult_comm, <-Rmult_assoc, Rinv_l ; [reflexivity |].
+        rewrite dist_defined. symmetry. now apply qs_not_w.
+  - rewrite Rabs_right by pos_R.
+    assert (Hrw : (1 / 4 = Rsqr (1 / 2))%R) by (unfold Rsqr ; lra). rewrite Hrw ; clear Hrw.
+    apply R_sqr.Rsqr_incr_1.
+    * transitivity (1 / 8)%R ; [| lra].
+      transitivity (norm eps / dmin)%R.
+      { unfold Rdiv. apply Rmult_le_compat_l ; [pos_R |].
+        apply Rinv_le_contravar ; [apply dmin_pos | now apply dmin_def]. }
+      unfold Rdiv at 1.
+      transitivity ((dmin / 8) * / dmin)%R.
+      { apply Rmult_le_compat_r ; [generalize dmin_pos ; pos_R | lra]. }
+      unfold Rdiv. rewrite Rmult_comm, <-Rmult_assoc, Rinv_l ; [reflexivity |].
+      generalize dmin_pos ; lra.    
+    * generalize (qs_dist_pos Hin). pos_R.
+    * pos_R.
+Qed.
+  
+Lemma dist_bound_helper_2 : 
+  exists c d, 
+    (0 < c)%R /\
+    (0 < d)%R /\ 
+    forall x eps, 
+      InA equiv x qs -> 
+      (norm eps < c)%R -> 
+      (Rsqr (-2 * inner_product eps (u x) / dist w x + Rsqr (norm eps / dist w x)) <= d * Rsqr (norm eps))%R.
+Proof using Type. 
+enough (H : 
+  exists c d, 
+    (0 < c)%R /\
+    (0 < d)%R /\ 
+    forall x eps, 
+      InA equiv x qs -> 
+      (norm eps < c)%R -> 
+      (Rabs (-2 * inner_product eps (u x) / dist w x + Rsqr (norm eps / dist w x)) <= d * norm eps)%R).
+{
+  destruct H as [c [d [Hc [Hd H]]]]. exists c. exists (Rsqr d).
+  split ; [| split].
+  + assumption.
+  + pos_R. 
+  + intros x eps Hin Heps. specialize (H x eps Hin Heps).
+    remember (-2 * 〈eps, u x〉%VS / dist w x + (norm eps / dist w x)²)%R as X.
+    rewrite <-R_sqr.Rsqr_mult. apply R_sqr.Rsqr_le_abs_1.
+    rewrite (Rabs_right (d * _)%R).
+    - assumption.
+    - pos_R.  
+}
+exists 1%R. exists (2 / dmin + / Rsqr dmin)%R.
+split ; [| split].
++ lra.
++ generalize dmin_pos ; pos_R.
++ intros x eps Hin Heps. rewrite Rabs_triang, Rmult_plus_distr_r.
+  apply Rplus_le_compat.
+  - unfold Rdiv. repeat rewrite Rabs_mult.
+    assert (Hrw : Rabs (-2)%R = 2%R) by (rewrite Rabs_left ; lra).
+    rewrite Hrw. clear Hrw.
+    repeat rewrite Rmult_assoc. apply Rmult_le_compat_l ; [lra |].
+    rewrite Rmult_comm. apply Rmult_le_compat.
+    * pos_R. 
+    * pos_R.
+    * rewrite Rabs_right by (generalize (qs_dist_pos Hin) ; pos_R).
+      apply Rinv_le_contravar ; [now apply dmin_pos | now apply dmin_def].
+    * rewrite Cauchy_Schwarz. unfold u. rewrite norm_unitary ; [lra |].
+      intros Heq. rewrite R2sub_origin in Heq. revert Heq. now apply qs_not_w.
+  - rewrite Rabs_right by pos_R.
+    unfold Rsqr at 1. unfold Rdiv. rewrite Rmult_assoc, <-(Rmult_1_l (/ dmin² * norm eps)%R).
+    apply Rmult_le_compat. 
+    * pos_R. 
+    * generalize (qs_dist_pos Hin). pos_R.
+    * lra.
+    * rewrite (Rmult_comm (norm _)), <-Rmult_assoc. 
+      apply Rmult_le_compat_r ; [pos_R |].
+      unfold Rsqr. rewrite Rinv_mult. apply Rmult_le_compat.
+      ++generalize (qs_dist_pos Hin) ; pos_R.
+      ++generalize (qs_dist_pos Hin) ; pos_R.
+      ++apply Rinv_le_contravar ; [apply dmin_pos | now apply dmin_def].
+      ++apply Rinv_le_contravar ; [apply dmin_pos | now apply dmin_def].
+Qed.
+
+(* Uses sqrt_bound. *)
+Lemma dist_bound : 
+  exists c d : R, 
+    (0 < c)%R /\ 
+    (0 < d)%R /\
+    forall (x eps : R2), 
+      InA equiv x qs ->
+      (norm eps < c)%R ->
+      (dist (w + eps) x <= dist w x - inner_product eps (u x) + d * Rsqr (norm eps))%R /\
+      (dist (w + eps) x >= dist w x - inner_product eps (u x) - d * Rsqr (norm eps))%R.
+Proof using Type.  
+destruct dist_bound_helper as [c1 [Hc1 H1]].
+destruct dist_bound_helper_2 as [c2 [d2 [Hc2 [Hd2 H2]]]].
+destruct sqrt_bound as [d3 [Hd3 Hsqrt]].
+  
+exists (Rmin c1 c2). exists (/ (2 * dmin) + d2 * d3 * dmax)%R. split ; [| split].
++ pos_R.
++ generalize dmin_pos, dmax_nonneg. pos_R.
++ intros x eps Hqs Heps. 
+  specialize (H1 x eps Hqs). feed H1. { eapply Rlt_le_trans ; [exact Heps | apply Rmin_l]. }
+  specialize (H2 x eps Hqs). feed H2. { eapply Rlt_le_trans ; [exact Heps | apply Rmin_r]. }
+  remember (-2 * 〈eps, u x〉%VS / dist w x + (norm eps / dist w x)²)%R as delta.
+  rewrite dist_formula_2 ; [| now apply qs_not_w].
+  assert (Hrw : (1 - 2 * 〈eps, u x〉%VS / dist w x + (norm eps / dist w x)² = 1 + delta)%R).
+  { unfold Rminus at 1. rewrite Rplus_assoc, Heqdelta. f_equal. f_equal.
+    rewrite <-Ropp_div, Ropp_mult_distr_l. f_equal. }
+  rewrite Hrw. clear Hrw.
+
+  specialize (Hsqrt delta H1). destruct Hsqrt as [Hsqrt_upper Hsqrt_lower].
+
+  split.
+  - transitivity (dist w x * (1 + delta / 2 + d3 * Rsqr delta))%R.
+    { 
+      apply Rmult_le_compat_l ; [apply dist_nonneg | exact Hsqrt_upper]. 
+    }
+    transitivity (dist w x * (1 + delta / 2 + d2 * d3 * Rsqr (norm eps)))%R.
+    { 
+      apply Rmult_le_compat_l ; [apply dist_nonneg |].
+      apply Rplus_le_compat_l. rewrite (Rmult_comm d2 d3), Rmult_assoc.
+      apply Rmult_le_compat_l ; [lra | exact H2]. 
+    }
+    assert (Hrw : (dist w x * (1 + delta / 2 + d2 * d3 * (norm eps)²) =
+      dist w x - 〈eps, u x〉%VS + (/ (2 * dist w x) + d2 * d3 * dist w x) * (norm eps)²)%R). 
+    { 
+      do 2 rewrite Rmult_plus_distr_l. rewrite Rmult_1_r.
+      unfold Rminus. rewrite Rplus_assoc, Rplus_assoc. f_equal.
+      rewrite Rmult_plus_distr_r, <-Rplus_assoc. f_equal.
+      + rewrite Heqdelta. unfold Rdiv. rewrite Rmult_plus_distr_r, Rmult_plus_distr_l.
+        f_equal.
+        - rewrite Rmult_assoc, Rmult_comm, Rmult_assoc.
+          rewrite (Rmult_comm _ (/ 2)%R), (Rmult_assoc (/ 2)%R).
+          rewrite Rinv_l ; [lra |].
+          rewrite dist_defined. symmetry. now apply qs_not_w.
+        - rewrite <-Rmult_assoc, (Rmult_comm _ (/ 2)%R), Rinv_mult, Rmult_assoc.
+          apply Rmult_eq_compat_l. rewrite R_sqr.Rsqr_mult. 
+          rewrite (Rmult_comm (Rsqr (norm _))), <-Rmult_assoc. 
+          apply Rmult_eq_compat_r. unfold Rsqr. 
+          rewrite <-Rmult_assoc, Rinv_r, Rmult_1_l ; [reflexivity|].
+          rewrite dist_defined. symmetry. now apply qs_not_w.
+      + lra. 
+    }
+    rewrite Hrw. clear Hrw.
+    
+    unfold Rminus. apply Rplus_le_compat_l. 
+    apply Rmult_le_compat_r ; [apply Rle_0_sqr|]. apply Rplus_le_compat.
+    * apply Rinv_le_contravar. 
+      ++generalize dmin_pos ; pos_R.
+      ++apply Rmult_le_compat_l ; [lra | now apply dmin_def].
+    * apply Rmult_le_compat_l ; [pos_R |].
+      now apply dmax_def.
+  - apply Rle_ge. 
+    transitivity (dist w x * (1 + delta / 2 - d3 * delta²))%R ; 
+    [transitivity (dist w x * (1 + delta / 2 - d2 * d3 * Rsqr (norm eps)))%R |].
+    * assert (Hrw : (dist w x * (1 + delta / 2 - d2 * d3 * (norm eps)²) =
+        dist w x - 〈eps, u x〉%VS + (/ (2 * dist w x) - d2 * d3 * dist w x) * (norm eps)²)%R). 
+      {
+        unfold Rminus. do 2 rewrite Rmult_plus_distr_l. rewrite Rmult_1_r.
+        repeat rewrite Rplus_assoc. f_equal.
+        rewrite Rmult_plus_distr_r, <-Rplus_assoc. f_equal.
+        + rewrite Heqdelta. unfold Rdiv. rewrite Rmult_plus_distr_r, Rmult_plus_distr_l.
+          f_equal.
+          - rewrite Rmult_assoc, Rmult_comm, Rmult_assoc.
+            rewrite (Rmult_comm _ (/ 2)%R), (Rmult_assoc (/ 2)%R).
+            rewrite Rinv_l ; [lra |].
+            rewrite dist_defined. symmetry. now apply qs_not_w.
+          - rewrite <-Rmult_assoc, (Rmult_comm _ (/ 2)%R), Rinv_mult, Rmult_assoc.
+            apply Rmult_eq_compat_l. rewrite R_sqr.Rsqr_mult. 
+            rewrite (Rmult_comm (Rsqr (norm _))), <-Rmult_assoc. 
+            apply Rmult_eq_compat_r. unfold Rsqr. 
+            rewrite <-Rmult_assoc, Rinv_r, Rmult_1_l ; [reflexivity|].
+            rewrite dist_defined. symmetry. now apply qs_not_w.
+        + lra.
+      }
+      rewrite Hrw. clear Hrw.
+
+      unfold Rminus. apply Rplus_le_compat_l.
+      rewrite Rmult_plus_distr_r, Ropp_plus_distr, Rmult_plus_distr_r.
+      apply Rplus_le_compat.
+      ++transitivity 0%R.
+        --rewrite <-Ropp_0. apply Ropp_le_contravar.
+          generalize dmin_pos ; pos_R.
+        --generalize (qs_dist_pos Hqs) ; pos_R.
+      ++rewrite <-Ropp_mult_distr_l. apply Ropp_le_contravar.
+        repeat rewrite Rmult_assoc. 
+        apply Rmult_le_compat_l ; [pos_R |].
+        apply Rmult_le_compat_l ; [pos_R |].
+        apply Rmult_le_compat_r ; [pos_R |].
+        now apply dmax_def.
+    * apply Rmult_le_compat_l ; [pos_R |].
+      apply Rplus_le_compat_l. apply Ropp_le_contravar.
+      rewrite (Rmult_comm d2 d3), Rmult_assoc. apply Rmult_le_compat_l ; [lra |].
+      apply H2. 
+    * apply Rmult_le_compat_l ; [pos_R |]. 
+      apply Rge_le, Hsqrt_lower.
+Qed.
+
+
+(* The derivative of dist_sum along the vector eps. *)
+Definition deriv (eps : R2) : R := 
+  (INR (mult w) * norm eps - inner_product eps (list_sumVS (List.map u ps)))%R.
+
+Existing Instance PermutationA_map.
+
+Lemma dist_sum_bound : 
+  exists c d : R, 
+  (0 < c)%R /\ 
+  (0 < d)%R /\
+  forall (eps : R2), 
+    (norm eps < c)%R ->
+    (dist_sum ps (w + eps) <= dist_sum ps w + deriv eps + d * Rsqr (norm eps))%R /\
+    (dist_sum ps (w + eps) >= dist_sum ps w + deriv eps - d * Rsqr (norm eps))%R.
+Proof using Type. 
+destruct dist_bound as [c [d [Hc [Hd H]]]].
+exists c. exists ((1 + INR (length qs)) * d)%R. 
+split ; [assumption | split ; [pos_R |]].
+intros eps Heps.
+
+remember (List.filter (fun x : R2 => if x =?= w then true else false) ps) as rs.
+assert (Happ : PermutationA equiv ps (qs ++ rs)).
+{ 
+  unfold qs. rewrite Heqrs. clear Heqrs H.
+  induction ps as [| p ps' IH].
+  + simpl. constructor.
+  + simpl. case (p =?= w) as [Heq | HNeq].
+    - rewrite IH at 1. apply PermutationA_middle. intuition.
+    - rewrite IH at 1. cbn [app]. reflexivity. 
+}
+assert (Hdist : (dist_sum ps w = dist_sum qs w)%R).
+{ 
+  rewrite Happ. rewrite <-(Rplus_0_r (dist_sum qs w)). unfold dist_sum. rewrite map_app, list_sum_app.
+  apply Rplus_eq_compat_l. rewrite Heqrs. clear dependent rs.
+  induction ps as [ | p ps' IH].
+  + simpl. reflexivity.
+  + cbn -[dist]. case (p =?= w) as [Hp | Hp].
+    - cbn [map list_sum]. rewrite Hp, dist_same, Rplus_0_l. exact IH.
+    - exact IH.  
+}
+assert (Hdist' : (dist_sum ps (w + eps) = dist_sum qs (w + eps) + INR (mult w) * norm eps)%R).
+{ 
+  rewrite Happ. unfold dist_sum. rewrite map_app, list_sum_app. apply Rplus_eq_compat_l.
+  rewrite Heqrs. clear dependent rs. clear Hdist. unfold mult.
+  induction ps as [ | p ps' IH].
+  + simpl. lra.
+  + cbn [filter countA_occ]. 
+    destruct (p =?= w) as [Heq | HNeq] ; destruct (R2_EqDec p w) as [Htmp | Htmp] 
+    ; intuition ; clear Htmp.
+    - rewrite Heq. cbn [map list_sum]. rewrite IH.
+      rewrite S_INR, Rmult_plus_distr_r, Rmult_1_l, Rplus_comm.
+      apply Rplus_eq_compat_l. rewrite norm_dist. f_equal.
+      rewrite (add_comm w), <-add_assoc, add_opp, add_origin_r. 
+      reflexivity.
+}
+
+assert (Hsum_rs : list_sumVS (map u rs) = 0%VS).
+{ 
+  rewrite Heqrs. clear dependent rs. clear Hdist Hdist'.
+  induction ps as [| p ps' IH].
+  + cbn [filter map list_sumVS]. reflexivity.
+  + cbn [filter]. case (p =?= w) as [Heq | HNeq].
+    - cbn [map list_sumVS]. rewrite IH. unfold u.
+      rewrite Heq, add_opp, unitary_origin, add_origin_l. reflexivity.
+    - rewrite IH. reflexivity.
+}
+
+rewrite Hdist, Hdist'. clear Hdist Hdist'. split.
++ enough (Henough : (dist_sum qs (w + eps) <= dist_sum qs w 
+    - inner_product eps (list_sumVS (map u ps)) 
+    + INR (length qs) * d * Rsqr (norm eps))%R).
+  {
+    unfold deriv. rewrite Rplus_comm. unfold Rminus. rewrite <-Rplus_assoc.
+    rewrite (Rplus_comm (dist_sum qs w)). repeat rewrite Rplus_assoc.
+    apply Rplus_le_compat_l. rewrite Henough. unfold Rminus.
+    repeat rewrite Rplus_assoc. apply Rplus_le_compat_l. apply Rplus_le_compat_l.
+    apply Rmult_le_compat_r ; [pos_R |].
+    apply Rmult_le_compat_r ; [pos_R |].
+    lra.
+  }
+  rewrite Happ, map_app, list_sumVS_app. rewrite Hsum_rs, add_origin_r.
+  clear dependent rs.
+  induction qs as [| q qs IH] ; unfold dist_sum.
+  - cbn [list_sum map length INR list_sumVS]. 
+    rewrite Rmult_0_l, Rmult_0_l, inner_product_origin_r. lra.
+  - cbn [map list_sum length list_sumVS]. 
+    rewrite S_INR, Rmult_plus_distr_r, Rmult_1_l, Rmult_plus_distr_r.
+    rewrite inner_product_add_r. unfold Rminus.
+    enough (dist (w + eps) q <= dist w q - inner_product eps (u q) + d * Rsqr (norm eps) /\
+            list_sum (map (dist (w + eps)) qs) <= list_sum (map (dist w) qs) 
+            - inner_product eps (list_sumVS (map u qs)) + INR (length qs) * d * Rsqr (norm eps))%R by lra.
+    split. 
+    * apply H ; [now apply InA_cons_hd | assumption].
+    * apply IH. intros x' eps' Hin'. apply H. apply InA_cons_tl. assumption.
++ enough (Henough : (dist_sum qs (w + eps) >= dist_sum qs w 
+    - inner_product eps (list_sumVS (map u ps)) 
+    - INR (length qs) * d * Rsqr (norm eps))%R).
+  {
+    unfold deriv. rewrite Rplus_comm. unfold Rminus. rewrite <-Rplus_assoc.
+    rewrite (Rplus_comm (dist_sum qs w)). repeat rewrite Rplus_assoc.
+    apply Rplus_ge_compat_l. apply Rle_ge. apply Rge_le in Henough.
+    transitivity (dist_sum qs w - 〈eps, list_sumVS (map u ps)〉%VS
+      - INR (length qs) * d * (norm eps)²)%R ; [| apply Henough].
+    unfold Rminus.
+    repeat rewrite Rplus_assoc. apply Rplus_le_compat_l. apply Rplus_le_compat_l.
+    apply Ropp_le_contravar.
+    apply Rmult_le_compat_r ; [pos_R |].
+    apply Rmult_le_compat_r ; [pos_R |].
+    lra.
+  }
+  rewrite Happ, map_app, list_sumVS_app. rewrite Hsum_rs, add_origin_r.
+  clear dependent rs.
+  induction qs as [| q qs IH] ; unfold dist_sum.
+  - cbn [list_sum map length INR list_sumVS]. 
+    rewrite Rmult_0_l, Rmult_0_l, inner_product_origin_r. lra.
+  - cbn [map list_sum length list_sumVS]. 
+    rewrite S_INR, Rmult_plus_distr_r, Rmult_1_l, Rmult_plus_distr_r.
+    rewrite inner_product_add_r. unfold Rminus.
+    enough (dist (w + eps) q >= dist w q - inner_product eps (u q) - d * Rsqr (norm eps) /\
+            list_sum (map (dist (w + eps)) qs) >= list_sum (map (dist w) qs) 
+            - inner_product eps (list_sumVS (map u qs)) - INR (length qs) * d * Rsqr (norm eps))%R by lra.
+    split. 
+    * apply H ; [now apply InA_cons_hd | assumption].
+    * apply IH. intros x' eps' Hin'. apply H. apply InA_cons_tl. assumption.
+Qed.
+
+Lemma deriv_scale (eps : R2) (s : R) : 
+  (0 <= s)%R ->
+  (deriv (s * eps) = s * deriv eps)%R.
+Proof using Type. 
+intros Hs. unfold deriv. 
+rewrite norm_mul, Rabs_right by pos_R.
+rewrite inner_product_mul_l. lra.
+Qed.
+
+Lemma deriv_0 : (deriv 0 = 0)%R.
+Proof using Type.
+unfold deriv. rewrite inner_product_origin_l, norm_origin. lra.
+Qed. 
+
+Lemma deriv_neg_strengthen_1 c : 
+  (0 < c)%R -> 
+  (exists eps, deriv eps < 0)%R ->
+  (exists eps', deriv eps' < 0 /\ norm eps' < c)%R.
+Proof using Type.
+intros Hc [eps Heps].
+assert (Heps_pos : (0 < norm eps)%R).
+{ 
+  case (norm_nonneg eps) as [Hpos | Heq]. 
+  + assumption.
+  + symmetry in Heq. rewrite norm_defined in Heq.
+    rewrite Heq, deriv_0 in Heps. lra.
+}
+
+set (s := (c / (2 * norm eps))%R).
+assert (0 < s)%R by (unfold s ; pos_R).
+
+exists (s * eps)%VS. split.
++ rewrite deriv_scale by pos_R. nra.
++ rewrite norm_mul, Rabs_right by pos_R. 
+  unfold s, Rdiv. rewrite Rinv_mult. repeat rewrite Rmult_assoc.
+  rewrite Rinv_l by lra. lra.
+Qed. 
+
+Lemma deriv_neg_strengthen_2 c d :
+  (0 < c)%R ->
+  (0 < d)%R ->
+  (exists eps, deriv eps < 0)%R -> 
+  (exists eps',
+    deriv eps' < 0 /\
+    norm eps' < c /\
+    deriv eps' + d * Rsqr (norm eps') < 0)%R.
+Proof using Type. 
+intros Hc Hd Heps. apply (@deriv_neg_strengthen_1 (Rmin 1 c)%R) in Heps ; [| pos_R].
+destruct Heps as [eps [Hderiv Hnorm]].
+assert (Heps_pos : (0 < norm eps)%R).
+{ 
+  case (norm_nonneg eps) as [Hpos | Heq]. 
+  + assumption.
+  + symmetry in Heq. rewrite norm_defined in Heq.
+    rewrite Heq, deriv_0 in Hderiv. lra.
+} 
+
+set (s := (-deriv eps / (2 * d * Rsqr (norm eps)))%R).
+assert (Hs_pos : (0 < s)%R) by (unfold s ; pos_R).
+
+set (eps' := (Rmin s 1 * eps)%VS).
+exists eps'. split ; [| split].
++ unfold eps'. rewrite deriv_scale by pos_R.
+  enough (0 < Rmin s 1)%R by nra.
+  pos_R.
++ unfold eps'. rewrite norm_mul. rewrite Rabs_right by pos_R.
+  rewrite <-(Rmult_1_l c).
+  enough (0 < Rmin s 1 <= 1 /\ 0 < norm eps < c)%R by nra.
+  split ; split.
+  - pos_R.
+  - apply Rmin_r.
+  - assumption.
+  - eapply Rlt_le_trans ; [apply Hnorm | apply Rmin_r].
++ enough (d * Rsqr (norm eps') < -deriv eps')%R by lra.
+  unfold eps'. rewrite deriv_scale by pos_R.
+  rewrite norm_mul, Rabs_right by pos_R.
+  rewrite R_sqr.Rsqr_mult. unfold Rsqr at 1. repeat rewrite <-Rmult_assoc.
+  rewrite (Rmult_comm d). repeat rewrite Rmult_assoc. rewrite Ropp_mult_distr_r.
+  apply Rmult_lt_compat_l ; [pos_R |].
+  rewrite <-Rmult_assoc, (Rmult_comm d), Rmult_assoc. 
+  apply (@Rlt_le_trans _ ((- deriv eps / (d * (norm eps)²)) * (d * Rsqr (norm eps)))%R).
+  - apply Rmult_lt_compat_r ; [pos_R |].
+    eapply Rle_lt_trans ; [apply Rmin_l |].
+    enough (Hrw : (- deriv eps / (d * (norm eps)²) = 2 * s)%R).
+    { rewrite Hrw. lra. }
+    unfold s, Rdiv. repeat rewrite Rinv_mult. repeat rewrite <-Rmult_assoc.
+    rewrite (Rmult_comm 2). repeat rewrite Rmult_assoc. f_equal.
+    repeat rewrite <-Rmult_assoc. rewrite Rinv_r by lra. lra.
+  - unfold Rdiv. rewrite Rmult_assoc, Rinv_l.
+    * lra.
+    * enough (0 < d * Rsqr (norm eps))%R by lra. pos_R. 
+Qed.
+
+(* If we can find a direction eps such that deriv eps < 0, 
+ * then w is not a weber point. *)
+Lemma deriv_neg_not_weber : 
+  (exists eps, deriv eps < 0)%R -> ~Weber ps w.
+Proof using Type.
+destruct dist_sum_bound as [c [d [Hc [Hd Hbound]]]]. intros Heps.
+destruct (deriv_neg_strengthen_2 Hc Hd Heps) as [eps [Hderiv [Hnorm Hsum]]].
+assert (Hdist : (dist_sum ps (w + eps) < dist_sum ps w)%R).
+{  
+  destruct (Hbound eps Hnorm) as [Hbound' _].
+  eapply Rle_lt_trans ; [exact Hbound' |].
+  rewrite <-(Rplus_0_r (dist_sum ps w)) at 2. repeat rewrite Rplus_assoc.
+  apply Rplus_lt_compat_l. assumption.
+}
+unfold Weber, argmin. apply Classical_Pred_Type.ex_not_not_all.
+exists (w + eps)%VS. lra.
+Qed.
+
+Lemma classical_contra (P Q : Prop) : (~Q -> ~P) -> (P -> Q).
+Proof using Type. tauto. Qed. 
+
+Lemma not_weber_deriv_neg :
+  ~Weber ps w -> exists eps, (deriv eps < 0)%R.  
+Proof using Type.
+intros Hweber. unfold Weber, argmin in Hweber. apply Classical_Pred_Type.not_all_ex_not in Hweber.
+destruct Hweber as [r Hr]. apply Rnot_le_gt, Rgt_lt in Hr.
+assert (Hdist_wr : (0 < dist w r)%R).
+{
+  case (dist_nonneg w r) as [Hpos | Hzero] ; [tauto |].
+  symmetry in Hzero. rewrite dist_defined in Hzero. rewrite Hzero in Hr. lra.
+}
+
+destruct dist_sum_bound as [c [d [Hc [Hd Hbound]]]]. 
+set (t_max := Rmin 1 (c / dist w r)%R).
+assert (Ht_max_pos : (0 < t_max)%R) by (unfold t_max ; pos_R).
+
+assert (Hderiv : forall t, (0 < t < t_max)%R -> 
+  (deriv (r - w) <= d * t * (dist w r)² + dist_sum ps r - dist_sum ps w)%R).
+{
+  intros t [Ht1 Ht2].
+  specialize (Hbound (t * (r - w))%VS). feed Hbound.
+  { 
+    rewrite norm_mul, Rabs_right, <-norm_dist, dist_sym by pos_R.
+    rewrite <-(Rmult_1_r c), <-(Rinv_l (dist w r)), <-Rmult_assoc by lra.
+    apply Rmult_lt_compat_r ; [assumption |].
+    eapply Rlt_le_trans ; [apply Ht2 |].
+    unfold t_max. rewrite Rmin_r. reflexivity.
+  }
+  destruct Hbound as [_ Hbound]. apply Rge_le in Hbound.
+  rewrite dist_sum_convex in Hbound ; [| split].
+  - unfold Rminus in Hbound.
+    rewrite Rplus_assoc in Hbound. apply Rplus_le_reg_l in Hbound.
+    assert (Hrw : (deriv (t * (r - w)) + - (d * (norm (t * (r - w)))²) = 
+      t * (deriv (r - w)%VS - d * t * (dist w r)²))%R).
+    {
+      unfold Rminus. rewrite Rmult_plus_distr_l. f_equal.
+      + rewrite deriv_scale by pos_R. reflexivity.
+      + rewrite norm_mul, <-norm_dist, dist_sym, Rabs_right by pos_R.
+        rewrite <-Ropp_mult_distr_r. f_equal.
+        rewrite R_sqr.Rsqr_mult. unfold Rsqr at 1.
+        repeat rewrite <-Rmult_assoc. f_equal. f_equal.
+        now rewrite Rmult_comm.
+    }
+    rewrite Hrw in Hbound ; clear Hrw.
+    apply Rmult_le_reg_l in Hbound ; [| assumption].
+    lra.
+  - lra.
+  - transitivity t_max.
+    * lra.
+    * unfold t_max. rewrite Rmin_l. reflexivity.   
+}
+
+exists (r - w)%VS. 
+set (t := (Rmin (t_max * / 2) ((dist_sum ps w - dist_sum ps r) * / (d * (dist w r)²) * / 2))%R).
+assert (Ht_pos : (0 < t)%R) by (unfold t ; pos_R). 
+
+assert (Ht1 : (t < t_max)%R).
+{ 
+  unfold t. eapply Rle_lt_trans ; [apply Rmin_l |].
+  lra.
+} 
+
+assert (Ht2 : (d * t * (dist w r)² < dist_sum ps w - dist_sum ps r)%R).
+{ 
+  rewrite (Rmult_comm d t), Rmult_assoc.
+  apply (Rle_lt_trans _ (((dist_sum ps w - dist_sum ps r) * / (d * (dist w r)²) *
+  / 2) * (d * (dist w r)²))%R).
+  + apply Rmult_le_compat_r ; [pos_R |].
+    unfold t. rewrite Rmin_r. reflexivity.
+  + rewrite <-(Rmult_1_r). repeat rewrite Rmult_assoc. apply Rmult_lt_compat_l ; [lra |].
+    rewrite Rmult_comm, Rmult_assoc, Rinv_r. 
+    - lra.
+    - enough (0 < d * (dist w r)²)%R by lra. pos_R.
+}
+
+feed (Hderiv t). { lra. }
+eapply Rle_lt_trans ; [apply Hderiv |]. lra.
+Qed.
+
+(* This is a consequence of the cauchy-schwartz inequality. *)
+Lemma deriv_neg_iff : 
+  (norm (list_sumVS (List.map u ps)) > INR (mult w))%R <-> (exists eps, deriv eps < 0)%R.
+Proof using Type. 
+split.
++ intros H. 
+  assert (H' : list_sumVS (map u ps) =/= 0%VS).
+  { intros Heq. rewrite Heq, norm_origin in H. enough (0 <= INR (mult w))%R by lra. pos_R. }
+
+  exists (unitary (list_sumVS (map u ps))). unfold deriv.
+  rewrite norm_unitary, Rmult_1_r by assumption.
+  rewrite (unitary_id (list_sumVS (map u ps))) at 2.
+  rewrite inner_product_mul_r, <-squared_norm_product, norm_unitary by assumption.
+  rewrite R_sqr.Rsqr_1. lra.
++ apply classical_contra. intros H.
+  apply Classical_Pred_Type.all_not_not_ex. intros eps.
+  apply Rnot_gt_le in H. apply Rge_not_lt, Rle_ge.
+  unfold deriv.
+  enough (inner_product eps (list_sumVS (map u ps)) <= INR (mult w) * norm eps)%R by lra.
+  rewrite Rle_abs, Cauchy_Schwarz, Rmult_comm. 
+  apply Rmult_le_compat_r ; [pos_R | assumption].
+Qed.
+
+(* This is the first order condition for weber points. *)
+Theorem weber_first_order : 
+  Weber ps w <-> (norm (list_sumVS (List.map u ps)) <= INR (countA_occ equiv R2_EqDec w ps))%R.
+Proof using Type.
+split.
++ apply classical_contra. intros H. apply deriv_neg_not_weber. 
+  rewrite <-deriv_neg_iff. apply Rnot_le_gt in H. assumption.
++ apply classical_contra. intros H. apply Rgt_not_le. 
+  rewrite deriv_neg_iff. apply not_weber_deriv_neg. assumption.
+Qed.
+
+Lemma u_aligned_left L x :
+  line_contains L w ->
+  line_contains L x -> 
+  (L^P x < L^P w)%R ->
+  (u x = -unitary (line_dir L))%VS.
+Proof using Type. 
+intros Hw Hx Hleft. unfold u. rewrite <-Hw, <-Hx. unfold line_proj_inv.
+rewrite opp_distr_add, (add_comm (line_org L)). rewrite add_assoc.
+rewrite <-(add_assoc _ (line_org L)), add_opp, add_origin_r. 
+rewrite <-minus_morph, add_morph. rewrite <-opp_opp. f_equal.
+rewrite <-unitary_opp, <-minus_morph. rewrite unitary_mul.
++ reflexivity.
++ rewrite <-Ropp_div, <-Rdiv_plus_distr, <-Ropp_div. 
+  enough (0 <> norm (line_dir L))%R by pos_R.
+  symmetry. rewrite norm_defined. 
+  intros Hnull. unfold line_proj in Hleft. 
+  rewrite Hnull, inner_product_origin_r, inner_product_origin_r in Hleft.
+  lra.
+Qed.
+
+Lemma u_aligned_right L x :
+  line_contains L w ->
+  line_contains L x -> 
+  (L^P x > L^P w)%R ->
+  (u x = unitary (line_dir L))%VS.
+Proof using Type. 
+intros Hw Hx Hright. unfold u. rewrite <-Hw, <-Hx. unfold line_proj_inv.
+rewrite opp_distr_add, (add_comm (line_org L)). rewrite add_assoc.
+rewrite <-(add_assoc _ (line_org L)), add_opp, add_origin_r. 
+rewrite <-minus_morph, add_morph. rewrite unitary_mul.
++ reflexivity.
++ rewrite <-Ropp_div, <-Rdiv_plus_distr. 
+  enough (0 <> norm (line_dir L))%R by pos_R.
+  symmetry. rewrite norm_defined. 
+  intros Hnull. unfold line_proj in Hright. 
+  rewrite Hnull, inner_product_origin_r, inner_product_origin_r in Hright.
+  lra.
+Qed.
+
+Lemma aligned_sum_us L : 
+  aligned_on L (w :: ps) -> 
+  (list_sumVS (map u ps) = (INR (length (L^right w ps)) - INR (length (L^left w ps))) * unitary (line_dir L))%VS.
+Proof using Type.
+intros Halign. rewrite aligned_on_cons_iff in Halign. destruct Halign as [Hw Halign].
+rewrite (line_left_on_right_partition L w ps) at 1.
+do 2 rewrite map_app, list_sumVS_app.
+induction ps as [| p ps' IH].
++ simpl. now simpl_R.
++ rewrite aligned_on_cons_iff in Halign. destruct Halign as [Hp Halign].
+  specialize (IH Halign).
+  unfold line_left, line_right, line_on. cbn [filter map list_sumVS].
+  case (Rtotal_order (L^P p) (L^P w)) as [Hpw | [Hpw | Hpw]].
+  - pose proof (Hrw := Hpw). rewrite <-Rltb_true in Hrw. rewrite Hrw ; clear Hrw.
+    assert (Hrw : ~(L^P w == L^P p)) by (change (L^P w <> L^P p) ; lra).
+    rewrite <-eqb_false_iff in Hrw. rewrite Hrw ; clear Hrw.
+    assert (Hrw : ~(L^P w < L^P p)%R) by lra. rewrite <-Rltb_false in Hrw.
+    rewrite Hrw ; clear Hrw. 
+    cbn [map list_sumVS length]. rewrite <-add_assoc.
+    unfold line_on, line_right, line_left in IH. rewrite IH. clear IH.
+    enough (Hu : (u p = (-1) * unitary (line_dir L))%VS).
+    { rewrite Hu. rewrite add_morph. f_equal. rewrite S_INR. lra. }
+    change (-1)%R with (Ropp 1%R). rewrite minus_morph, mul_1.
+    apply u_aligned_left ; assumption.
+  - assert (Hrw : L^P w == L^P p) by now simpl. rewrite <-eqb_true_iff in Hrw. 
+    rewrite Hrw ; clear Hrw.
+    assert (Hrw : ~(L^P w < L^P p)%R) by lra. rewrite <-Rltb_false in Hrw.
+    rewrite Hrw ; clear Hrw.
+    assert (Hrw : ~(L^P p < L^P w)%R) by lra. rewrite <-Rltb_false in Hrw.
+    rewrite Hrw ; clear Hrw. 
+    cbn [map list_sumVS length]. repeat rewrite add_assoc. rewrite (add_comm _ (u p)).
+    repeat rewrite <-add_assoc.
+    unfold line_on, line_right, line_left in IH. rewrite IH. clear IH.
+    enough (Heq : p == w).
+    { unfold u. rewrite Heq, add_opp, unitary_origin, add_origin_l. reflexivity. }
+    rewrite <-Hw, <-Hp, Hpw. reflexivity.
+  - assert (Hrw : ~(L^P w == L^P p)) by (change (L^P w <> L^P p) ; lra). rewrite <-eqb_false_iff in Hrw. 
+    rewrite Hrw ; clear Hrw.
+    assert (Hrw : (L^P w < L^P p)%R) by lra. rewrite <-Rltb_true in Hrw.
+    rewrite Hrw ; clear Hrw.
+    assert (Hrw : ~(L^P p < L^P w)%R) by lra. rewrite <-Rltb_false in Hrw.
+    rewrite Hrw ; clear Hrw. 
+    cbn [map list_sumVS length]. rewrite (add_comm (u p)).
+    repeat rewrite add_assoc. repeat rewrite add_assoc in IH.
+    unfold line_on, line_right, line_left in IH. rewrite IH. clear IH.
+    enough (Hu : (u p = 1 * unitary (line_dir L))%VS).
+    { rewrite Hu. rewrite add_morph. f_equal. rewrite S_INR. lra. }
+    rewrite mul_1. apply u_aligned_right ; assumption.
+Qed.
+
+End WeberFirstOrder.
+
+
+(* This is an application of weber_first_order. *)
+Lemma weber_majority_weak ps w : 
+  countA_occ equiv R2_EqDec w ps >= (Nat.div2 (length ps + 1)) -> Weber ps w.
+Proof.
+intros Hcount. rewrite ge_le_iff in Hcount. rewrite weber_first_order.
+assert (Hineq : (norm (list_sumVS (map (u w) ps)) + INR (countA_occ equiv R2_EqDec w ps) <= INR (length ps))%R).
+{
+  clear Hcount.
+  induction ps as [| p ps IH].
+  + simpl. rewrite Rmult_0_l, Rplus_0_l, sqrt_0. lra.
+  + cbn [countA_occ length map list_sumVS]. case (R2_EqDec p w) as [Hpw | Hpw].
+    - unfold u at 1. rewrite Hpw, add_opp, unitary_origin, add_origin_l.
+      rewrite S_INR, S_INR. lra.
+    - transitivity (1%R + norm (list_sumVS (map (u w) ps)) + INR (countA_occ equiv R2_EqDec w ps))%R.
+      * apply Rplus_le_compat_r. rewrite triang_ineq. apply Rplus_le_compat_r.
+        unfold u. rewrite norm_unitary ; [lra |]. 
+        intros H. rewrite R2sub_origin in H. intuition. 
+      * rewrite S_INR. lra.  
+}
+transitivity (INR (Nat.div2 (length ps + 1))) ; [| now apply le_INR].
+eapply Rplus_le_reg_r ; rewrite Hineq.
+rewrite <-(Rplus_le_compat_l (INR (Nat.div2 (length ps + 1))) _ _ (le_INR _ _ Hcount)).
+rewrite <-plus_INR. apply le_INR. apply div2_sum_p1_ge.
+Qed.
+
+(* This is the first order condition for weber points,
+ * in the special case of aligned points. *)
+Lemma weber_aligned_spec_weak L ps w : 
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+    Weber ps w <-> 
+    (Rabs (INR (length (L^left w ps)) - INR (length (L^right w ps))) <= INR (length (L^on w ps)))%R.
+Proof. 
+intros Hdir Halign.
+rewrite weber_first_order. rewrite (aligned_sum_us Halign).
+rewrite norm_mul, norm_unitary, Rmult_1_r by assumption.
+rewrite <-(line_on_length_aligned L _ _ Halign). 
+f_equiv. rewrite <-Rabs_Ropp. f_equiv. lra.
+Qed.
+
+(* A variant for the left to right part with a different hypothesis. *)
+Lemma weber_aligned_spec_weak_right L ps w : 
+  line_dir L =/= 0%VS ->
+  ps <> nil ->
+  aligned_on L ps -> 
+  Weber ps w -> 
+  (Rabs (INR (length (L^left w ps)) - INR (length (L^right w ps))) <= INR (length (L^on w ps)))%R.
+Proof. 
+intros Hdir Halign' Hnonnil HWeber. 
+apply weber_aligned_spec_weak;auto.
+apply aligned_on_cons_iff;split;auto.
+eapply weber_aligned;eauto.
+Qed.
+
+Lemma line_on_app: forall (L: line) (l l':list R2) x,
+    PermutationA equiv ((L ^on) x (l++l')) (((L ^on) x l) ++ ((L ^on) x l')) .
+Proof.
+  intros L l l' x. 
+  unfold line_on.
+  rewrite filter_app.
+  reflexivity.
+Qed.
+Lemma line_on_perm_compat: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T}
+                                  {H : EuclideanSpace T},
+    Proper (equiv ==> equiv ==> PermutationA equiv ==> PermutationA equiv) line_on.
+Proof.
+  repeat intro.
+  unfold line_on.
+  rewrite H2.
+  rewrite (filter_ext (fun y2 : T => (x ^P) x0 ==b (x ^P) y2) (fun y2 : T => (y ^P) y0 ==b (y ^P) y2)).
+  { reflexivity. }
+  intros a. 
+  rewrite H1.
+  rewrite H0.
+  reflexivity.
+Qed.
+
+(* Notation "L '^lftof' x '^in' ps" := (INR (Rlength (line_left L x ps))) (at level 100, no associativity). *)
+
+Notation "'^R' x" := (INR (length x)) (at level 180, no associativity).
+
+Lemma weber_aligned_sufficient L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) = ^R (L^on w ps))%R ->
+  ((^R (L^left w ps)) - (^R (L^right w ps)) < 0)%R ->
+  Weber ps (L^-P (L^min (L^right w ps))).
+Proof.
+  intros Hpsnonnil Hlinedir Halign H_eq_on H_right. 
+  remember (L^-P (L^min (L^right w ps))) as w'.
+  assert(InA equiv w' ((L ^right) w ps)) as h_InA.
+  { subst w'.
+    apply line_iP_min_InA.
+    - rewrite <- length_zero_iff_nil.
+
+      assert (0 <= (INR (length (line_left L w ps))))%R. 
+      { apply pos_INR. }
+      assert (INR (length (line_right L w ps)) > 0)%R as h_lgth.
+      { lra. }
+      intro abs.
+      rewrite  abs in h_lgth.
+      cbn in h_lgth.
+      lra.
+    - eapply aligned_on_inclA with (ps':=ps).
+      + apply line_right_inclA.
+      + rewrite aligned_on_cons_iff in Halign.
+        apply Halign. }
+  specialize (bipartition_min L ((L ^right) w ps)) as h.
+  apply PermutationA_length in h.
+  rewrite app_length in h.
+  (* apply (f_equal INR) in h. *)
+  (* rewrite plus_INR in h. *)
+  (* rewrite <- Heqw' in h. *)
+  rewrite <- Heqw' in h.
+
+  specialize (line_left_on_right_partition L w ps) as h_w.
+  specialize (line_left_on_right_partition L w' ps) as h_w'.
+  rewrite h_w' in h_w at 1.
+  clear h_w'.
+  apply PermutationA_length in h_w.
+  rewrite 4 app_length in h_w.
+  rewrite weber_aligned_spec_weak with (L:=L);auto.
+  2:{ rewrite aligned_on_cons_iff in *|-*.
+      split.
+      - subst w'.
+        apply line_contains_proj_inv.
+        assumption.
+      - apply Halign. }
+  
+  assert ((L ^P) w < (L ^P) w')%R.
+  { subst w'.
+    rewrite line_P_iP_min.
+    (* TODO: have a lemma *)
+    unfold line_right in h_InA at 2.
+    apply filter_InA in h_InA.
+    2:{ intros ? ? h_eq.
+        f_equal.
+        now rewrite h_eq. }
+    rewrite Rltb_true in h_InA.
+    rewrite line_P_iP_min in h_InA.            
+    apply h_InA. }
+  rewrite <- aggravate_right in h;auto.
+  rewrite <- aggravate_on_right in h;auto.
+  rewrite <- h in h_w.
+  apply (f_equal INR) in h, h_w.
+  rewrite plus_INR in h,h_w.
+  rewrite plus_INR in h_w.
+  rewrite plus_INR in h_w.
+  (* rewrite <- Heqw' in h. *)
+  assert (((^R (L ^left) w' ps) - (^R (L ^right) w' ps)) = (^R (L ^on) w' ps))%R.
+  { lra. }
+  assert (0 <= (^R (L ^on) w' ps))%R.
+  { apply pos_INR. }
+  rewrite Rabs_pos_eq;
+    lra.
+Qed.
+
+
+Lemma weber_aligned_single_weber_sufficient L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) = ^R (L^on w ps))%R ->
+  ((^R (L^left w ps)) - (^R (L^right w ps)) < 0)%R ->
+  ~ OnlyWeber ps w.
+Proof.
+  intros Hpsnonnil Hlinedir Halign H_eq_on H_right. 
+  intro abs.
+  remember (L^-P (L^min (L^right w ps))) as w'.
+  assert (Weber ps w') as Hweberw'.
+  { subst.
+    apply weber_aligned_sufficient;auto. }
+  assert(InA equiv w' ((L ^right) w ps)) as h_w_right.
+  { subst w'.
+    apply line_iP_min_InA.
+    - rewrite <- length_zero_iff_nil.
+
+      assert (0 <= (INR (length (line_left L w ps))))%R. 
+      { apply pos_INR. }
+      assert (INR (length (line_right L w ps)) > 0)%R as h_lgth.
+      { lra. }
+      intro abs3.
+      rewrite abs3 in h_lgth.
+      cbn in h_lgth.
+      lra.
+    - eapply aligned_on_inclA with (ps':=ps).
+      + apply line_right_inclA.
+      + rewrite aligned_on_cons_iff in Halign.
+        apply Halign. }
+  assert (abs2:w'<>w).
+  { specialize (line_right_spec L w w' ps) as h.
+    red in abs.
+    rewrite h in h_w_right.
+    intro abs4.
+    rewrite  abs4 in h_w_right.
+    lra. }
+  apply abs2.
+  now apply abs.
+Qed.
+
+
+(* there is at least one point in ps at the right of w. Let us
+   take the closest one p. hypothesis (left w - right w = on w)
+   implies that the point w' = (w+p)/2 is such that: (left w' -
+   right w' = 0 = on w') because left w' = left w + on w, and
+   right w' = right w. so by 1st order spec w' is a weber point
+   distinct from w, which contradict OnlyWeber w. *)
+
+Lemma weber_aligned_spec_right_bigger L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^left w ps)) - (^R (L^right w ps)) < 0)%R ->
+  OnlyWeber ps w -> 
+    (Rabs ((^R (L^left w ps)) - (^R (L^right w ps))) < ^R (L^on w ps))%R.
+Proof. 
+  intros Hpsnonnil Hlinedir Halign H_right HonlyWeber. 
+  specialize (weber_aligned_spec_weak Hlinedir Halign) as h_weak. 
+  destruct (HonlyWeber) as [hweber honly].
+  apply h_weak in hweber.
+  destruct (Rle_lt_or_eq_dec _ _ hweber) as [hlt | h_exact];clear hweber.
+    + auto.
+    + exfalso.
+      apply (@weber_aligned_single_weber_sufficient L ps w);auto.
+      rewrite Rabs_left in h_exact.
+      lra.
+      assumption.
+Qed.
+
+Lemma weber_aligned_spec_left_bigger L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) < 0)%R ->
+  OnlyWeber ps w -> 
+    (Rabs ((^R (L^left w ps)) - (^R (L^right w ps))) < ^R (L^on w ps))%R.
+Proof. 
+  intros Hpsnonnil Hlinedir Halign H_right HonlyWeber. 
+  specialize (weber_aligned_spec_weak Hlinedir Halign) as h_weak. 
+  destruct (HonlyWeber) as [hweber honly].
+  apply h_weak in hweber.
+  destruct (Rle_lt_or_eq_dec _ _ hweber) as [hlt | h_exact];clear hweber.
+    + auto.
+    + exfalso.
+      
+      apply (@weber_aligned_single_weber_sufficient (line_reverse L) ps w);auto.
+      * cbn -[opp equiv].
+        intros abs.
+        rewrite <- opp_origin in abs.
+        apply opp_reg in abs.
+        contradiction.
+      * now rewrite <- aligned_on_reverse.
+      * rewrite line_reverse_right, line_reverse_left,line_reverse_on.
+        rewrite Rabs_right in h_exact.
+        -- lra.
+        -- lra.
+      * now rewrite line_reverse_right, line_reverse_left.
+Qed.
+
+
+Lemma INR_length_0: forall A (l:list A) , (INR (length l)) = 0%R -> l = nil.
+Proof.
+  intros A l H. 
+  destruct (Nat.eq_decidable (length l) 0).
+  - now apply length_0.
+  - assert (0 < length l).
+    { lia. }
+    apply (lt_0_INR (length l)) in H1.
+    exfalso.
+    lra.
+Qed.
+
+Lemma weber_aligned_spec_on L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) = 0)%R ->
+  OnlyWeber ps w -> 
+  (0 < ^R (L^on w ps))%R.
+Proof.
+  intros h_neq_ps_nil h_nequiv_line_dir_origin h_align_L_cons h_eq_Rminus_ h_OnlyWeber_ps_w_.
+  destruct (Rlt_dec 0 (^R (L ^on) w ps)%R) as [h|h];auto.
+  exfalso.
+  assert (aligned_on L ps) as h_align_ps.
+  { apply aligned_on_cons_iff in h_align_L_cons.
+    apply h_align_L_cons. }
+  assert (((^R (L ^on) w ps) = 0)%R) as h_w_not_occupied.
+  { assert (0 <= (^R (L ^on) w ps))%R.
+    { apply pos_INR. }
+    lra. }
+  remember ((L ^-P) ((L ^max) ((L ^left) w ps))) as w'.
+  assert ((L ^left) w ps <> nil) as h_left_w_nonnil.
+  { specialize (line_left_on_right_partition L w ps) as h_tri.
+    apply PermutationA_length in h_tri.
+    repeat rewrite app_length in h_tri. 
+    apply (f_equal INR) in h_tri.
+    repeat rewrite plus_INR in h_tri.
+    intro abs. 
+    apply (f_equal (@length R2)) in abs.
+    apply (f_equal INR) in abs.
+    cbn [length INR] in abs.
+    rewrite abs in h_tri.
+    assert ((^R ps) > 0)%R.
+    { apply lt_0_INR.
+      apply Nat.neq_0_lt_0.
+      intro abs2.
+      apply length_zero_iff_nil in abs2.
+      contradiction. }
+    lra. }
+  assert (Weber ps w') as h_weber_w'.
+  { assert (aligned_on L (w' :: ps)) as h_align_w'.
+    { apply aligned_on_cons_iff.
+      split;auto.
+      subst w'.
+      apply line_contains_proj_inv;auto. }
+    assert (((L ^P) w' < (L ^P) w)%R) as h_lt_on_w_w'.
+    { destruct (line_left_spec L w w' ps) as [ h_left1 h_left2 ].
+        apply h_left1.
+        rewrite Heqw'.
+        apply line_iP_max_InA;auto.
+        apply aligned_on_inclA with (ps':=ps).
+        - apply line_left_inclA.
+        - now apply (aligned_on_cons_iff L w). }
+    apply (@weber_aligned_spec_weak L ps w');auto.
+    specialize (bipartition_max L ((L^left) w ps)) as h''.
+    rewrite <- Heqw' in h''.
+    rewrite <- aggravate_left in h'';auto.
+    assert (PermutationA equiv
+              ((L ^on) w' ((L ^left) w ps))
+              ((L ^on) w' ps))%R as h_simpl_on.
+    { rewrite <- aggravate_on_left;auto.
+      reflexivity. }
+    rewrite h_simpl_on in h''.
+    assert (PermutationA equiv ((L ^right) w' ps) ((L ^right) w ps))%R as h_right_same.
+    { assert (PermutationA equiv ((L ^right) w' ps) ((L ^right) w ps)) as h_right_w_w'.
+      { apply aggravate_right';auto.
+        - rewrite Heqw'.
+          rewrite line_right_iP_max;auto.
+        - now apply INR_length_0. }
+      rewrite h_right_w_w'.
+      reflexivity. }
+    assert ((^R (L ^on) w' ps) > 0)%R.
+    { specialize (line_iP_max_InA L ((L ^left) w ps) h_left_w_nonnil) as h_max.
+      rewrite <- Heqw' in h_max.
+      assert(aligned_on L ((L ^left) w ps)).
+      { apply aligned_on_inclA with ps; auto.
+        now apply line_left_inclA. }
+      specialize h_max with (1:=H).
+      assert(((L ^P) w') == ((L ^P) w')) as h_tmp by reflexivity.
+      pose (conj h_max h_tmp).
+      clearbody a.
+      
+      rewrite <- line_on_spec in a.
+      assert(InA equiv w' ((L ^on) w' ps)).
+      { assert (inclA equiv ((L ^on) w' ((L ^left) w ps)) ((L ^on) w' ps)) as h_incl.
+        { apply line_on_inclA_ps.
+          apply line_left_inclA. }
+        now apply h_incl. }
+      apply (@InA_nth _ equiv 0%VS w' ((L ^on) w' ps)) in H0.
+      destruct H0 as [n [y [h1 h2]]].
+      assert (0 < length ((L ^on) w' ps)) by lia.
+      apply lt_INR in H0.
+      cbn in H0.
+      lra. }
+    assert ((^R (L ^left) w' ps) < (^R (L ^right) w' ps))%R as h_left_less.
+    { apply PermutationA_length in h_right_same.
+      rewrite h_right_same.
+      assert ((^R (L ^right) w ps) = (^R (L ^left) w ps))%R as h_eq_ww'.
+      { lra. }
+      rewrite h_eq_ww'.
+      rewrite h''.
+      rewrite app_length,plus_INR.
+      lra. }
+    apply Req_le.
+    rewrite <- h_right_same in h_eq_Rminus_.
+    rewrite h'' in h_eq_Rminus_.
+    rewrite app_length in h_eq_Rminus_. 
+    rewrite plus_INR in h_eq_Rminus_.
+    rewrite Rabs_left1;
+    lra. }
+  assert (w =/= w') as h_neq.
+  { intro abs.
+    destruct ((L ^left) w ps) eqn:heq.
+    { now destruct h_left_w_nonnil. }
+    apply (line_left_diff L w ps);auto.
+    - rewrite heq.
+      auto.
+    - now rewrite heq, <- Heqw'. }
+  destruct h_OnlyWeber_ps_w_ as [hweb hwebeq].
+  apply h_neq.
+  symmetry.
+  eapply hwebeq;eauto.
+Qed.
+
+Lemma right_lt: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} 
+                       (L : line) (ps : list T) (w w' : T), 
+    ((L ^P) w' < (L ^P) w)%R ->
+    ((^R (L ^right) w' ps) >= (^R (L ^right) w ps) + (^R (L ^on) w ps))%R.
+Proof.
+  intros T S0 EQ0 VS ES L ps w w' h_lt. 
+  specialize (line_left_on_right_partition L w ps) as h.
+  assert (PermutationA equiv ((L ^right) w' ps) ((L ^right) w' ((L ^left) w ps ++ (L ^on) w ps ++ (L ^right) w ps)))
+    as heq.
+  { unfold line_right at 1 2.
+    apply filter_PermutationA_compat;auto.
+    typeclasses eauto. }
+  unfold line_right at 2 in heq.
+  repeat rewrite filter_app in heq.
+  change (filter (fun y : T => Rltb ((L ^P) w') ((L ^P) y)) ((L ^on) w ps))
+    with ((L ^right) w' ((L ^on) w ps))
+    in heq.
+  change (filter (fun y : T => Rltb ((L ^P) w') ((L ^P) y)) ((L ^right) w ps))
+    with ((L ^right) w' ((L ^right) w ps))
+    in heq.
+  assert (PermutationA equiv ((L ^right) w' ((L ^right) w ps)) ((L ^right) w ((L ^right) w' ps))) as h_perm_right.
+  { unfold line_right. 
+    rewrite filter_comm.
+    reflexivity. }
+  rewrite h_perm_right in heq.
+  rewrite <- aggravate_right in heq;auto.
+  assert (PermutationA equiv ((L ^right) w' ((L ^on) w ps)) ((L ^on) w ps)) as h_on_right.
+  { unfold line_right, line_on.
+    rewrite filter_weakenedA;auto.
+    - reflexivity.
+    -  intros x h1 h2.
+       rewrite Rltb_true.
+       rewrite eqb_true_iff in h2.
+       rewrite <- h2.
+       assumption. }
+  rewrite h_on_right in heq.
+  rewrite heq.
+  repeat rewrite app_length.
+  repeat rewrite plus_INR.
+  match goal with |- (?A + _ >= _)%R => remember A as hh
+  end.
+  assert (hh >= 0)%R.
+  { rewrite Heqhh.
+    apply Rle_ge.
+    apply pos_INR. }
+  lra.
+Qed.
+
+Lemma left_lt: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} 
+                       (L : line) (ps : list T) (w w' : T), 
+    ((L ^P) w < (L ^P) w')%R ->
+    ((^R (L ^left) w' ps) >= (^R (L ^left) w ps) + (^R (L ^on) w ps))%R.
+Proof.
+  intros T S0 EQ0 VS ES L ps w w' h_lt. 
+  specialize (line_left_on_right_partition L w ps) as h.
+  assert (PermutationA equiv ((L ^left) w' ps) ((L ^left) w' ((L ^left) w ps ++ (L ^on) w ps ++ (L ^right) w ps)))
+    as heq.
+  { unfold line_left at 1 2.
+    apply filter_PermutationA_compat;auto.
+    repeat red.
+    intros x y heqxy. 
+    now rewrite heqxy. }
+  unfold line_left at 2 in heq.
+  repeat rewrite filter_app in heq.
+  change (filter (fun y : T => Rltb ((L ^P) y) ((L ^P) w')) ((L ^on) w ps))
+    with ((L ^left) w' ((L ^on) w ps))
+    in heq.
+  change (filter (fun y : T => Rltb ((L ^P) y) ((L ^P) w')) ((L ^left) w ps))
+    with ((L ^left) w' ((L ^left) w ps))
+    in heq.
+  assert (PermutationA equiv ((L ^left) w' ((L ^left) w ps)) ((L ^left) w ((L ^left) w' ps))) as h_perm_left.
+  { unfold line_left.
+    rewrite filter_comm.
+    reflexivity. }
+  rewrite h_perm_left in heq.
+  rewrite <- aggravate_left in heq;auto.
+  assert (PermutationA equiv ((L ^left) w' ((L ^on) w ps)) ((L ^on) w ps)) as h_on_left.
+  { unfold line_left, line_on.
+    rewrite filter_weakenedA;auto.
+    - reflexivity.
+    -  intros x h1 h2.
+       rewrite Rltb_true.
+       rewrite eqb_true_iff in h2.
+       rewrite <- h2.
+       assumption. }
+  rewrite h_on_left in heq.
+  rewrite heq.
+  repeat rewrite app_length.
+  repeat rewrite plus_INR.
+  match goal with |- (_ + (_+?A) >= _)%R => remember A as hh
+  end.
+  assert (hh >= 0)%R.
+  { rewrite Heqhh.
+    apply Rle_ge.
+    apply pos_INR. }
+  lra.
+Qed.
+
+Lemma left_weaker: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} 
+                          (L : line) (ps : list T) (w w' : T), 
+    ((L ^P) w' < (L ^P) w)%R -> 
+    ((^R (L ^left) w' ps) <= (^R (L ^left) w ps))%R.
+Proof.
+  intros T S0 EQ0 VS ES L ps w w' h_Rlt_line_proj.
+  unfold line_left.
+  rewrite <- (filter_weakenedA (fun y => Rltb ((L ^P) y) ((L ^P) w')) (fun y => Rltb ((L ^P) y) ((L ^P) w))).
+  - rewrite filter_comm.
+    apply le_INR.
+    rewrite filter_length at 1.
+    lia.
+  - intros x0 H0 H1.
+    rewrite Rltb_true in H1 |- *.
+    lra.
+Qed.
+
+Lemma right_weaker: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} 
+                           (L : line) (ps : list T) (w w' : T), 
+    ((L ^P) w < (L ^P) w')%R -> 
+    ((^R (L ^right) w' ps) <= (^R (L ^right) w ps))%R.
+Proof.
+  intros T S0 EQ0 VS ES L ps w w' h_Rlt_line_proj.
+  unfold line_right.
+  rewrite <- (filter_weakenedA (fun y => Rltb ((L ^P) w') ((L ^P) y)) (fun y => Rltb ((L ^P) w) ((L ^P) y))).
+  - rewrite filter_comm.
+    apply le_INR.
+    rewrite filter_length at 1.
+    lia.
+  - intros x0 H0 H1.
+    rewrite Rltb_true in H1 |- *.
+    lra.
+Qed.
+
+Lemma INR_add_le: forall x y r: nat,
+    x + y < r ->
+    (((INR x) + (INR y) + 1 <= INR r)%R) .
+Proof.
+  intros x y r h_Rlt.
+  rewrite <- plus_INR.
+  now apply NumberComplements.INR_add_le.
+Qed.
+
+Lemma INR_add_le': forall x y r: nat,
+    ((INR x) + (INR y) < INR r)%R ->
+    (((INR x) + (INR y) + 1 <= INR r)%R) .
+Proof.
+  intros x y r h_Rlt.
+  apply INR_add_le.
+  rewrite <- plus_INR in h_Rlt;auto.
+  now apply INR_lt.
+Qed.
+
+Lemma INR_minus_le: forall x y r: nat,
+    y <= x ->
+    x - y < r ->
+    (((INR x) - (INR y) + 1 <= INR r)%R) .
+Proof.
+  intros x y r h_le h_Rlt.
+  rewrite <- minus_INR;auto.
+  now apply NumberComplements.INR_add_le.
+Qed.
+
+Lemma INR_minus_le': forall x y r: nat,
+    y <= x ->
+    ((INR x) - (INR y) < INR r)%R ->
+    (((INR x) - (INR y) + 1 <= INR r)%R) .
+Proof.
+  intros x y r h_le h_Rlt.
+  apply INR_minus_le;auto.  
+  rewrite <- minus_INR in h_Rlt;auto.
+  now apply INR_lt.
+Qed.
+
+Lemma weber_aligned_spec_right_bigger_2 L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^left w ps)) - (^R (L^right w ps)) < 0)%R ->
+  (Rabs ((^R (L^left w ps)) - (^R (L^right w ps))) < ^R (L^on w ps))%R ->
+  OnlyWeber ps w.
+Proof. 
+  intros h_neq_ps_nil h_line_dir h_align h_left_right h_Rlt_Rabs.
+  specialize (weber_aligned_spec_weak h_line_dir h_align) as h_weak.
+  
+  alias (^R (L ^left) w ps) as leftw after h_neq_ps_nil.
+  alias (^R (L ^right) w ps) as rightw after h_neq_ps_nil.
+  alias (^R (L ^on) w ps) as onw after h_neq_ps_nil.
+
+  apply Rlt_le in h_Rlt_Rabs as h_Rle_Rabs.
+  apply <- h_weak in h_Rle_Rabs.
+  clear h_weak.
+  split;[now auto|].
+  intros x h_weber.
+  destruct (equiv_dec x w);auto.
+  exfalso.
+  apply aligned_on_cons_iff in h_align as h.
+  destruct h as [h_line h_align_ps].
+  assert (aligned_on L (x::ps)) as h_align_ps'.
+  { apply aligned_on_cons_iff.
+    split;auto.
+    eapply weber_aligned;eauto. }
+  specialize (weber_aligned h_neq_ps_nil h_align_ps h_weber) as h_align_x.
+  specialize (weber_aligned_spec_weak_right h_line_dir h_neq_ps_nil h_align_ps h_weber) as h_rabs.
+  apply Rabs_left in h_left_right as hrabs_eq.
+  rewrite hrabs_eq in *.
+  clear hrabs_eq.
+  rewrite Ropp_minus_distr' in h_Rlt_Rabs.
+  assert (((^R (L ^right) w ps) - (^R (L ^left) w ps) + 1 <= (^R (L ^on) w ps))%R)
+    as h_weber_strong.
+  { rewrite <- minus_INR.
+    - apply INR_add_le_INR.
+      rewrite minus_INR.
+      + reremember.
+        assumption.
+      + reremember.
+        apply INR_le;lra.
+    - apply INR_le;lra. }
+  clear h_Rlt_Rabs.
+  move x before w.
+  alias (^R (L ^left) x ps) as leftx after h_neq_ps_nil.
+  alias (^R (L ^right) x ps) as rightx after h_neq_ps_nil.
+  alias (^R (L ^on) x ps) as onx after h_neq_ps_nil.
+  assert (onx+leftx+rightx = onw+leftw+rightw)%R as h_eq_partition.
+  { specialize (line_left_on_right_partition L x ps) as h1.
+    specialize (line_left_on_right_partition L w ps) as h2.
+    rewrite h1 in h2 at 1.
+    apply PermutationA_length in h2.
+    apply (f_equal INR) in h2.
+    repeat rewrite app_length in h2.
+    repeat rewrite plus_INR in h2.
+    reremember in h2.
+    lra. }
+  assert (L^P x <> L^P w) as h_proj.
+  { intro abs.
+    apply (f_equal (line_proj_inv L)) in abs.
+    rewrite <- h_line in c.
+    rewrite <- h_align_x in c.
+    contradiction. }
+  specialize (Rdichotomy _ _ h_proj) as [h_proj' | h_proj'].
+  - assert (rightx >= rightw + onw )%R as h_le_rightxw.
+    { unremember.
+      now apply right_lt. }
+    assert (onx >= 0)%R as h_onx_pos.
+    { rewrite Heqonx.
+      apply Rle_ge.
+      apply pos_INR. }
+    assert ((rightx - leftx) <= onx)%R as h_balance.
+    { rewrite Rabs_left1 in h_rabs.
+      - lra.
+      - lra. }
+    lra.
+  - assert (onw > 0)%R by lra.
+    assert (leftx >= leftw + onw )%R.
+    { unremember.
+      now apply left_lt. }
+    assert (rightx <= rightw)%R.
+    { unremember.
+      now apply right_weaker. }
+    assert (leftx > rightx)%R by lra.
+    assert (leftx - rightx <= onx)%R as h_balance.
+    { rewrite  Rabs_pos_eq in h_rabs;lra. }
+    clear h_rabs.
+    clear h_proj.
+    assert (rightx + onx <= rightw)%R by lra.
+    lra.    
+Qed.
+
+Lemma weber_aligned_spec_left_bigger_2 L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) < 0)%R ->
+  (Rabs ((^R (L^left w ps)) - (^R (L^right w ps))) < ^R (L^on w ps))%R ->
+  OnlyWeber ps w.
+Proof.
+  intros h_neq_ps_nil h_line_dir h_align h_left_right h_Rlt_Rabs.
+  specialize (@weber_aligned_spec_right_bigger_2 (line_reverse L) ps w h_neq_ps_nil) as h.
+  apply h;clear h.
+  - intro abs.
+    apply h_line_dir.
+    clear h_line_dir.
+    cbn in abs.
+    inversion abs.
+    rewrite H0 in H1.
+    destruct L;cbn in H0, H1.
+    cbn.
+    apply pair_eqE;simpl.
+    split;lra.
+  - now apply aligned_on_reverse in h_align.
+  - rewrite line_reverse_left, line_reverse_right.
+    assumption.
+  - rewrite line_reverse_left, line_reverse_right, line_reverse_on.
+    rewrite Rabs_minus_sym.
+    assumption.
+Qed.    
+
+Lemma weber_aligned_spec_on_2 L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  ((^R (L^right w ps)) - (^R (L^left w ps)) = 0)%R ->
+  (0 < ^R (L^on w ps))%R ->
+  OnlyWeber ps w.
+Proof.
+  intros h_neq_ps_nil h_line_dir h_align h_left_right h_Rlt_onw.
+  specialize (weber_aligned_spec_weak h_line_dir h_align) as h_weak.  
+  alias (^R (L ^left) w ps) as leftw after h_neq_ps_nil.
+  alias (^R (L ^right) w ps) as rightw after h_neq_ps_nil.
+  alias (^R (L ^on) w ps) as onw after h_neq_ps_nil.
+  split.
+  { apply h_weak.
+    rewrite Rabs_minus_sym.
+    rewrite h_left_right.
+    rewrite Rabs_R0.
+    lra. }
+  clear h_weak.
+  intros x h_weber.
+  move x before w.
+  destruct (equiv_dec x w);auto.
+  exfalso.
+  apply aligned_on_cons_iff in h_align as h.
+  destruct h as [h_line h_align_ps].
+  assert (aligned_on L (x::ps)) as h_align_ps'.
+  { apply aligned_on_cons_iff.
+    split;auto.
+    eapply weber_aligned;eauto. }
+  specialize (weber_aligned h_neq_ps_nil h_align_ps h_weber) as h_align_x.
+  specialize (weber_aligned_spec_weak_right h_line_dir h_neq_ps_nil h_align_ps h_weber) as h_rabs.
+  alias (^R (L ^left) x ps) as leftx after h_neq_ps_nil.
+  alias (^R (L ^right) x ps) as rightx after h_neq_ps_nil.
+  alias (^R (L ^on) x ps) as onx after h_neq_ps_nil.
+  assert (onx+leftx+rightx = onw+leftw+rightw)%R as h_eq_partition.
+  { specialize (line_left_on_right_partition L x ps) as h1.
+    specialize (line_left_on_right_partition L w ps) as h2.
+    rewrite h1 in h2 at 1.
+    apply PermutationA_length in h2.
+    apply (f_equal INR) in h2.
+    repeat rewrite app_length in h2.
+    repeat rewrite plus_INR in h2.
+    reremember in h2.
+    lra. }
+  assert (L^P x <> L^P w) as h_proj.
+  { intro abs.
+    apply (f_equal (line_proj_inv L)) in abs.
+    rewrite <- h_line in c.
+    rewrite <- h_align_x in c.
+    contradiction. }
+  specialize (Rdichotomy _ _ h_proj) as [h_proj' | h_proj']; clear h_proj.
+  - assert (rightx >= rightw + onw )%R as h_le_rightxw.
+    { unremember.
+      now apply right_lt. }
+    assert (onx >= 0)%R as h_onx_pos.
+    { rewrite Heqonx.
+      apply Rle_ge.
+      apply pos_INR. }
+    assert ((rightx - leftx) <= onx)%R as h_balance.
+    { rewrite Rabs_left1 in h_rabs.
+      - lra.
+      - lra. }
+    lra.
+  - assert (onw > 0)%R by lra.
+    assert (leftx >= leftw + onw )%R.
+    { unremember.
+      now apply left_lt. }
+    assert (rightx <= rightw)%R.
+    { unremember.
+      now apply right_weaker. }
+    assert (leftx > rightx)%R by lra.
+    assert (leftx - rightx <= onx)%R as h_balance.
+    { rewrite  Rabs_pos_eq in h_rabs;lra. }
+    clear h_rabs.
+    assert (rightx + onx <= rightw)%R by lra.
+    lra.
+Qed.
+
+Lemma weber_aligned_spec L ps w : 
+  ps <> nil ->                  (* Seems necessary *)
+  line_dir L =/= 0%VS ->
+  aligned_on L (w :: ps) -> 
+  OnlyWeber ps w <-> 
+  (Rabs ((^R (L^left w ps)) - (^R (L^right w ps))) < ^R (L^on w ps))%R.
+    (* (Rabs (INR (length (L^left w ps)) - INR (length (L^right w ps))) < INR (length (L^on w ps)))%R. *)
+Proof.
+  intros Hpsnonnil Hlinedir Halign. 
+  split.
+  { intro HonlyWeber.
+    destruct ((Req_appart_dec ((^R (L^right w ps)) - (^R (L^left w ps))) 0)%R) as [ | [| ]].
+    - rewrite Rabs_right;try lra.
+      assert ((^R (L ^left) w ps) - (^R (L ^right) w ps) = 0)%R as heq.
+      { lra. }
+      rewrite heq.
+      apply weber_aligned_spec_on;auto.
+    - apply weber_aligned_spec_left_bigger;auto.
+    - apply weber_aligned_spec_right_bigger;auto.
+      lra. }
+  { intro Hlt.
+    destruct ((Req_appart_dec ((^R (L^right w ps)) - (^R (L^left w ps))) 0)%R) as [ | [| ]].
+    - apply weber_aligned_spec_on_2 with L;auto.
+      assert (((^R (L ^left) w ps) - (^R (L ^right) w ps))%R = 0%R) by lra.
+      rewrite H,Rabs_R0 in Hlt.
+      assumption.
+    - apply weber_aligned_spec_left_bigger_2 with L;auto.
+    - apply weber_aligned_spec_right_bigger_2 with L;auto.
+      lra. }
+Qed.
+
+Lemma weber_majority ps w : 
+  countA_occ equiv R2_EqDec w ps > (Nat.div2 (length ps + 1)) -> OnlyWeber ps w. 
+Proof.
+  intros h_count. 
+  assert (Weber ps w) as h_weberw.
+  { apply weber_majority_weak.
+    lia. }
+  destruct (aligned_dec ps) as [h_align | h_nalign].
+  2:{ apply weber_Naligned_unique;auto. }
+  alias (line_calc ps) as L after h_count.
+  assert (aligned_on L ps) as h_aligned_on.
+  { unremember.
+    now apply line_calc_spec. }
+  assert (ps <>nil) as h_nonnil.
+  { intro abs.
+    rewrite <- length_zero_iff_nil in abs.
+    specialize (countA_occ_length_le equiv R2_EqDec ps w) as h.
+    rewrite abs in *.
+    cbn -[equiv] in h_count.
+    lia. }
+  destruct (equiv_dec (line_dir L) 0%VS) as [h_eq0 | h_neq0].
+  { assert (forall x : R2, {w == x} + {~ w == x}) as h_dec.
+    { intros x.
+      destruct (R2dec_bool w x) eqn:heq.
+      - rewrite R2dec_bool_true_iff in heq.
+        now left.
+      - rewrite R2dec_bool_false_iff in heq.
+        now right. }
+    edestruct (Forall_dec (equiv w) h_dec ps) as [h_forall | h_nforall].
+    - assert (exists ps', PermutationA equiv ps (w::ps')).
+      { apply PermutationA_split;auto.
+        assert (countA_occ equiv R2_EqDec w ps > 0).
+        { set (M:=Nat.div2 (length ps + 1)) in *. (* regression  for lia *)
+          lia. }
+        rewrite <- countA_occ_pos;eauto. }
+      destruct H.
+      rewrite H in h_forall.
+      apply weber_gathered in h_forall.
+      rewrite H.
+      assumption.
+    - exfalso.
+      apply neg_Forall_Exists_neg in h_nforall;auto.
+      apply Exists_exists in h_nforall.
+      destruct h_nforall as [e [h_in_e h_neq_e]].
+      assert (InA equiv e ps).
+      { now apply InA_Leibniz. }
+      assert (line_contains L w) as h_contains_w.
+      { apply weber_aligned with ps;auto. }
+      assert (line_contains L e) as h_contains_e.
+      { apply aligned_on_InA_contains with ps;auto. }
+      specialize (line_dir_nonnul L w e h_neq_e h_contains_w h_contains_e) as h.
+      now apply h. }
+  specialize (line_left_on_right_partition L w ps) as h_partition.
+  apply PermutationA_length in h_partition.
+  repeat rewrite app_length in h_partition.
+  assert (aligned_on L (w :: ps)) as h_alignwps.
+  { apply aligned_on_cons_iff.
+    split;auto.
+    apply weber_aligned with ps;auto. }
+  specialize (line_on_length_aligned L w ps h_alignwps) as h_on_countA.
+  assert (length ((L ^left) w ps) + length ((L ^right) w ps) < Nat.div2 (length ps + 1)).
+  { specialize (div2_sum_p1_ge (length ps)) as h_le.
+    lia. }
+  apply weber_aligned_spec with (L:=L);auto.
+  - assert (0 <= (^R (L ^left) w ps) <  (^R (L ^on) w ps))%R.
+    { split.
+      - apply pos_INR.
+      - apply lt_INR.
+        lia. }
+    assert (0 <= (^R (L ^right) w ps) <  (^R (L ^on) w ps))%R.
+    { split.
+      - apply pos_INR.
+      - apply lt_INR.
+        lia. }
+    apply Rabs_def1;lra.
+Qed.  
+
+Ltac duplic h name :=
+  set (name := h);
+  clearbody name;
+  move name before h.
+
+Ltac duplic_autoname h :=
+  let name := fresh h "'" in
+  duplic h name.
+
+Tactic Notation "dup" hyp(h) := duplic_autoname h.
+Tactic Notation "dup" hyp(h) "as" ident(name) := duplic h name.
+
+
+Lemma left_right_balanced_even ps (w:R2) L:
+  (Rabs ((^R (L ^left) w ps) - (^R (L ^right) w ps)) = 0)%R -> 
+  (^R (L ^on) w ps) = 0%R -> 
+  Nat.Even (length ps).
+Proof.
+  intros h_eq0 h_on_0.
+  assert (((^R (L ^left) w ps) = (^R (L ^right) w ps))) as h_left_right.
+  { destruct (Rle_lt_dec (^R (L ^left) w ps)  (^R (L ^right) w ps))%R as [h | h].
+    - rewrite Rabs_left1 in h_eq0;try lra.
+    - rewrite Rabs_pos_eq in h_eq0;try lra. }
+  specialize (line_left_on_right_partition L w ps) as h_partition.
+
+  apply PermutationA_length in h_partition.
+  repeat rewrite app_length in h_partition.
+  apply (f_equal INR) in h_partition.
+  repeat rewrite plus_INR in h_partition.
+  rewrite h_on_0, h_left_right in h_partition.
+  assert ((^R ps) = (2 * (^R (L ^right) w ps))%R) by lra.
+  change 2%R with (INR 2%nat) in H.
+  rewrite <- mult_INR in H.
+  apply INR_injection in H.
+  rewrite H.
+  eexists ;eauto.
+Qed.
+
+Lemma middle_diff_l: forall x y:R2, x<>y -> middle x y <> x.
+Proof.
+  intros x y H. 
+  specialize (middle_diff H) as h.
+  intro abs.
+  rewrite abs in h.
+  apply h.
+  constructor;auto.
+Qed.
+
+Lemma middle_diff_r: forall x y:R2, x<>y -> middle x y <> y.
+Proof.
+  intros x y H. 
+  specialize (middle_diff H) as h.
+  intro abs.
+  rewrite abs in h.
+  apply h.
+  constructor 2.
+  constructor;auto.
+Qed.
+
+Lemma weber_nonunique_segment_l L ps x w:
+  ps <> nil ->
+  aligned_on L ps ->
+  Weber ps w ->
+  Weber ps x ->
+  (L ^P x <  L ^P w)%R ->
+  (0 < ^R (L^on x ps))%R ->
+  (0 < ^R (L^on w ps))%R ->
+  forall m,
+    line_contains L m
+    -> (L ^P x < L ^P m)%R
+    -> (L ^P m < L ^P w)%R
+    -> Weber ps m /\ ((^R (L ^on) m ps) = 0)%R.
+Proof.
+  intros h_nonnil  h_aligned_on h_weber_w h_weber_x h_proj' h_x_gt0 h_w_gt0
+    m h_contain_m h_ltxm h_lt_mw.
+  alias (^R (L ^left) m ps) as leftm after h_nonnil.
+  alias (^R (L ^right) m ps) as rightm after h_nonnil.
+  alias (^R (L ^on) m ps) as onm after h_nonnil.
+  alias (^R (L ^left) w ps) as leftw after h_nonnil.
+  alias (^R (L ^right) w ps) as rightw after h_nonnil.
+  alias (^R (L ^on) w ps) as onw after h_nonnil.
+  alias (^R (L ^left) x ps) as leftx after h_nonnil.
+  alias (^R (L ^right) x ps) as rightx after h_nonnil.
+  alias (^R (L ^on) x ps) as onx after h_nonnil.
+  assert (0<=leftw)%R as h_lwpos by (unremember; apply pos_INR). 
+  assert (0<=leftx)%R as h_lxpos by (unremember; apply pos_INR).
+  assert (0<=rightw)%R as h_rwpos by (unremember; apply pos_INR).
+  assert (0<=rightx)%R as h_rxpos by (unremember; apply pos_INR).
+  assert (0<=leftm)%R as h_lmpos by (unremember; apply pos_INR).
+  assert (0<=rightm)%R as h_rmpos by (unremember; apply pos_INR).
+  assert (0<=onm)%R as onmpos by (unremember; apply pos_INR).
+
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_w) as ha_align_w.
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_x) as ha_align_x.
+  assert ((aligned_on L (w::ps))) as h_align_wps.
+  { apply aligned_on_cons_iff;split;auto. }
+  assert ((aligned_on L (x::ps)))  as h_align_xps.
+  { apply aligned_on_cons_iff;split;auto. }
+  assert (x=/=w) as h_neqxw.
+  { intros abs.
+    rewrite abs in h_proj'.
+    lra. }
+  assert (L^P x <> L^P w) as h_proj.
+  { intro abs.
+    apply (f_equal (line_proj_inv L)) in abs.
+    rewrite <- ha_align_w, <-ha_align_x in h_neqxw.
+    contradiction. }
+  assert (line_dir L =/= 0%VS) as h_dir_nonnull.
+  { apply line_dir_nonnul with x w;auto. }
+  dup h_weber_x.
+  apply (@weber_aligned_spec_weak L ps x h_dir_nonnull h_align_xps) in h_weber_x'.
+  reremember in h_weber_x'.
+  assert (leftx + onx + rightx = leftw + onw + rightw)%R as h_eq_xw.
+  {  transitivity (^R ps).
+     - rewrite (line_left_on_right_partition L x ps).
+       repeat rewrite app_length.
+       repeat rewrite plus_INR.
+       reremember;lra.
+     - rewrite (line_left_on_right_partition L w ps).
+       repeat rewrite app_length.
+       repeat rewrite plus_INR.
+       reremember;lra. }
+  
+  assert (leftw + onw + rightw = leftm + onm + rightm)%R.
+  {  transitivity (^R ps).
+     - rewrite (line_left_on_right_partition L w ps).
+       repeat rewrite app_length.
+       repeat rewrite plus_INR.
+       reremember;lra.
+     - rewrite (line_left_on_right_partition L m ps).
+       repeat rewrite app_length.
+       repeat rewrite plus_INR.
+       reremember;lra. }
+  assert (leftx + onm <= leftw)%R.
+  { subst leftx leftw onm.
+    transitivity ((^R (L ^left) m ps) + (^R (L ^on) m ps))%R.
+    - apply Rplus_le_compat_r.
+      apply left_weaker.
+      assumption.
+    - apply Rge_le.
+      apply left_lt;lra. }
+  assert (rightw + onm <= rightx)%R.
+  { subst rightw rightx onm.
+    transitivity ((^R (L ^right) m ps) + (^R (L ^on) m ps))%R.
+    - apply Rplus_le_compat_r.
+      apply right_weaker.
+      assumption.
+    - apply Rge_le.
+      apply right_lt;lra. }
+  assert (rightm >= rightw + onw)%R.
+  { subst rightm rightw onw.
+    apply right_lt;lra. }
+  assert (leftm >= leftx + onx)%R.
+  { subst leftm onx leftx.
+    apply left_lt;lra. }
+  assert (~Rabs(leftx - rightx) < onx)%R as h_lt_balancex.
+  { intro abs.
+    unremember in abs.
+    apply <- weber_aligned_spec in abs;auto; reremember;auto.
+    red in abs.
+    destruct abs as [ _ h_uniq].
+    apply h_neqxw;auto.
+    symmetry;auto. }
+  assert (~Rabs(leftw - rightw) < onw)%R as h_lt_balancew.
+  { intro abs.
+    unremember in abs.
+    apply <- weber_aligned_spec in abs;auto; reremember;auto.
+    red in abs.
+    destruct abs as [ _ h_uniq].
+    apply h_neqxw;auto. }
+  assert (Rabs (leftx - rightx) = onx)%R as h_equilib_x by lra.
+  dup h_weber_w.
+  apply (@weber_aligned_spec_weak L ps w h_dir_nonnull h_align_wps) in h_weber_w'.
+  reremember in h_weber_w'.
+  assert (Rabs (leftw - rightw) = onw)%R as h_equilib_w by lra.
+
+  rewrite Rabs_left1 in h_equilib_x.
+  2:{ destruct (Rle_lt_dec leftx rightx);try lra.
+      exfalso.
+      rewrite Rabs_right in h_equilib_w; try lra. }
+  rewrite Rabs_right in h_equilib_w; try lra.
+  split;[|lra].
+  apply (@weber_aligned_spec_weak L);auto; auto.
+  - apply aligned_on_cons_iff;auto.
+  - reremember.
+    apply Rabs_le;lra.
+Qed.
+
+Lemma weber_nonunique_segment L ps x w:
+  ps <> nil ->
+  aligned_on L ps ->
+  Weber ps w ->
+  Weber ps x ->
+  (L ^P x <> L ^P w)%R ->
+  (0 < ^R (L^on x ps))%R ->
+  (0 < ^R (L^on w ps))%R ->
+  forall m,
+    line_contains L m           (* FIXME: should not be needed *)
+    -> strict_segment x w m
+    -> Weber ps m /\ ((^R (L ^on) m ps) = 0)%R.
+Proof.
+  intros h_nonnil h_aligned_on h_weber_w h_weber_x h_diff_xw h_x_gt0 h_w_gt0 m h_contains_m h_strict_seg.
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_w) as ha_align_w.
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_x) as ha_align_x.
+  dup h_strict_seg.
+  rewrite (strict_segment_line L x w m) in h_strict_seg';auto.
+  destruct h_strict_seg as [h_seg _].
+  destruct h_strict_seg' as [h_ineq | h_ineq].
+  - destruct  h_ineq as [h_lt_xm h_lt_mw].
+    assert (((L ^P) x < (L ^P) w)%R) as h_lt_xw by lra.
+    apply (weber_nonunique_segment_l h_nonnil h_aligned_on h_weber_w h_weber_x h_lt_xw
+             h_x_gt0 h_w_gt0 h_contains_m h_lt_xm h_lt_mw).
+  - destruct h_ineq as [h_lt_wm h_lt_mx].
+    assert (((L ^P) w < (L ^P) x)%R) as h_lt_wx by lra.
+    apply (weber_nonunique_segment_l h_nonnil h_aligned_on h_weber_x h_weber_w h_lt_wx
+             h_w_gt0 h_x_gt0 h_contains_m h_lt_wm h_lt_mx).
+Qed.
+
+
+(* When there are an odd number of points, the weber point is always unique. *)
+Lemma weber_odd_unique ps w : 
+  Nat.Odd (length ps) -> Weber ps w -> OnlyWeber ps w.
+Proof.
+(* Proof sketch : if the points are not aligned, then the weber point is unique. 
+ * If they are aligned, take two distinct weber points w1 and w2. 
+ * 1. Show that w1 and w2 are also on the line using weber_aligned. 
+ * 2. Show that there is a weber point on the segment [w1, w2] 
+ *    that is not in ps (i.e. with multiplicity 0).
+ * 3. Using weber_aligned_spec show that n is even. *)
+  intros h_Odd_ps h_weber_w.
+  assert (ps<>nil) as h_nonnil.
+  { intro abs.
+    rewrite abs in h_Odd_ps.
+    cbn in h_Odd_ps.
+    inversion h_Odd_ps.
+    lia. }
+  destruct (aligned_dec ps) as [h_align | h_nalign].
+  2:{ apply weber_Naligned_unique;auto. }
+  split;auto.
+  intros x h_weber_x. 
+  move x at top.
+  destruct (equiv_dec x w) as [ | h_neqxw ];auto.
+  exfalso.
+  alias (line_calc ps) as L after h_weber_w.
+  assert (aligned_on L ps) as h_aligned_on.
+  { unremember.
+    now apply line_calc_spec. }
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_w) as ha_align_w.
+  specialize (weber_aligned h_nonnil h_aligned_on h_weber_x) as ha_align_x.
+  assert (line_dir L =/= 0%VS) as h_dir_nonnull.
+  { apply line_dir_nonnul with x w;auto. }
+  assert (L^P x <> L^P w) as h_proj.
+  { intro abs.
+    apply (f_equal (line_proj_inv L)) in abs.
+    rewrite <- ha_align_w, <-ha_align_x in h_neqxw.
+    contradiction. }
+  alias (^R (L ^left) w ps) as leftw after h_Odd_ps.
+  alias (^R (L ^right) w ps) as rightw after h_Odd_ps.
+  alias (^R (L ^on) w ps) as onw after h_Odd_ps.
+  alias (^R (L ^left) x ps) as leftx after h_Odd_ps.
+  alias (^R (L ^right) x ps) as rightx after h_Odd_ps.
+  alias (^R (L ^on) x ps) as onx after h_Odd_ps.
+(*  
+  assert (exists pt, Weber ps pt /\ ((^R (L ^on) pt ps) = 0)%R) as h_ex_pt.
+  { specialize (Rdichotomy _ _ h_proj) as [h_proj' | h_proj'];clear h_proj.
+    - destruct (equiv_dec onw  0)%R  as [? | h_w_neq0].
+      { exists w;subst onw;auto. }
+      destruct (equiv_dec onx 0)%R as [? | h_x_neq0].
+      { exists x;subst onx;auto. }
+      assert (0 < onx)%R  as h_x_gt0.
+      { assert (onx <> 0 /\ onx >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      assert (0 < onw)%R  as h_w_gt0.
+      { assert (onw <> 0 /\ onw >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      clear h_w_neq0 h_x_neq0.
+      assert ((L ^P) x <> (L ^P) w)%R as h_diff_xw by lra.
+      specialize (weber_nonunique_segment h_nonnil h_aligned_on h_weber_w h_weber_x h_diff_xw) as h.
+      reremember in h.
+      specialize (h h_x_gt0 h_w_gt0).
+      exists (middle x w).
+      assert (line_contains L (middle x w)).
+      { apply line_contains_middle;auto. }
+      apply h;auto.
+      apply bar;auto.
+    -
+      
+*)
+  assert ((aligned_on L (w::ps))) as h_align_wps.
+  { apply aligned_on_cons_iff;split;auto. }
+  assert ((aligned_on L (x::ps)))  as h_align_xps.
+  { apply aligned_on_cons_iff;split;auto. }
+  (* assert (line_dir L =/= 0%VS) as h_dir_nonnull. *)
+  (* { apply line_dir_nonnul with x w;auto. } *)
+  dup h_weber_w.
+  rewrite -> (@weber_aligned_spec_weak L) in h_weber_w';auto.
+  reremember in h_weber_w'.
+  assert (0<=leftw)%R as h_lwpos by (unremember; apply pos_INR). 
+  assert (0<=leftx)%R as h_lxpos by (unremember; apply pos_INR).
+  assert (0<=rightw)%R as h_rwpos by (unremember; apply pos_INR).
+  assert (0<=rightx)%R as h_rxpos by (unremember; apply pos_INR).
+  dup h_weber_x.
+  apply (@weber_aligned_spec_weak L ps x h_dir_nonnull h_align_xps) in h_weber_x'.
+  reremember in h_weber_x'.
+  assert (~Rabs(leftx - rightx) < onx)%R as h_lt_balancex.
+  { intro abs.
+    unremember in abs.
+    apply <- weber_aligned_spec in abs;auto; reremember;auto.
+    red in abs.
+    destruct abs as [ _ h_uniq].
+    apply h_neqxw;auto.
+    symmetry;auto. }
+  assert (~Rabs(leftw - rightw) < onw)%R as h_lt_balancew.
+  { intro abs.
+    unremember in abs.
+    apply <- weber_aligned_spec in abs;auto; reremember;auto.
+    red in abs.
+    destruct abs as [ _ h_uniq].
+    apply h_neqxw;auto. }
+  assert (Rabs (leftx - rightx) = onx)%R as h_equilib_x by lra.
+  assert (Rabs (leftw - rightw) = onw)%R as h_equilib_w by lra.
+  clear h_weber_w' h_weber_x' h_lt_balancex h_lt_balancew.
+  assert (exists pt, Weber ps pt /\ ((^R (L ^on) pt ps) = 0)%R) as h_ex_pt.
+  { specialize (Rdichotomy _ _ h_proj) as [h_proj' | h_proj'];clear h_proj.
+    - destruct (equiv_dec onw  0)%R  as [? | h_w_neq0].
+      { exists w;subst onw;auto. }
+      destruct (equiv_dec onx 0)%R as [? | h_x_neq0].
+      { exists x;subst onx;auto. }
+      assert (0 < onx)%R  as h_x_gt0.
+      { assert (onx <> 0 /\ onx >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      assert (0 < onw)%R  as h_w_gt0.
+      { assert (onw <> 0 /\ onw >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      clear h_w_neq0 h_x_neq0.
+      assert (forall m,
+                 line_contains L m
+                 -> (L ^P x < L ^P m)%R
+                 -> (L ^P m < L ^P w)%R
+                 -> Weber ps m /\ ((^R (L ^on) m ps) = 0)%R) as h_segment_weber.
+      { intros m h_contain_m h_ltxm h_lt_mw.
+        alias (^R (L ^left) m ps) as leftm after h_Odd_ps.
+        alias (^R (L ^right) m ps) as rightm after h_Odd_ps.
+        alias (^R (L ^on) m ps) as onm after h_Odd_ps.            
+        assert (0<=leftm)%R as h_lmpos by (unremember; apply pos_INR).
+        assert (0<=rightm)%R as h_rmpos by (unremember; apply pos_INR).
+        assert (0<=onm)%R as onmpos by (unremember; apply pos_INR).
+        assert (leftx + onx + rightx = leftw + onw + rightw)%R as h_eq_xw.
+        {  transitivity (^R ps).
+           - rewrite (line_left_on_right_partition L x ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra.
+           - rewrite (line_left_on_right_partition L w ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra. }
+          
+        assert (leftw + onw + rightw = leftm + onm + rightm)%R.
+        {  transitivity (^R ps).
+           - rewrite (line_left_on_right_partition L w ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra.
+           - rewrite (line_left_on_right_partition L m ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra. }
+        assert (leftx + onm <= leftw)%R.
+        { subst leftx leftw onm.
+          transitivity ((^R (L ^left) m ps) + (^R (L ^on) m ps))%R.
+          - apply Rplus_le_compat_r.
+            apply left_weaker.
+            assumption.
+          - apply Rge_le.
+            apply left_lt;lra. }
+        assert (rightw + onm <= rightx)%R.
+        { subst rightw rightx onm.
+          transitivity ((^R (L ^right) m ps) + (^R (L ^on) m ps))%R.
+          - apply Rplus_le_compat_r.
+            apply right_weaker.
+            assumption.
+          - apply Rge_le.
+            apply right_lt;lra. }
+        assert (rightm >= rightw + onw)%R.
+        { subst rightm rightw onw.
+          apply right_lt;lra. }
+        assert (leftm >= leftx + onx)%R.
+        { subst leftm onx leftx.
+          apply left_lt;lra. }
+        rewrite Rabs_left1 in h_equilib_x.
+        all:swap 1 2.
+        { destruct (Rle_lt_dec leftx rightx);try lra.
+          exfalso.
+          rewrite Rabs_right in h_equilib_w; try lra. }
+        rewrite Rabs_right in h_equilib_w; try lra.
+        split;[|lra].
+        apply (@weber_aligned_spec_weak L);auto; auto.
+        - apply aligned_on_cons_iff;auto.
+        - reremember.
+          apply Rabs_le;lra. }
+      exists (L^-P ((L^P x + L^P w) * 1/2))%R.
+      apply h_segment_weber;auto.
+      + apply line_contains_proj_inv;auto.
+      + rewrite line_P_iP;auto.
+        lra.
+      + rewrite line_P_iP;auto.
+        lra.
+    - destruct (equiv_dec onw  0)%R  as [? | h_w_neq0].
+      { exists w;subst onw;auto. }
+      destruct (equiv_dec onx 0)%R as [? | h_x_neq0].
+      { exists x;subst onx;auto. }
+      assert (0 < onx)%R  as h_x_gt0.
+      { assert (onx <> 0 /\ onx >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      assert (0 < onw)%R  as h_w_gt0.
+      { assert (onw <> 0 /\ onw >= 0)%R;try split;try lra.
+        - assumption.
+        - unremember.          
+          apply Rle_ge , pos_INR. }
+      clear h_w_neq0 h_x_neq0.
+      assert (forall m,
+                 line_contains L m
+                 -> (L ^P w < L ^P m)%R
+                 -> (L ^P m < L ^P x)%R
+                 -> Weber ps m /\ ((^R (L ^on) m ps) = 0)%R) as h_segment_weber.
+      { intros m h_contain_m h_ltwm h_lt_mx.
+        alias (^R (L ^left) m ps) as leftm after h_Odd_ps.
+        alias (^R (L ^right) m ps) as rightm after h_Odd_ps.
+        alias (^R (L ^on) m ps) as onm after h_Odd_ps.            
+        assert (0<=leftm)%R as h_lmpos by (unremember; apply pos_INR).
+        assert (0<=rightm)%R as h_rmpos by (unremember; apply pos_INR).
+        assert (0<=onm)%R as onmpos by (unremember; apply pos_INR).
+        assert (leftx + onx + rightx = leftw + onw + rightw)%R as h_eq_xw.
+        {  transitivity (^R ps).
+           - rewrite (line_left_on_right_partition L x ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra.
+           - rewrite (line_left_on_right_partition L w ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra. }
+          
+        assert (leftw + onw + rightw = leftm + onm + rightm)%R.
+        {  transitivity (^R ps).
+           - rewrite (line_left_on_right_partition L w ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra.
+           - rewrite (line_left_on_right_partition L m ps).
+             repeat rewrite app_length.
+             repeat rewrite plus_INR.
+             reremember;lra. }
+        assert (leftw + onm <= leftx)%R.
+        { subst leftx leftw onm.
+          transitivity ((^R (L ^left) m ps) + (^R (L ^on) m ps))%R.
+          - apply Rplus_le_compat_r.
+            apply left_weaker.
+            assumption.
+          - apply Rge_le.
+            apply left_lt;lra. }
+        assert (rightx + onm <= rightw)%R.
+        { subst rightw rightx onm.
+          transitivity ((^R (L ^right) m ps) + (^R (L ^on) m ps))%R.
+          - apply Rplus_le_compat_r.
+            apply right_weaker.
+            assumption.
+          - apply Rge_le.
+            apply right_lt;lra. }
+        assert (rightm >= rightx + onx)%R.
+        { subst rightm rightx onx.
+          apply right_lt;lra. }
+        assert (leftm >= leftw + onw)%R.
+        { subst leftm onw leftw.
+          apply left_lt;lra. }
+
+
+        rewrite Rabs_left1 in h_equilib_w.
+        all:swap 1 2.
+        { destruct (Rle_lt_dec leftw rightw);try lra.
+          exfalso.
+          rewrite Rabs_right in h_equilib_x; try lra. }
+        rewrite Rabs_right in h_equilib_x; try lra.
+        split;[|lra].
+        apply (@weber_aligned_spec_weak L);auto; auto.
+        - apply aligned_on_cons_iff;auto.
+        - reremember.
+          apply Rabs_le;lra. }
+      exists (L^-P ((L^P x + L^P w) * 1/2))%R.
+      apply h_segment_weber;auto.
+      + apply line_contains_proj_inv;auto.
+      + rewrite line_P_iP;auto.
+        lra.
+      + rewrite line_P_iP;auto.
+        lra. }
+  destruct h_ex_pt as [pt [hpt1 hpt2]].
+  rewrite -> (@weber_aligned_spec_weak L) in hpt1;auto.
+  2:{ specialize (weber_aligned h_nonnil h_aligned_on hpt1) as ha_align_pt.
+      apply aligned_on_cons_iff;split;auto. }
+  assert (Rabs ((^R (L ^left) pt ps) - (^R (L ^right) pt ps)) >= 0)%R.
+  { apply Rle_ge.
+    apply Rabs_pos. }
+
+  assert (Nat.Even (length ps)).
+  { eapply left_right_balanced_even with (w := pt) (L:=L);try lra. }
+  apply Nat.Even_Odd_False with (length ps);auto.
+Qed.
+
+
+Lemma line_contains_def (L: @line R2) w:
+  line_contains L w -> 
+  L^-P (L ^P w) = w.
+Proof.
+  intros h_contain.
+  apply h_contain.
+Qed.
+
+
+Lemma weber_same L ps w:
+  ps <> nil ->
+  aligned_on L ps ->
+  Weber ps w ->
+  forall w', 
+  line_contains L w' ->
+  (L^on w' ps = L^on w ps)%R ->
+  (L^left w' ps = L^left w ps)%R ->
+  (L^right w' ps = L^right w ps)%R ->
+  Weber ps w'.
+Proof.
+  intros H H0 H1 w' H2 H3 H4 H5. 
+  destruct (equiv_dec w w').
+  { now rewrite <- e. }
+  assert (line_dir L =/= 0%VS) as h_dir_nonnull.
+  { apply line_dir_nonnul with w w';auto.
+    apply (@weber_aligned L ps) ;auto. }
+  apply (@weber_aligned_spec_weak L);auto.
+  { apply aligned_on_cons_iff;auto. }
+  rewrite H3,H4,H5.
+  apply (@weber_aligned_spec_weak L);auto.
+  {  apply aligned_on_cons_iff;split;auto.
+     apply (@weber_aligned L ps) ;auto. }
+Qed.
+
+Lemma Rabs_zero: forall r, Rabs r = 0%R -> r = 0%R.
+Proof.
+  intros r H.
+  destruct (equiv_dec r 0%R);auto.
+  exfalso.
+  assert (r <> 0%R) by auto.
+  apply Rabs_no_R0 in c.
+  contradiction.
+Qed.
+
+Lemma INR_zero n: INR n = 0%R <-> n = 0.
+Proof.
+  split.
+  - intros H. 
+    destruct n;auto.
+    rewrite S_INR in H.
+    exfalso.
+    assert (0 <= INR n)%R.
+    { apply pos_INR. }
+    lra.
+  - intros H. 
+    subst.
+    reflexivity.
+Qed.
+
+
+Lemma balance_conseq: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} (ps1 ps2 : list T),
+    (((^R ps1) - (^R ps2))%R = 0%R) ->
+    ps1 = nil -> ps2 = nil.
+Proof.
+  intros T S0 EQ0 VS ES ps1 ps2 h_balance_w H1.
+  subst.
+  cbn in h_balance_w.
+  rewrite Rminus_0_l in h_balance_w.
+  destruct (ps2);auto.
+  cbn [length] in h_balance_w.
+  apply Ropp_eq_0_compat in h_balance_w.
+  rewrite Ropp_involutive in h_balance_w.
+  apply INR_zero in h_balance_w.
+  discriminate.
+Qed.
+
+Lemma balance_conseq_2: forall {T : Type} {S0 : Setoid T} {EQ0 : EqDec S0} {VS : RealVectorSpace T} {H : EuclideanSpace T} (ps1 ps2 : list T),
+    (((^R ps1) - (^R ps2))%R = 0%R) ->
+    ps2 = nil -> ps1 = nil.
+Proof.
+  intros T S0 EQ0 VS ES ps1 ps2 h_balance_w H1.
+  subst.
+  cbn in h_balance_w.
+  rewrite Rminus_0_r in h_balance_w.
+  destruct (ps1);auto.
+  cbn [length] in h_balance_w.
+  apply INR_zero in h_balance_w.
+  discriminate.
+Qed.
+
+Lemma balance_nonnul_left: forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    (^R (L ^on) w ps) = 0%R ->
+    ((^R (L ^left) w ps) - (^R (L ^right) w ps))%R = 0%R ->
+    (L ^left) w ps <> nil.
+Proof.
+  intros L ps w h_nonnil h_unoccupied H1. 
+  intro abs.
+  remember ((L ^left)w ps) as leftw in *.
+  remember ((L ^right)w ps) as rightw in *.
+  remember ((L ^on)w ps) as onw in *.
+  assert (rightw=nil).
+  { apply balance_conseq with leftw;auto. }
+  apply h_nonnil.
+  apply Preliminary.PermutationA_nil with (eqA:=eq);auto.
+  rewrite (line_left_on_right_partition L w ps).
+  reremember.
+  rewrite abs, H.
+  rewrite app_nil_l,app_nil_r.
+  apply INR_zero in h_unoccupied.
+  rewrite length_zero_iff_nil in h_unoccupied.
+  now rewrite h_unoccupied.    
+Qed.
+
+Lemma balance_nonnul_right: forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    (^R (L ^on) w ps) = 0%R ->
+    ((^R (L ^left) w ps) - (^R (L ^right) w ps))%R = 0%R ->
+    (L ^right) w ps <> nil.
+Proof.
+  intros L ps w h_nonnil h_unoccupied H1.
+  specialize (balance_nonnul_left h_nonnil h_unoccupied H1) as h.
+  intro abs.
+  apply length_zero_iff_nil in abs.
+  apply (f_equal INR) in abs.
+  change (INR 0) with 0%R in abs.
+  apply h.
+  rewrite abs in H1.
+  apply length_zero_iff_nil.
+  apply INR_zero.
+  lra.
+Qed.
+
+Lemma weber_aligned_non_occupied:
+  forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    line_dir L =/= 0%VS ->
+    aligned_on L (w :: ps) ->
+    (^R (L ^on) w ps) = 0%R ->
+    Weber ps w ->
+    Weber ps ((L ^-P) ((L ^min) ((L ^right) w ps))). (* also treu for left *)
+Proof.
+  intros L ps w h_nonnil h_dir_nonnul h_align h_unoccupied h_weber_w.
+  assert (Rabs ((^R (L ^left) w ps) - (^R (L ^right) w ps)) = 0)%R as h_balance_w.
+  { specialize (@weber_aligned_spec_weak L ps w h_dir_nonnul h_align) as h.
+    apply h in h_weber_w.
+    specialize (Rabs_pos ((^R (L ^left) w ps) - (^R (L ^right) w ps))) as h'.
+    lra. }
+  apply Rabs_zero in h_balance_w.
+  remember ((L ^left)w ps) as leftw in *.
+  remember ((L ^right)w ps) as rightw in *.
+  remember ((L ^on)w ps) as onw in *.
+  remember ((L ^-P) ((L ^min) rightw)) as neighbor in *.
+
+  assert (leftw <> nil) as h_left_nonil.
+  { subst leftw.
+    apply balance_nonnul_left;subst;auto. }
+
+  assert (rightw <> nil) as h_right_nonil.
+  { subst.
+    apply balance_nonnul_right;auto. }
+
+  assert (aligned_on L ps) as h_align_ps.
+  { apply (aligned_on_cons_iff L w ps).
+    assumption. }      
+
+  assert (aligned_on L (neighbor :: ps)) as h_align_neigh.
+  { apply aligned_on_cons_iff;split;auto.
+    unremember. apply line_contains_proj_inv;auto. }
+
+  apply <- (@weber_aligned_spec_weak L);auto.
+  specialize (line_right_spec L w (L^-P((L ^min) ((L ^right) w ps))) ps) as hhh.
+  assert (((L ^P) w < (L ^P) neighbor)%R).
+  { subst.
+    apply hhh.
+    apply line_iP_min_InA;auto.
+    apply aligned_on_inclA with (ps':=ps);auto.
+    apply line_right_inclA. }
+
+  assert (PermutationA equiv ((L^left) (neighbor) ps) leftw) as h_perm_left_neigh.
+  { rewrite Heqleftw.
+    apply (aggravate_left' L ps w neighbor).
+    - assumption.
+    - reremember.
+      rewrite Heqneighbor.
+      apply line_left_iP_min.
+    - apply length_0.
+      apply INR_zero.
+      reremember.
+      lra. }
+
+  assert (PermutationA equiv ((L^right) w ps) 
+            ((L^on) (neighbor) ps ++ (L^right) (neighbor) rightw)) as h_perm_right_neigh.
+  { assert ((L ^on) ((L ^-P) ((L ^min) rightw)) ps
+            = (L ^on) ((L ^-P) ((L ^min) rightw))  rightw)%R as heq_rightw.
+    { unremember.
+      apply aggravate_on_right;auto.
+      reremember.
+      assumption. }
+    subst neighbor.
+    rewrite heq_rightw at 1.
+    reremember.
+    apply bipartition_min. }
+  apply PermutationA_length in h_perm_right_neigh.
+  rewrite app_length in h_perm_right_neigh.
+  apply (f_equal INR) in h_perm_right_neigh.
+  rewrite plus_INR in h_perm_right_neigh.
+  
+  assert ((^R (L ^on) neighbor ps) = (^R (L ^right) w ps) - (^R (L ^right) neighbor rightw))%R as h_balance' by lra.
+  rewrite h_balance'.
+  reremember in h_perm_right_neigh.
+  reremember.
+  assert (((^R leftw) = (^R rightw)))%R as h_balance_w' by lra.
+  reremember.
+  (* rewrite h_balance_w'. *)
+  apply Req_le.
+  assert ((^R (L ^right) neighbor ps)=(^R (L ^right) neighbor rightw))%R as h_same.
+  { unremember.
+    rewrite <- aggravate_right;auto.
+    rewrite line_P_iP;auto.
+    reremember.
+    specialize (line_right_spec L w ((L^-P) ((L ^min) rightw)) ps) as h_right.
+    assert (InA equiv ((L ^-P) ((L ^min) rightw)) ((L ^right) w ps)) as h_InA_minright.
+    { rewrite <- Heqrightw.
+      apply line_iP_min_InA;auto.
+      apply aligned_on_inclA with ps.
+      - unremember. apply line_right_inclA.
+      - apply (aligned_on_cons_iff L w ps);auto. }
+    rewrite h_right in h_InA_minright.
+    destruct h_InA_minright as [h_InA h_gt].
+    rewrite line_P_iP in h_gt;auto. }
+
+  rewrite Rabs_pos_eq.
+  - rewrite h_same.
+    rewrite h_perm_left_neigh.
+    rewrite h_balance_w'.
+    reflexivity.
+  - apply Rle_minus_sim.
+    
+    reremember in h_balance'.
+    rewrite h_same.
+    rewrite h_perm_left_neigh.
+    rewrite  h_balance_w'.
+    rewrite (line_left_on_right_partition L neighbor rightw) at 2.
+    assert ((^R (L ^left) neighbor rightw) >=0)%R.
+    { subst neighbor.
+      rewrite line_left_iP_min.
+      cbn.
+      lra. }
+
+    assert ((^R (L ^on) neighbor rightw) >=0)%R.
+    { apply Rle_ge.
+      apply pos_INR. }
+    repeat rewrite app_length.
+    repeat rewrite plus_INR.
+    lra.
+Qed.
+
+
+Lemma weber_aligned_non_occupied_2:
+  forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    line_dir L =/= 0%VS ->
+    aligned_on L (w :: ps) ->
+    (^R (L ^on) w ps) = 0%R ->
+    Weber ps w ->
+    Weber ps ((L ^-P) ((L ^max) ((L ^left) w ps))).
+Proof.
+  intros L' ps w h_nonnil h_dir_nonnul h_align h_unoccupied h_weber_w.
+  remember (line_reverse L') as L.
+  rewrite <- line_reverse_right.
+  rewrite <- line_reverse_iP_min.
+  reremember.
+  apply weber_aligned_non_occupied;auto.
+  - subst.
+    destruct L';cbn.
+    cbn in h_dir_nonnul.
+    intro abs.
+    apply h_dir_nonnul.
+    inversion abs.
+    destruct line_dir;cbn in *.
+    assert (r = 0%R) by lra.
+    assert (r0 = r%R) by lra.
+    rewrite H2,H.
+    reflexivity.
+  - subst.
+    rewrite <- aligned_on_reverse;auto.
+  - rewrite <- line_reverse_on in h_unoccupied.
+    subst.
+    assumption.
+Qed.
+
+Lemma weber_aligned_non_occupied_neighbor_l:
+  forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    line_dir L =/= 0%VS ->
+    aligned_on L (w :: ps) ->
+    (^R (L ^on) w ps) = 0%R ->
+    Weber ps w ->
+    forall x,
+      line_contains L x ->
+      PermutationA equiv (L^right x ps) (L^right w ps) ->
+      (^R (L ^on) x ps) = 0%R
+      -> Weber ps x.
+Proof.
+  intros L ps w h_nonnil h_dir_nonnul h_align h_unoccupied h_weber_w x h_line_contains_x h_neighbor h_on_0.
+  assert ( (^R (L ^on) w ps) =  ^R (L ^on) x ps) as heq_on.
+  { now rewrite h_on_0, h_unoccupied. }
+  assert ( (^R (L ^left) w ps) =  ^R (L ^left) x ps).
+  { specialize (line_left_on_right_partition L x ps) as h.
+    rewrite (line_left_on_right_partition L w ps) in h at 1.
+    apply PermutationA_length in h.
+    apply (f_equal INR) in h.
+    repeat rewrite app_length in h.
+    repeat rewrite plus_INR in h.
+    apply PermutationA_length in h_neighbor.
+    apply (f_equal INR) in h_neighbor.
+    lra. }
+  apply (@weber_aligned_spec_weak L);auto.
+  - apply aligned_on_cons_iff;split;auto.
+    apply aligned_on_cons_iff in h_align;apply h_align.
+  - rewrite <- H, h_neighbor, <-heq_on.
+    apply (@weber_aligned_spec_weak L);auto.
+Qed.
+
+Lemma weber_aligned_non_occupied_neighbor_r:
+  forall [L : line] [ps : list R2] [w : R2],
+    ps <> nil ->
+    line_dir L =/= 0%VS ->
+    aligned_on L (w :: ps) ->
+    (^R (L ^on) w ps) = 0%R ->
+    Weber ps w ->
+    forall x,
+      line_contains L x ->
+      PermutationA equiv (L^left x ps) (L^left w ps) ->
+      (^R (L ^on) x ps) = 0%R ->
+      Weber ps x.
+Proof.
+  intros L ps w h_nonnil h_dir_nonnul h_align h_unoccupied h_weber_w x h_line_contains_x h_neighbor h_on_0.
+  apply (@weber_aligned_non_occupied_neighbor_l (line_reverse L) ps w);auto.
+  - intro abs.
+    inversion abs.
+    apply Ropp_eq_0_compat in H0.
+    rewrite Ropp_involutive in H0.
+    apply Ropp_eq_compat in H1. 
+    repeat rewrite Ropp_involutive in H1.
+    apply h_dir_nonnul.
+    cbn.
+    destruct L;cbn in *. 
+    destruct line_dir;cbn in *.
+    now subst.
+  - now rewrite <- aligned_on_reverse.
+  - now rewrite line_reverse_on.
+  - now rewrite <- line_reverse_contains.
+  - repeat rewrite line_reverse_right.
+    apply h_neighbor.
+  - now rewrite line_reverse_on.
+Qed.
+
+
+
+Lemma symetric_left_outside_seg: forall L (w1 w2 w:R2),
+    line_dir L =/= 0%VS ->
+    ((L ^P) w1 < (L ^P) w2)%R ->
+    line_contains L w1 ->
+    line_contains L w2 ->
+    w = (w1 + w1 - w2)%VS ->
+    ~ segment w1 w2 w.
+Proof.
+  intros L w1 w2 w h_dir_nonnull h_lt_w1_w2 h_contain_w1 h_contain_w2 heqw.
+  intro abs.
+  (* segment means w is a unary linear combination of w1 and w2
+           which is contradictory with w = (w1 + w1 - w2)%VS. *)
+  red in abs.
+  destruct abs as [t [ht1 ht2]].
+  change equiv with (@eq R2) in ht2.
+  rewrite heqw in ht2.
+  assert ((t * w1 + (2 - t) * w1 + (1 - t) * w2 + (t - 2) * w2) 
+          == ((t * w1 + (1 - t) * w2) + (2 - t) * w1 + (t - 2) * w2)
+         )%VS as heq_expr.
+  { setoid_rewrite <- add_assoc at 2.
+    setoid_rewrite <- add_comm at 3. 
+    setoid_rewrite add_assoc at 1.
+    reflexivity. }
+  assert ((w1 + w1 - w2) = (t * w1 + ((2 - t) * w1)) + (1 - t) * w2 + (t - 2) * w2)%VS as heq_expr_mul.
+  { setoid_rewrite <- add_assoc at 2.
+    repeat rewrite add_morph.
+    setoid_rewrite Rplus_comm at 1.
+    setoid_rewrite Rplus_assoc at 1.
+    rewrite Rplus_opp_l.
+    rewrite Rplus_0_r.
+    setoid_rewrite <- Rplus_assoc.
+    setoid_rewrite Rplus_assoc at 2.
+    rewrite Rplus_opp_l.
+    rewrite Rplus_0_r.
+    unfold R2.
+    assert ((1 + - (2)) = -1)%R as hR by lra.
+    rewrite hR.
+    setoid_rewrite minus_morph.
+    rewrite mul_1.
+    setoid_rewrite  <- (mul_1 w1) at 1 2.
+    rewrite add_morph.
+    assert (1+1 = 2)%R as hR1 by lra.
+    rewrite hR1;auto. }
+  change equiv with (@eq R2) in heq_expr.
+  rewrite heq_expr in heq_expr_mul.
+  rewrite <- ht2 in heq_expr_mul.
+  clear heq_expr.
+  assert (((2 - t) * w1 - (2 - t) * w2)%VS == 0%VS) as heq_exp_t.
+  { apply (R2plus_compat_eq_r (opp (w1 + w1 - w2)%VS)) in heq_expr_mul.
+    assert ((w1 + w1 - w2 - (w1 + w1 - w2))%VS == 0)%VS as h_eq_zero.
+    { apply R2sub_origin. reflexivity. }
+    setoid_rewrite h_eq_zero in heq_expr_mul.
+    rewrite add_comm in heq_expr_mul.
+    setoid_rewrite add_assoc at 1 in heq_expr_mul.
+    setoid_rewrite add_assoc at 1 in heq_expr_mul.
+    setoid_rewrite add_comm at 3 in heq_expr_mul.
+    setoid_rewrite h_eq_zero in heq_expr_mul.
+    rewrite add_origin_l in heq_expr_mul. 
+    rewrite <- (opp_opp ((t - 2) * w2))%VS in heq_expr_mul.
+    rewrite <- minus_morph in heq_expr_mul.
+    rewrite Ropp_minus_distr in heq_expr_mul.
+    symmetry.
+    apply heq_expr_mul. }
+  assert (w1 = w2) as heq_w1_w2.
+  { rewrite R2sub_origin in heq_exp_t.
+    apply mul_reg_l in heq_exp_t;auto.
+    lra. }
+  rewrite heq_w1_w2 in h_lt_w1_w2.
+  apply Rlt_irrefl with (1:=h_lt_w1_w2).
+Qed.
+
+Lemma symetric_left: forall L (w1 w2 w:R2),
+    line_dir L =/= 0%VS ->
+    ((L ^P) w1 < (L ^P) w2)%R ->
+    line_contains L w1 ->
+    line_contains L w2 ->
+    w = (w1 + w1 - w2)%VS ->
+    ((L^P w) < L^P w1)%R.
+Proof.
+  intros L w1 w2 w  h_dir_nonnull  h_lt_w1_w2 h_contain_w1 h_contain_w2 heqw.
+  red in h_contain_w1.
+  red in h_contain_w2.
+  cbn in *.
+  remember (fst w1) as x1 in *.
+  remember (fst w2) as x2 in *.
+  remember (snd w1) as y1 in *.
+  remember (snd w2) as y2 in *.
+  remember (line_dir L) as d.
+  remember (line_org L) as o.
+  remember (fst o) as xo.
+  remember (snd o) as yo.
+  remember (fst d) as xd.
+  remember (snd d) as yd.
+  remember ((sqrt (xd * xd + yd * yd))²)%R as nrm.
+  subst w.
+  cbn in *.
+  subst w1 w2.
+  cbn in *.
+  lra.
+Qed.
+
+Lemma line_on_zero L ps (w:R2):
+  aligned_on L ps ->
+  line_contains L w ->
+  ~InA equiv w ps -> 
+  (^R (L ^on) w ps) = 0%R.
+Proof.
+  intros h_align h_contain h_notInA.
+  apply INR_zero.
+  apply length_zero_iff_nil.
+  unfold line_on.
+
+  rewrite -> filter_nilA with (eqA := eq);auto; try typeclasses eauto.
+  intros x H. 
+  apply eqb_false_iff.
+  intro abs.
+  assert ((L ^P) w = (L ^P) x) as h_eq by auto.
+  apply (f_equal (line_proj_inv L)) in h_eq.
+  rewrite line_contains_def in h_eq;auto.
+  rewrite line_contains_def in h_eq;auto.
+  apply h_notInA.
+  subst.
+  assumption.
+Qed.
+
+Lemma weber_segment_InA_left ps w1 w2 L : 
+  ps <> nil ->
+  line_dir L =/= 0%VS ->
+  w1 =/= w2 ->
+  aligned_on L ps ->
+  line_contains L w1 ->
+  line_contains L w2 ->
+  Weber ps w1 ->
+  (L^P w1 < L^P w2)%R ->
+  ~ InA equiv w1 ps ->
+  (exists w, ~segment w1 w2 w /\ Weber ps w).
+Proof.
+  intros h_nonnil h_dir_nonnull h_neq h_align_ps h_contain_w1 h_contain_w2 h_weber_w1 h_lt_w1_w2 absInA. 
+  destruct (eqlistA_dec equiv_dec ((L^left w1 ps)) nil) as [heql_left_nil | hneql_left_nil]. 
+  - inversion heql_left_nil as [heq_left_nil | ].
+    alias (w1 + w1 - w2)%VS as w before h_neq.
+    assert (((L ^P) w < (L ^P) w1)%R) as hlt_proj.
+    { apply symetric_left with (w2 := w2);auto. }
+    exists w.
+    { split.
+      - apply symetric_left_outside_seg with (w1:=w1)(w2:=w2)(w:=w)(L:=L);auto.
+      - apply (@weber_aligned_non_occupied_neighbor_l L ps w1);auto.
+        + apply aligned_on_cons_iff;auto.
+        + apply line_on_zero;auto.
+        + subst w.
+          assert (w1 + (-1)%R * (w2 - w1) = (w1 + w1 - w2))%VS as heq.
+          { setoid_rewrite minus_morph.
+            setoid_rewrite <- mul_opp.
+            rewrite mul_1.
+            rewrite R2_opp_dist.
+            rewrite opp_opp.
+            setoid_rewrite add_comm at 2.
+            rewrite add_assoc.
+            reflexivity. }
+          rewrite <- heq.
+          apply line_contains_combine;auto. 
+        + rewrite (aggravate_right' L ps w1 w); try lra.
+          * reflexivity.
+          * rewrite <- heq_left_nil.
+            reflexivity.
+          * apply length_zero_iff_nil.
+            apply INR_zero.
+            apply line_on_zero;auto.
+        + specialize (left_lt L ps w w1 hlt_proj) as h.
+          rewrite <- heq_left_nil in h.
+          cbn [length INR] in h.
+          assert (0 <= (^R (L ^left) w ps))%R.
+          { apply pos_INR. }
+          assert (0 <= (^R (L ^on) w ps))%R.
+          { apply pos_INR. }
+          lra. }
+  - assert (((L ^left) w1 ps) <> nil) as hneq_left_nil.
+    { rewrite <- eqlistA_Leibniz.
+      assumption. }
+    clear hneql_left_nil.
+    alias (L^-P (L^max (L^left w1 ps))) as w after h_nonnil.
+    specialize (line_left_diff L w1 ps hneq_left_nil h_align_ps) as h.
+    specialize (line_contains_proj_inv L ((L ^max) ((L ^left) w1 ps)) h_dir_nonnull) as h'.
+    reremember in h'.
+    exists w;split.
+    + intro abs.
+      rewrite -> (segment_line L) in abs;auto.
+      destruct abs as [[hle_w1_w hle_w_w2] | [hle_w2_w hle_w_w1] ].
+      2:{ exfalso;lra. }
+      specialize (line_left_spec L w1 w ps) as h''.
+      assert (((L ^P) w < (L ^P) w1)%R).
+      { clear hle_w1_w hle_w_w2.
+        apply -> h''.
+        unremember.
+        apply line_iP_max_InA;auto.
+        apply aligned_on_inclA with (ps':=ps);auto.
+        apply line_left_inclA. }
+      exfalso.
+      lra.
+    + subst w.
+      apply weber_aligned_non_occupied_2;auto.
+      * apply aligned_on_cons_iff;auto.
+      * apply line_on_zero;auto.
+Qed.
+
+
+Lemma weber_segment_InA_1 ps w1 w2 L: 
+  ps <> nil ->
+  w1 =/= w2 ->
+  aligned_on L ps ->
+  (L^P w1 < L^P w2)%R ->
+  (forall w, Weber ps w <-> segment w1 w2 w) -> 
+  InA equiv w1 ps. 
+Proof.
+(* Proof sketch: suppose w1 is not occupied, then there is a point
+   close to w1 but outside the segment that is also a weber point
+  (since left, right and on would be the same). This would be a
+  contradiction. *)
+
+  intros h_nonnil h_neq h_align h_lt_w1_w2 h_forall. 
+
+  (* Frmo now on ps is supposed aligned. And w1 and w2 are also akigned with it. And they are weber points. *)
+  assert (Weber ps w1) as h_weber_w1.
+  { apply h_forall.
+    apply segment_start. }
+  assert (Weber ps w2) as h_weber_w2.
+  { apply h_forall.
+    apply segment_end. }
+  assert (line_contains L w1) as h_contain_w1.
+  { apply weber_aligned with ps;auto. }
+  assert (line_contains L w2) as h_contain_w2.
+  { apply weber_aligned with ps;auto. }
+  assert (line_dir L =/= 0%VS) as h_dir_nonnull.
+  { apply line_dir_nonnul with w1 w2;auto. }
+
+  apply decidable_not_not_iff.
+  { unfold Decidable.decidable.
+    destruct (InA_dec equiv_dec w1 ps);auto. }
+  intro absInA.
+  exfalso.
+  specialize (@weber_segment_InA_left ps w1 w2 L) as [w [hw1 hw2]];auto.
+  apply h_forall in hw2.
+  contradiction.
+Qed. 
+
+Lemma weber_segment_InA_2 ps w1 w2 L: 
+  ps <> nil ->
+  w1 =/= w2 ->
+  aligned_on L ps ->
+  (L^P w1 < L^P w2)%R ->
+  (forall w, Weber ps w <-> segment w1 w2 w) -> 
+  InA equiv w2 ps. 
+Proof.
+  intros h_nonnil h_neq HeqL h_lt_w1_w2 h_forall. 
+  remember (line_reverse L) as L'.
+  apply (@weber_segment_InA_1 ps w2 w1 L').
+  - assumption.
+  - now symmetry.
+  - subst.
+    now rewrite <- aligned_on_reverse.
+  - subst.
+    repeat rewrite line_reverse_proj.
+    lra.
+  - now setoid_rewrite segment_sym.
+Qed.
+
+
+
+Lemma weber_segment_InA_3 ps w1 w2 L: 
+  ps <> nil ->
+  w1 =/= w2 ->
+  aligned_on L ps ->
+  (L^P w2 < L^P w1)%R ->
+  (forall w, Weber ps w <-> segment w1 w2 w) -> 
+  InA equiv w2 ps. 
+Proof.
+  intros h_nonnil h_neq HeqL h_lt_w1_w2 h_forall. 
+  apply (@weber_segment_InA_1 ps w2 w1 L).
+  - assumption.
+  - now symmetry.
+  - assumption.
+  - lra.
+  - now setoid_rewrite segment_sym.
+Qed.
+
+
+Lemma weber_segment_InA_4 ps w1 w2 L: 
+  ps <> nil ->
+  w1 =/= w2 ->
+  aligned_on L ps ->
+  (L^P w2 < L^P w1)%R ->
+  (forall w, Weber ps w <-> segment w1 w2 w) -> 
+  InA equiv w1 ps. 
+Proof.
+  intros h_nonnil h_neq HeqL h_lt_w1_w2 h_forall. 
+  remember (line_reverse L) as L'.
+  apply (@weber_segment_InA_1 ps w1 w2 L').
+  - assumption.
+  - now symmetry.
+  - subst.
+    now rewrite <- aligned_on_reverse.
+  - subst.
+    repeat rewrite line_reverse_proj.
+    lra.
+  - assumption.
+Qed.
+
+
+
+Lemma weber_segment_InA ps w1 w2 : 
+  ps <> nil ->
+  w1 =/= w2 ->
+  (forall w, Weber ps w <-> segment w1 w2 w) -> 
+  InA equiv w1 ps /\ InA equiv w2 ps. 
+Proof.
+  intros h_nonnil h_neq h_forall.   
+  remember (line_calc ps) as L.
+  destruct (aligned_on_dec L ps) as [h_align_ps | h_nalign_ps].
+  all:swap 1 2. 
+  { (* If ps is not aligned, w1 and w2 can't be different *)
+    unremember in h_nalign_ps.
+    exfalso.
+    apply h_neq.
+    dup h_nalign_ps.
+    rewrite <- line_calc_spec in h_nalign_ps'.
+    specialize (@weber_Naligned_unique _ w1 h_nalign_ps') as h1.
+    specialize (@weber_Naligned_unique _ w2 h_nalign_ps') as h2.
+    rewrite h_forall in  h1,h2.
+    specialize (h1 (segment_start w1 w2)) as [h11 h12].
+    specialize (h2 (segment_end w1 w2)) as [h21 h22].
+    { auto. } }
+  assert (Weber ps w1) as h_weber_w1.
+  { apply h_forall.
+    apply segment_start. }
+  assert (Weber ps w2) as h_weber_w2.
+  { apply h_forall.
+    apply segment_end. }
+  assert (line_contains L w1) as h_contain_w1.
+  { apply weber_aligned with ps;auto. }
+  assert (line_contains L w2) as h_contain_w2.
+  { apply weber_aligned with ps;auto. }
+  assert (line_dir L =/= 0%VS) as h_dir_nonnull.
+  { apply line_dir_nonnul with w1 w2;auto. }
+  
+  destruct (Rlt_le_dec (L^P w1) (L^P w2)) as [h_lt_w1_w2|h_le_w2_w1].
+  { split.
+    - apply weber_segment_InA_1 with (w2:=w2)(L:=L);auto.
+    - apply weber_segment_InA_2 with (w1:=w1)(L:=L);auto. }
+  destruct (Rle_lt_or_eq_dec _ _ h_le_w2_w1) as [h_lt_w2_w1|h_eq_w2_w1].
+  { split.
+    - apply weber_segment_InA_4 with (w2:=w2)(L:=L);auto.
+    - apply weber_segment_InA_3 with (w1:=w1)(L:=L);auto. }
+  exfalso.
+  apply h_neq.
+  rewrite <- h_contain_w1.
+  rewrite <- h_contain_w2.
+  rewrite h_eq_w2_w1.
+  reflexivity.
+Qed.
+
+(*
+Lemma weber_segment_InA' ps w1 w2 : 
+  (forall w, Weber ps w <-> segment w1 w2 w)
+  <-> InA equiv w1 ps /\ InA equiv w2 ps /\ Weber ps w1 /\ Weber ps w2. 
+Proof. Admitted. 
+*)
+
+End WeberPoint.
+
+
+
+
+(*
+Lemma foo (x y m: R2): segment x y (middle x y) .
+Proof.
+  remember (x::y::nil) as ps.
+  remember (line_calc ps) as L.
+  assert (ps <> nil) as h_nonnil.
+  { now intro abs;subst. }
+  assert (aligned ps) as h_align_ps.
+  { unremember.
+    apply aligned_spec.
+    exists (y-x)%VS.
+    intros p H.
+    apply InA_singleton in H.
+    rewrite H.
+    exists 1%R.
+    rewrite mul_1.
+    symmetry.
+    apply add_sub. }
+  assert (line_contains L x) as h_line_x.
+  { apply aligned_on_InA_contains with (ps:=ps);auto.
+    + unremember.
+      auto.
+    + unremember.
+      unremember in h_align_ps.
+      apply line_calc_spec;auto. }
+  assert (line_contains L y) as h_line_y.
+  { apply aligned_on_InA_contains with (ps:=ps);auto.
+    + unremember.
+      auto.
+    + unremember.
+      unremember in h_align_ps.
+      apply line_calc_spec;auto. }
+  
+  dup h_align_ps.
+  rewrite (line_calc_spec ps) in h_align_ps'.
+  reremember in h_align_ps'.
+  specialize (line_iP_min_InA L ps h_nonnil h_align_ps') as h_min.
+  specialize (line_iP_max_InA L ps h_nonnil h_align_ps') as h_max.
+
+  apply (line_min_max_spec L (middle x y::ps)).
+  { intro abs.
+    discriminate. }
+  3:{ auto. }
+  - admit.
+  - red.
+    specialize (line_middle L (middle x y :: ps)) as h_mid.
+
+  
+  assert (((L ^min) ps) = (L^P) x /\ (L ^max) ps = (L ^P) y \/ ((L ^min) ps) = ((L^P) y) /\ (L ^max) ps = (L ^P) x).
+  { rewrite Heqps in h_min at 2.
+    rewrite Heqps in h_max at 2.
+    rewrite InA_cons in h_min,h_max.
+    rewrite InA_cons in h_min,h_max.
+    rewrite InA_nil in h_min,h_max.
+
+    assert (InA equiv x ps) as h_in_x.
+    { unremember; auto. }
+    assert (InA equiv y ps) as h_in_y.
+    { unremember; auto. }
+
+    specialize (line_bounds L ps x h_in_x) as h_inx.
+    specialize (line_bounds L ps y h_in_y) as h_iny.
+
+    destruct h_min as [h_min | [h_min | abs]]; try contradiction;
+      destruct h_max as [h_max | [h_max | abs]]; try contradiction.
+    - assert ((L ^-P) ((L ^min) ps) = x) as h_min'  by auto.
+      assert ((L ^-P) ((L ^max) ps) = x) as h_max' by auto.
+      rewrite <- h_max' in h_inx.
+      rewrite line_P_iP_max in h_inx.
+      assert ((L ^P) x =  (L ^P) y).
+      { subst x.
+
+        lra. }
+      left.
+      
+.
+
+    specialize (line_min_le_max L ps) as h_minmax.
+    admit. }
+  destruct H as [ [h_min_x h_max_y] | [h_min_y h_max_x]].
+  - rewrite h_min_x,h_max_y in h_mid.
+    red.
+    exists (1/2)%R.
+    split; try lra.
+    assert ((1 - 1 / 2 = 1/2)%R).
+    { lra. }
+    rewrite H.
+    transitivity (1 / 2 * (x + y))%VS.
+    2:{ apply mul_distr_add. }
+    reflexivity.
+  - rewrite h_min_y,h_max_x in h_mid.
+    red.
+    exists (1/2)%R.
+    split; try lra.
+    assert ((1 - 1 / 2 = 1/2)%R).
+    { lra. }
+    rewrite H.
+    transitivity (1 / 2 * (x + y))%VS.
+    2:{ apply mul_distr_add. }
+    reflexivity.
+ 
+  
+
+
+Lemma foo (x y m: R2): segment x y (middle x y) .
+Proof.
+  remember (x::y::nil) as ps.
+  remember (line_calc ps) as L.
+  assert (aligned ps).
+  { unremember.
+    apply aligned_spec.
+    exists (y-x)%VS.
+    intros p H.
+    apply InA_singleton in H.
+    rewrite H.
+    exists 1%R.
+    rewrite mul_1.
+    symmetry.
+    apply add_sub. }
+  assert (line_contains L x).
+  { apply aligned_on_InA_contains with (ps:=ps);auto.
+    + unremember.
+      auto.
+    + unremember.
+      unremember in H.
+      apply line_calc_spec;auto. }
+  assert (line_contains L y).
+  { apply aligned_on_InA_contains with (ps:=ps);auto.
+    + unremember.
+      auto.
+    + unremember.
+      unremember in H.
+      apply line_calc_spec;auto. }
+  assert (line_contains L (middle x y)).
+  { apply line_contains_middle;auto. }
+
+  apply (segment_line L);auto.
+  destruct (Rle_dec ((L ^P) x) ( (L ^P) y));[left|right].
+  - split.
+    + cbn.
+      unremember.
+      cbn.
+      destruct (equiv_dec x y);cbn.
+      * simpl.
+        repeat rewrite Rmult_0_r.
+        lra.
+      * repeat rewrite Rmult_plus_distr_r.
+        repeat rewrite Rmult_plus_distr_l.
+        repeat rewrite Rmult_plus_distr_r.
+        remember (fst x) as a.
+        remember (fst y) as b.
+        remember (snd x) as d.
+        remember (snd y) as e.
+        repeat rewrite Ropp_mult_distr_r_reverse.
+        repeat rewrite Ropp_mult_distr_l_reverse.
+        repeat rewrite  Ropp_involutive.
+        repeat rewrite Ropp_mult_distr_l_reverse.
+        assert (a * b + - (a * a) + (- (a * b) + a * a) = 0)%R as h_simpl_1.
+        { lra. }
+        assert (d * e + - (d * d) + (- (d * e) + d * d) = 0)%R as h_simpl_2.
+        { lra. }
+        rewrite h_simpl_1, h_simpl_2.
+        setoid_rewrite Rplus_comm at 8. 
+        setoid_rewrite Rplus_assoc at 3.
+        setoid_rewrite Rplus_comm at 6. 
+        setoid_rewrite Rplus_assoc at 3.
+        setoid_rewrite <- Rplus_assoc at 3.
+        assert ((1 / 2 * a * b) + - (1 / 2 * b * a) = 0)%R as h_simpl_3 .
+        { admit. 
+        }
+        setoid_rewrite h_simpl_3.
+        
+        assert ((- (1 / 2 * a * a) + 1 / 2 * b * b) + (- (a * b) + a * a) == 1/2 * ((a + -b)* (a+ -b)))%R.
+        { setoid_rewrite Rmult_plus_distr_r.
+          repeat setoid_rewrite Rmult_plus_distr_l.
+          (* setoid_rewrite Rmult_plus_distr_l. *)
+          (* setoid_rewrite Rmult_plus_distr_r. *)
+          admit. }
+        admit.
+    + admit.
+  - admit.
+Admitted.
+ 
+Lemma bar (x y m: R2): x<>y -> strict_segment x y (middle x y) .
+Proof.
+
+Admitted.
+*)
diff --git a/Core/Identifiers.v b/Core/Identifiers.v
index 6e0f1691f0338358c34ebdc91cc640819065f45f..8834c51cdfca790568322bba17a5de32f7485075 100644
--- a/Core/Identifiers.v
+++ b/Core/Identifiers.v
@@ -20,6 +20,17 @@ Require Import Pactole.Util.Enum.
 
 Set Implicit Arguments.
 
+(* Lemma nth_enum i m d : *)
+(*   forall Him : i < m, nth i (enum m) d = exist (fun x => x < m) i Him. *)
+(* Proof using . *)
+(* intros Him. apply eq_proj1, Nat.le_antisymm ; cbn. *)
+(* + apply lt_n_Sm_le, firstn_enum_spec. rewrite <-(firstn_skipn (S i)) at 1. *)
+(*   rewrite app_nth1 ; [apply nth_In|] ; rewrite firstn_length_le ; try rewrite enum_length ; lia.  *)
+(* + apply skipn_enum_spec. rewrite <-(firstn_skipn i) at 1. *)
+(*   rewrite app_nth2 ; [apply nth_In|] ; rewrite firstn_length_le by (rewrite enum_length ; lia) ; auto. *)
+(*   rewrite Nat.sub_diag, skipn_length, enum_length. lia. *)
+(* Qed. *)
+
 (** ** Byzantine Robots *)
 
 (** We have finitely many robots. Some are good, others are Byzantine.
diff --git a/Makefile.local b/Makefile.local
index 060b3d563d3f0159cf4d3dbcd34f7f7223355165..9a202a4a4869758778e4e5a6eabfcfde4d8c2014 100644
--- a/Makefile.local
+++ b/Makefile.local
@@ -5,13 +5,17 @@
 # called xxx_Assumptions.v. We create targets "vos-nocheck" and
 # "vok-nocheck" which ignore these files.
 
-.PHONY: vos-nocheck vok-nocheck
+.PHONY: vos-nocheck vok-nocheck vo-nocheck assumptions
 
 ## VOFILES is not defined before the "-include Makefile.local", so we
 ## use VFILES
-VOFILESNOASSUMPTION=$(filter-out %_Assumptions.v,$(VFILES))
-VOSFILESNOASSUMPTION=$(VOFILESNOASSUMPTION:%.v=%.vos)
-VOKFILESNOASSUMPTION=$(VOFILESNOASSUMPTION:%.v=%.vok)
+VFILESASSUMPTION=$(filter %_Assumptions.v,$(VFILES))
+VOFILESASSUMPTION=$(VFILESASSUMPTION:%.v=%.vo)
+
+VFILESNOASSUMPTION=$(filter-out %_Assumptions.v,$(VFILES))
+VOSFILESNOASSUMPTION=$(VFILESNOASSUMPTION:%.v=%.vos)
+VOKFILESNOASSUMPTION=$(VFILESNOASSUMPTION:%.v=%.vok)
+VOFILESNOASSUMPTION=$(VFILESNOASSUMPTION:%.v=%.vo)
 
 # special target for fast (and unsafe) compilation when developing
 # proofs. See also README.md.
@@ -22,6 +26,8 @@ vos-nocheck: $(VOSFILESNOASSUMPTION)
 # This replaces "make vok", it excludes xxx_Assumptions.v files
 vok-nocheck: $(VOKFILESNOASSUMPTION)
 
+vo-nocheck: $(VOFILESNOASSUMPTION)
+
 # This prints message when a "proof using" is missing. This cannot be
 # specified in _coqProject because spaces in args are not supported by
 # coq_makefile.
@@ -43,6 +49,8 @@ COQEXTRAFLAGS+= -w -deprecated-syntactic-definition
 
 core: Pactole_all.vo
 
+assumptions: $(VOFILESASSUMPTION)
+
 # make ../foo.check Ask to compile foo.v id necessary and then coqchk foo.vo
 %.check: %.vo
 	$(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS_NOML) $<
diff --git a/Models/GraphEquivalence.v b/Models/GraphEquivalence.v
index fde958ed810ad9358e679d6a47459994679f826f..27e7b5ef59928d00c5b054fb1f4f973997d4805e 100644
--- a/Models/GraphEquivalence.v
+++ b/Models/GraphEquivalence.v
@@ -1,4 +1,4 @@
- (**************************************************************************)
+(**************************************************************************)
 (*   Mechanised Framework for Local Interactions & Distributed Algorithms *)
 (*   C. Auger, P. Courtieu, L. Rieg, X. Urbain , R. Pelle                 *)
 (*   PACTOLE project                                                      *)
@@ -25,53 +25,6 @@ Require Import Pactole.Spaces.ThresholdIsomorphism.
 Require Import Pactole.Models.ContinuousGraph.
 Require Import Pactole.Models.DiscreteGraph.
 
-#[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.
 
@@ -222,6 +175,7 @@ Proof using .
     * f_equiv.
       rewrite Hconfig.
       now setoid_rewrite Hpt.
+    (* subst. do 3 (f_equiv; auto; []). now 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.
diff --git a/Observations/MultisetObservation.v b/Observations/MultisetObservation.v
index 8e9f7718eabed16205d72177676c3b432824db91..6c37adc7584678bf2e83491af6300d067112452e 100644
--- a/Observations/MultisetObservation.v
+++ b/Observations/MultisetObservation.v
@@ -13,7 +13,7 @@ Require Import Lia.
 Require Import SetoidList.
 Require Import SetoidDec.
 Require Import Pactole.Util.FMaps.FMapList.
-Require Import Pactole.Util.MMultiset.MMultisetWMap.
+Require Pactole.Util.MMultiset.MMultisetWMap.
 Require Export Pactole.Util.MMultiset.MMultisetInterface.
 Require Export Pactole.Util.MMultiset.MMultisetFacts.
 Require Export Pactole.Util.MMultiset.MMultisetExtraOps.
@@ -33,9 +33,8 @@ Context {loc : Type}.
 Context `{EqDec loc}.
 
 Existing Instance multiplicity_compat.
-Existing Instance FMOps_WMap.
-Existing Instance MakeFMultisetsFacts.
-
+Existing Instance MMultisetWMap.FMOps_WMap.
+Existing Instance MMultisetWMap.MakeFMultisetsFacts.
 
 (** **  Building multisets from lists  **)
 
@@ -173,7 +172,7 @@ Context `{St : State info}.
 Context `{Names}.
 Implicit Type config : configuration.
 
-Instance multiset_observation : Observation.
+#[export]Instance multiset_observation : Observation.
 simple refine {|
   observation := multiset location;
   
diff --git a/Pactole_all.v b/Pactole_all.v
index 78c1c8ef8701e73e91d51bbe0120ae94d86e25e1..c706df898629a558a166af4f463a9ff3a5852531 100644
--- a/Pactole_all.v
+++ b/Pactole_all.v
@@ -4,7 +4,7 @@ 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.GraphEquivalence.*)
 Require Import Pactole.Models.Flexible.
 Require Import Pactole.Models.Similarity.
 Require Import Pactole.Spaces.Isomorphism.
diff --git a/Spaces/R2.v b/Spaces/R2.v
index 44ec3a41fe08bee00075e221847c927600c263ee..b4dbb662b68cbacf546cde94d4f40be680da3bd0 100644
--- a/Spaces/R2.v
+++ b/Spaces/R2.v
@@ -26,6 +26,7 @@ Import List Permutation SetoidList.
 Require Import SetoidDec.
 Require Import FunInd.
 Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
 Require Import Pactole.Util.Bijection.
 Require Export Pactole.Spaces.EuclideanSpace.
 Require Import Pactole.Spaces.Similarity.
@@ -1290,6 +1291,26 @@ repeat rewrite mul_distr_add, ?mul_morph. repeat rewrite <- add_assoc. do 2 f_eq
 rewrite add_morph. ring_simplify (k + (1 - k)). now rewrite mul_1.
 Qed.
 
+
+Lemma straight_path_similarity (sim : similarity R2) x y (r : ratio) :
+  straight_path (sim x) (sim y) r == sim (straight_path x y r).
+Proof using .
+cbn -[mul opp RealVectorSpace.add]. 
+rewrite sim_add, <-RealVectorSpace.add_assoc. f_equal.
+rewrite sim_mul. unfold Rminus. rewrite <-add_morph, mul_1.
+rewrite (RealVectorSpace.add_comm (sim 0%VS) _), <-2 RealVectorSpace.add_assoc.
+rewrite add_opp, add_origin_r. 
+rewrite minus_morph, <-mul_opp, <-mul_distr_add. f_equal.
+rewrite sim_add, sim_opp, <-2 RealVectorSpace.add_assoc. f_equal.
+change 2%R with (1 + 1)%R. rewrite <-add_morph, mul_1. 
+rewrite RealVectorSpace.add_comm, 2 RealVectorSpace.add_assoc. 
+rewrite <-add_origin_l. f_equal.
++ rewrite <-(RealVectorSpace.add_assoc (- sim 0)%VS _), (RealVectorSpace.add_comm (- sim 0)%VS (sim 0)%VS), add_opp.
+  rewrite add_origin_r, RealVectorSpace.add_comm, add_opp.
+  reflexivity.
++ rewrite add_origin_l. reflexivity. 
+Qed.
+
 (** To prove that isobarycenter are invariant by similarities, we handle translations first. *)
 Lemma isobarycenter_translation_morph : forall t l, l <> nil ->
   isobarycenter (map (translation t) l) == translation t (isobarycenter l).
@@ -2571,7 +2592,7 @@ destruct (Exists_dec (fun x => x <> pt1 /\ on_circle (SEC (pt1 :: l)) x = true))
   assert (Hnew : enclosing_circle {| center := c'; radius := r' |} (pt1 :: l)).
   { intros pt Hin. simpl center. simpl radius. destruct Hin.
     * subst pt. unfold c'. rewrite R2dist_ref_0.
-      replace (pt1 - (c + (1 - r'/r) * (pt1 - c)))%VS with (r'/ r * pt1 - (r' / r * c))%VS
+      replace (pt1 - (c + (1 - r' / r) * (pt1 - c)))%VS with (r'/ r * pt1 - (r' / r * c))%VS
         by (destruct pt1, c; cbn; f_equal; ring).
       rewrite <- R2dist_ref_0. rewrite dist_homothecy.
       rewrite on_circle_true_iff in Hon1. unfold c. rewrite Hon1.
diff --git a/Spaces/RealNormedSpace.v b/Spaces/RealNormedSpace.v
index 3547ef0f6c3843fb765d5bebda2b73411c8d737e..429c0ee7a159ffbcc76cb742869eed78ba75dee7 100644
--- a/Spaces/RealNormedSpace.v
+++ b/Spaces/RealNormedSpace.v
@@ -15,6 +15,7 @@ Require Import SetoidList.
 Require Import SetoidDec.
 Require Import RelationPairs.
 Require Import Pactole.Util.Coqlib.
+Require Import Pactole.Util.Ratio.
 Require Export Pactole.Spaces.RealMetricSpace.
 Open Scope R_scope.
 Open Scope VectorSpace_scope.
@@ -421,4 +422,36 @@ Section BarycenterResults.
   
 End BarycenterResults.
 
+Section StraightPathResults.
+Context {T : Type}.
+Context `{RealNormedSpace T}.
+Implicit Types (x y : T).
+
+Lemma straight_path_dist_start x y (r : ratio) : 
+  dist x (straight_path x y r) == (r * dist x y)%R.
+Proof using .
+cbn -[norm mul opp add equiv].
+rewrite opp_distr_add, add_assoc, add_opp, add_origin_l.
+rewrite <-mul_opp, opp_distr_add, opp_opp, add_comm.
+rewrite norm_mul, Rabs_pos_eq ; [reflexivity | apply ratio_bounds].
+Qed.
+
+Lemma straight_path_dist_end x y (r : ratio) : 
+  dist y (straight_path x y r) == ((1 - r) * dist x y)%R.
+Proof using .  
+cbn -[norm mul opp add equiv].
+assert (Hr1 : (y - (x + r * (y - x)) == (r - 1) * (x - y))%VS).
+{
+  unfold Rminus. rewrite <-add_morph, minus_morph, mul_1, 2 opp_distr_add, opp_opp.
+  rewrite (add_comm _ (-x + y)%VS), add_assoc. apply add_compat.
+  + now rewrite add_comm.
+  + rewrite <-mul_opp. apply mul_compat ; auto. now rewrite opp_distr_add, opp_opp, add_comm.  
+}
+rewrite Hr1, norm_mul. f_equiv. rewrite Rabs_left1.
++ now rewrite Ropp_minus_distr.
++ generalize (ratio_bounds r). lra.
+Qed. 
+
+End StraightPathResults.
+
 Arguments unitary {T%_type} {_} {_} {_} {_} u%_VS.
diff --git a/Spaces/RealVectorSpace.v b/Spaces/RealVectorSpace.v
index 53c8c624460554a082c3a9bfa6666ff3aaf26c94..19098d1d64b23a498e32dad197909cb8494fedd2 100644
--- a/Spaces/RealVectorSpace.v
+++ b/Spaces/RealVectorSpace.v
@@ -15,7 +15,7 @@ Require Import RelationPairs.
 Require Import Pactole.Util.Coqlib.
 
 
-Typeclasses eauto := (bfs).
+(* Typeclasses eauto := (bfs). *)
 Remove Hints eq_setoid : typeclass_instances.
 
 
@@ -181,9 +181,48 @@ intros pt1 pt2 Heq pt1' pt2' Heq' x. simpl.
 now apply add_compat, mul_compat, add_compat, opp_compat.
 Qed.
 
+
+Lemma mul_eq0_iff `{RealVectorSpace} (k : R) x : (k * x == 0)%VS <-> (k == 0)%R \/ (x == 0)%VS.
+Proof.
+split ; intros Heq0.
++ case (k =?= 0%R) as [Hk | Hk] ; [now left | right].
+  apply mul_reg_l with k ; [intuition|].
+  rewrite mul_origin. exact Heq0.
++ destruct Heq0 as [Hk | Hx].
+  - now rewrite Hk, mul_0.
+  - now rewrite Hx, mul_origin.
+Qed.
+
+Lemma straight_path_0 `{RealVectorSpace} p p' : straight_path p p' ratio_0 == p%VS. 
+Proof using . cbn -[equiv mul opp add]. now rewrite mul_0, add_origin_r. Qed. 
+
 Lemma straight_path_1 `{RealVectorSpace} : forall pt pt', straight_path pt pt' ratio_1 == pt'.
 Proof using . intros. simpl. now rewrite mul_1, add_assoc, (add_comm pt), <- add_assoc, add_opp, add_origin. Qed.
 
+Lemma straight_path_same `{RealVectorSpace} x r : straight_path x x r == x.
+Proof using . cbn -[equiv mul opp add]. now rewrite add_opp, mul_origin, add_origin_r. Qed. 
+
+Lemma straight_path_end `{RealVectorSpace} x y r : 
+  straight_path x y r == y <-> (x == y \/ proj_ratio r == 1%R).
+Proof using .
+split. 
++ intros Hend. cbn -[mul opp add] in Hend.
+  assert (Hr1 : ((r - 1) * (y - x) == 0)%VS).
+  {
+    unfold Rminus. rewrite <-add_morph, minus_morph, mul_1, opp_distr_add, opp_opp.
+    rewrite (add_comm _ x), add_assoc, (add_comm _ x).
+    now rewrite Hend, add_opp.
+  }  
+  rewrite mul_eq0_iff in Hr1. case Hr1 as [H1 | H2].
+  - right. cbn in H1 |- *. lra.
+  - left. rewrite <-(add_opp x) in H2. apply add_reg_r in H2. now symmetry.
++ intros [H1 | H2].
+  - now rewrite H1, straight_path_same.
+  - cbn -[mul opp add equiv]. rewrite H2, mul_1.
+    now rewrite (add_comm y), add_assoc, add_opp, add_origin_l.
+Qed.
+
+
 (** We can simplify this definition in the local frame as we start from the origin. *)
 Definition local_straight_path {T} `{RVS : RealVectorSpace T} (pt : T) : path T.
 Proof.
diff --git a/Util/ListComplements.v b/Util/ListComplements.v
index 4783c9573bb166489c28f0e935c856839fbb66e6..91ad6a8e16b1ffcb1f8f56b13ccf3a8d866a0fe6 100644
--- a/Util/ListComplements.v
+++ b/Util/ListComplements.v
@@ -25,10 +25,13 @@ Require Export SetoidPermutation.
 Require Import Sorting.Permutation.
 Require Import Bool.
 Require Import Lia PeanoNat.
+Require Import Coq.Arith.Wf_nat.
 Require Import Psatz.
 Require Import SetoidDec.
 Require Import Pactole.Util.Preliminary.
+Require Import Pactole.Util.SetoidDefs.
 Require Import Pactole.Util.NumberComplements.
+Require Import Arith.
 
 (* this will become non default soon. *)
 Ltac Tauto.intuition_solver ::= auto with *.
@@ -41,8 +44,38 @@ Arguments PermutationA {A}%_type eqA%_signature l1%_list l2%_list.
 (** *  Some results on Lists  *)
 (******************************)
 
+(* TODO: remove this lemma and rely on the stdlib version
+ * once we drop support for Coq < v8.14 *)
+Lemma Forall_cons_iff A (P : A -> Prop) a l :
+  Forall P (a :: l) <-> P a /\ Forall P l.
+Proof. 
+split.
++ intros HF. inv HF. intuition.
++ intuition.
+Qed.
+
+(* TODO: remove this lemma and rely on the stdlib version 
+ * once we drop support for Coq < v8.14 *)
+Lemma Forall_map A B (f : A -> B) P l :
+  Forall P (map f l) <-> Forall (fun x => P (f x)) l.
+Proof.
+induction l as [|x l IH] ; cbn.
++ intuition. 
++ rewrite 2 Forall_cons_iff, <-IH. reflexivity.
+Qed.
+
+(* TODO: remove this lemma and rely on the stdlib version 
+ * once we drop support for Coq < v8.14 *)
+Lemma Exists_map A B (f : A -> B) P l :
+  List.Exists P (List.map f l) <-> List.Exists (fun x => P (f x)) l.
+Proof.
+induction l as [|x l IH] ; cbn.
++ split ; intros H ; inv H. 
++ rewrite 2 Exists_cons, <-IH. reflexivity.
+Qed. 
+
 Section List_results.
-Context (A B : Type).
+Context (A B C : Type).
 Context (eqA eqA' : relation A) (eqB : relation B).
 Context (HeqA : Equivalence eqA) (HeqB : Equivalence eqB).
 
@@ -103,6 +136,21 @@ intros d x l Hin. induction l as [| e l].
     exists (S n), y. repeat split; cbn; auto; lia.
 Qed.
 
+Lemma nth_InA i l d : Reflexive eqA -> 
+  i < length l -> InA eqA (nth i l d) l.
+Proof using .
+intros Hrefl Hi. induction l using rev_ind. 
++ cbn in Hi. lia.
++ rewrite app_length in Hi. cbn in Hi.
+  assert (Hil : i < length l \/ i >= length l) by lia.
+  case Hil as [Hi_lt | Hi_ge].
+  - rewrite InA_app_iff ; left.
+    rewrite app_nth1 by auto. now apply IHl.
+  - assert (Hi_eq : i = length l) by lia.
+    rewrite Hi_eq, nth_middle, InA_app_iff ; right.
+    now apply InA_cons_hd.
+Qed.
+
 (* Already exists as [List.In_nth] but with a reversed argument order
 Lemma In_nth : forall d x (l : list A), In x l -> exists n, (n < length l)%nat /\ nth n l d = x.
 Proof.
@@ -242,6 +290,48 @@ intros x y Hxy l1 l2 Hl. split; intro H; eapply InA_eqA; try eassumption.
   now rewrite Hl.
 Qed.
 
+Lemma combine_nil_iff (l : list A) (l' : list B) :
+  combine l l' = nil <-> (l = nil \/ l' = nil).
+Proof using . 
+split.
++ intros Hcom. destruct l as [|x l] ; destruct l' as [|x' l'] ; intuition.
+  discriminate Hcom.
++ intros [-> | ->] ; [|rewrite combine_nil] ; reflexivity.
+Qed.
+
+Lemma combine_map l l' (f : A -> C) (f' : B -> C) : 
+  combine (map f l) (map f' l') = map (fun '(x, x') => (f x, f' x')) (combine l l').
+Proof using .
+generalize l'. clear l'. induction l as [| x l IH] ; intros [|x' l'] ; cbn ; try reflexivity.
+f_equal. apply IH.
+Qed.
+
+Lemma in_combine_id (a b : A) l : 
+  In (a, b) (combine l l) <-> (a = b /\ In a l).
+Proof using . 
+induction l as [|x l IH].
++ cbn. intuition.
++ cbn. rewrite IH. split.
+  - intros [Hxx | [-> Hin]].
+    * inv Hxx. intuition.
+    * intuition.
+  - intros [-> [-> | Hin]] ; intuition.
+Qed.
+
+Lemma in_combine_sym (a : A) (b : B) l1 l2 : 
+  In (a, b) (combine l1 l2) <-> In (b, a) (combine l2 l1).
+Proof using .
+cut (forall (A B : Type) (x : A) (y : B) l l', 
+  In (x, y) (combine l l') -> In (y, x) (combine l' l)).
+{ intros H. split ; apply H. }
+clear a b l1 l2. intros A' B' a b l1.
+induction l1 as [|x1 l1 IH] ; intros l2 Hin ; [now cbn in Hin|].
+case l2 as [|x2 l2] ; [now cbn in Hin|]. cbn in Hin |- *. 
+case Hin as [Heq | Hin].
++ inv Heq. now left.
++ right. apply IH. exact Hin.
+Qed.
+
 
 (** ***  Results about [map]  **)
 
@@ -592,6 +682,29 @@ intros. rewrite <- negb_involutive, existsb_forallb. f_equal.
   - simpl. apply IHl.
 Qed.
 
+Lemma Forall2_Forall (R : A -> B -> Prop) l l' : length l = length l' -> 
+  (Forall2 R l l' <-> Forall (fun '(x, y) => R x y) (combine l l')).
+Proof using . 
+intros Hlen. split.
++ intros Hforall2. induction Hforall2 as [|x x' l l' Hx Hl IH] ; constructor ; auto.
++ intros Hforall. remember (combine l l') as c.
+  generalize dependent l'. generalize dependent l. generalize dependent c.
+  induction c as [|a c IH] ; intros Hforall l l' Hlen Hcom.
+  - symmetry in Hcom. rewrite combine_nil_iff in Hcom. 
+    destruct Hcom as [-> | ->] ; cbn in Hlen.
+    * case l' as [|? ?] ; [constructor | discriminate].
+    * case l as [|? ?] ; [constructor | discriminate].
+  - destruct a as [y y']. 
+    destruct l as [|x l] ; [discriminate|]. 
+    destruct l' as [|x' l'] ; [discriminate|].
+    cbn in Hcom. inv Hcom. rewrite Forall_cons_iff in Hforall. constructor ; intuition.
+Qed.
+
+Lemma Forall2_length (R : A -> B -> Prop) l1 l2 : 
+  Forall2 R l1 l2 -> length l1 = length l2.
+Proof using . intros HF2. induction HF2 ; cbn ; lia. Qed.
+
+
 (** ***  Results about [PermutationA]  **)
 
 Global Instance InA_perm_compat : Proper (eqA ==> PermutationA eqA ==> iff) (InA eqA).
@@ -659,6 +772,28 @@ Proof using . intros x y ? l l' Hl. subst. split; apply Permutation_in; assumpti
 Global Instance In_permA_compat : Proper (eq ==> @PermutationA A eq ==> iff) (@In A).
 Proof using . repeat intro. rewrite PermutationA_Leibniz in *. now apply In_perm_compat. Qed.
 
+Lemma flat_map_compat_equiv : 
+Proper ((eqA ==> PermutationA eqB) ==> eqlistA eqA ==> PermutationA eqB) (@flat_map A B).
+Proof using HeqA HeqB. 
+intros f f' Hff' l l' Hll'. elim Hll'.
++ cbn. now reflexivity.
++ intros x x' t t' Hxx' Htt' IH. cbn. now f_equiv ; auto.
+Qed.
+
+Global Instance flat_map_compat_perm : 
+Proper ((eqA ==> PermutationA eqB) ==> PermutationA eqA ==> PermutationA eqB) (@flat_map A B).
+Proof using HeqA HeqB. 
+intros f f' Hff' l l' Hll'. elim Hll'.
++ simpl. now reflexivity.
++ intros x x' t t' Hxx' Htt' IH. cbn. rewrite <-IH. f_equiv. now apply Hff'.
++ intros x y t. cbn. repeat rewrite app_assoc. f_equiv.
+- rewrite PermutationA_app_comm by auto. f_equiv ; now apply Hff'.
+- apply flat_map_compat_equiv ; auto. reflexivity.
++ intros t t' t'' _ IH1 _ IH2. rewrite IH1, <-IH2. 
+  apply flat_map_compat_equiv ; try reflexivity. 
+  intros x x' Hxx'. symmetry. now apply Hff'.
+Qed.
+
 Lemma Forall_Perm_trans : forall (l1 l2 : list A) (P Q : A -> Prop),
   (eq ==> iff)%signature P Q -> Permutation l1 l2 -> List.Forall P l1 -> List.Forall Q l2.
 Proof using .
@@ -884,6 +1019,34 @@ intro l1. induction l1 as [| e l1]; intro l2.
   + right. abstract (intro Habs; apply He; rewrite <- Habs; now left).
 Defined.
 
+
+Lemma nth_firstn2: forall (l : list A) k n df, k < n -> nth k l df = List.nth k (firstn n l) df.
+Proof using Type.
+  intros l k n df H. 
+  specialize firstn_le_length with (n:=n)(l:=l) as h.
+  destruct (le_lt_dec (length l) n).
+  - specialize firstn_all2 with (1:= l0) as h3. 
+    now rewrite h3.
+  - specialize firstn_length_le with (n:=n)(l:=l) as h4.
+    rewrite <- (firstn_skipn n l) at 1.
+    rewrite app_nth1 at 1.
+    + reflexivity.
+    + rewrite h4;auto with arith.
+Qed.
+
+Lemma combine_nth2:
+  forall (l : list A) (l' : list B) (k : nat) (x : A) (y : B),
+    k < length l -> k < length l' -> nth k (combine l l') (x, y) = (nth k l x, nth k l' y).
+Proof using Type.
+  intros until l; induction l; intro l'; destruct l';intros n x y h h'; try now inversion h; try now inversion h'.
+  destruct n; simpl in *.
+  + reflexivity.
+  + apply IHl.
+    * auto with arith.
+    * auto with arith.
+Qed.
+
+
 Lemma PermutationA_3_swap : forall l e1 e2 e3 pt1 pt2,
   complement eqA pt1 pt2 ->
   InA eqA pt1 (e1 :: e2 :: e3 :: l) ->
@@ -989,6 +1152,140 @@ Qed.
 
 End List_results.
 
+Section ListCombineResult.
+Context (A B C : Type).
+Context (SA : Setoid A)(SB : Setoid B).
+
+Lemma inA_combine_r: forall (l : list A) (l' : list B) (x : A) (y : B), InA equiv (x, y) (combine l l') -> InA equiv y l'.
+Proof using Type.
+  intros until l'. 
+  revert l.
+  induction l'.
+  - intros l x y hin.
+    rewrite combine_nil in hin.
+    apply InA_nil in hin.
+    contradiction.
+  - intros l x y hin.
+    destruct l.
+    + cbn [combine] in hin.
+      apply InA_nil in hin.
+      contradiction.
+    + cbn [combine] in hin.
+      inversion hin as [ ? ? heq | ? ? h];subst.
+      * destruct heq as [hx hy]; cbn in hx,hy.
+        rewrite hy.
+        now constructor 1.
+      * constructor 2.
+        eapply IHl';eauto.
+Qed.
+
+Lemma inA_combine_l: forall (l : list A) (l' : list B) (x : A) (y : B), InA equiv (x, y) (combine l l') -> InA equiv x l.
+Proof using Type.
+  intros until l.
+  induction l.
+  - intros l' x y hin.
+    cbn [combine] in hin.
+      apply InA_nil in hin.
+      contradiction.
+  - intros l' x y hin.
+    destruct l'.
+    + rewrite combine_nil in hin.
+      apply InA_nil in hin.
+      contradiction.
+    + cbn [combine] in hin.
+      inversion hin as [ ? ? heq | ? ? h];subst.
+      * destruct heq as [hx hy]; cbn in hx,hy.
+        rewrite hx.
+        now constructor 1.
+      * constructor 2.
+        eapply IHl;eauto.
+Qed.
+
+Lemma inA_combine_id: forall (l : list A) (a b:A),
+  InA equiv (a, b) (combine l l) <-> (a == b /\ InA equiv a l).
+Proof using . 
+induction l as [|x l IH];intros.
++ cbn. setoid_rewrite InA_nil. intuition.
++ cbn. setoid_rewrite InA_cons.
+  setoid_rewrite IH. split.
+  - intros [Hxx | [-> Hin]].
+    * inv Hxx;cbn in *;intuition. etransitivity;eauto.
+    * intuition.
+  - intros [-> [-> | Hin]] ; intuition.
+Qed.
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine_same_lgth (l : list A) x x' (l' : list B) dfA dfB : 
+  length l = length l' ->
+  InA equiv (x, x') (combine l l') -> 
+  exists n, (n < length l /\ x == List.nth n l dfA /\ x' == nth n l' dfB).
+Proof using Type.
+  intros Hlth Hin. 
+  specialize InA_nth with (1 := Hin) as h.
+  specialize (h (dfA,dfB)).
+  destruct h as [k [y [hlgth [heqxy hnth]]]].
+  exists k;split;[|split].
+  - rewrite combine_length in hlgth.
+    rewrite <- Hlth in hlgth.
+    rewrite Nat.min_id in hlgth.
+    assumption.
+  - (* clear heqxy hnth. *)
+    (* cbn in heq'. *)
+    (* destruct heq' as [heq'1 heq'2]. *)
+    (* rewrite <- heq'1. *)
+    rewrite combine_nth in hnth.
+    + rewrite <- hnth in heqxy.
+      cbn in heqxy.
+      destruct heqxy.
+      auto.
+    + assumption.
+  - rewrite combine_nth in hnth.
+    + rewrite <- hnth in heqxy.
+      cbn in heqxy.
+      destruct heqxy.
+      auto.
+    + assumption.
+Qed.
+
+
+(* This would have been much more pleasant to do with mathcomp's tuples. *)
+Lemma config_list_InA_combine_same_lgth2 (l : list A) x x' (l' : list B) dfA dfB : 
+  InA equiv (x, x') (combine l l') -> 
+  exists n, (n < length l /\ x == List.nth n l dfA /\ x' == nth n l' dfB).
+Proof using Type.
+  intros Hin. 
+  specialize InA_nth with (1 := Hin) as h.
+  specialize (h (dfA,dfB)).
+  destruct h as [k [y [hlgth [heqxy hnth]]]].
+  exists k;split;[|split].
+  - rewrite combine_length in hlgth.
+    eapply Nat.min_glb_l.
+    eassumption.
+  - rewrite combine_nth2 in hnth.
+    + rewrite <- hnth in heqxy.
+      now inversion heqxy.
+    + apply Nat.lt_le_trans with (length (combine l l')) ;auto.
+      rewrite combine_length.
+      apply Nat.le_min_l.
+    + apply Nat.lt_le_trans with (length (combine l l')) ;auto.
+      rewrite combine_length.
+      apply Nat.le_min_r.
+  - rewrite combine_nth2 in hnth.
+    + rewrite <- hnth in heqxy.
+      cbn in heqxy.
+      now inversion hnth.
+    + apply Nat.lt_le_trans with (length (combine l l')) ;auto.
+      rewrite combine_length.
+      apply Nat.le_min_l.
+    + apply Nat.lt_le_trans with (length (combine l l')) ;auto.
+      rewrite combine_length.
+      apply Nat.le_min_r.
+Qed.
+
+End ListCombineResult.
+
+
+
 Lemma map_alls {A B} : forall (f : A -> B) pt n, map f (alls pt n) = alls (f pt) n.
 Proof using . intros f pt n. induction n. reflexivity. simpl. now rewrite IHn. Qed.
 
@@ -1390,6 +1687,8 @@ Section CountA.
 Context (A : Type).
 Context (eqA : relation A).
 Context (HeqA : Equivalence eqA).
+(* Context [eqA : relation A].
+   Context {HeqA : Equivalence eqA}. *)
 Hypothesis eq_dec : forall x y : A, {eqA x y} + {~eqA x y}.
 
 Fixpoint countA_occ (x : A) l :=
@@ -1422,6 +1721,7 @@ intros a b Hab l. induction l as [| x l]; intros [| x' l'] Hl.
   - 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.
@@ -1507,7 +1807,6 @@ Corollary PermutationA_count_split : forall x l,
   PermutationA eqA l (alls x (countA_occ x l) ++ removeA eq_dec x l).
 Proof using HeqA. intros. now rewrite <- countA_occ_spec. Qed.
 
-
 Lemma count_filter_length: forall x l,
     countA_occ x l = length (List.filter (fun a => if eq_dec a x then true else false) l).
 Proof using.
@@ -1520,8 +1819,71 @@ Proof using.
     reflexivity.
 Qed.
 
+Lemma countA_occ_removeA_same x l :
+  countA_occ x (removeA eq_dec x l) = 0.
+Proof using HeqA. 
+induction l as [|y l IH].
++ reflexivity.
++ cbn. destruct_match.
+  - now rewrite IH.
+  - cbn. destruct_match ; [intuition | now rewrite IH].
+Qed.    
+
+Lemma countA_occ_removeA_other x y l :
+  ~ eqA x y -> countA_occ x (removeA eq_dec y l) = countA_occ x l.
+Proof using HeqA.
+intros Hxy. induction l as [|z l IH].
++ reflexivity.
++ cbn. repeat destruct_match.
+  - rewrite <-H in H0. intuition.
+  - now rewrite IH.
+  - rewrite H0. cbn. destruct_match ; [|intuition]. now rewrite IH.
+  - cbn. destruct_match ; [intuition|]. now rewrite IH.
+Qed.
+
+Theorem PermutationA_countA_occ l l' :
+  PermutationA eqA l l' <-> 
+  forall x, countA_occ x l == countA_occ x l'.
+Proof using HeqA. 
+split.
++ intros Hperm x. elim Hperm.
+  - now reflexivity.
+  - intros x1 x2 l1 l2 Hx Hl IH. cbn. 
+    repeat destruct_match ; try (now rewrite IH).
+    * rewrite <-H, Hx in * ; now intuition.
+    * rewrite <-H0, Hx in * ; now intuition.
+  - intros a b t. cbn. repeat destruct_match ; reflexivity.
+  - intros l1 l2 l3 _ H1 _ H2. now rewrite H1, <-H2.
++ intros Hcount. remember (length l) as m. generalize l l' Heqm Hcount.
+  pattern m. apply (well_founded_ind lt_wf). clear m l l' Heqm Hcount.
+  intros m IH [|x l] l' Hm Hcount.
+  -  cbn in *. destruct l' as [|y tl'] ; [now apply permA_nil|].
+    specialize (Hcount y). revert Hcount ; cbn. destruct_match ; [|intuition]. discriminate.
+  - rewrite (PermutationA_count_split x l).
+    rewrite (PermutationA_count_split x l').
+    rewrite app_comm_cons. f_equiv.
+    * apply eqlistA_PermutationA. rewrite <-Hcount. cbn. destruct_match ; [|intuition].
+      cbn. reflexivity.
+    * eapply IH ; [|reflexivity|].
+      ++apply (Nat.le_lt_trans _ (length l)).
+        --clear Hm Hcount. induction l as [|? l IHl] ; cbn ; try reflexivity.
+          destruct_match ; cbn ; revert IHl ; lia.
+        --rewrite Hm. cbn. apply Nat.lt_succ_diag_r.
+      ++intros y. case (eq_dec x y) as [Hxy | Hxy]. 
+        --rewrite <-Hxy. repeat rewrite countA_occ_removeA_same. reflexivity.
+        --repeat rewrite countA_occ_removeA_other by intuition.
+          rewrite <-Hcount. cbn. destruct_match ; intuition.
+Qed.
+
 End CountA.
 
+(* Global Instance countA_occ_compat_setoid A {S:Setoid A} {D:EqDec S}:  *)
+(*   Proper (@equiv _ S ==> PermutationA equiv ==> eq) (countA_occ D). *)
+(* Proof using . *)
+(*   apply countA_occ_compat. *)
+(*   autoclass. *)
+(* Qed. *)
+
 Lemma countA_occ_map {A B} (eqA : relation A) (eqB : relation B) Aeq_dec Beq_dec :
   forall (f : A -> B), Proper (eqA ==> eqB) f -> injective eqA eqB f ->
   forall x l, countA_occ eqB Beq_dec (f x) (map f l) = countA_occ eqA Aeq_dec x l.
diff --git a/Util/NumberComplements.v b/Util/NumberComplements.v
index ae022b589c53129db4c6cf69afbc98e03c6bcc2c..5addd52bbba185fb3b03b5a7c80923942fad783d 100644
--- a/Util/NumberComplements.v
+++ b/Util/NumberComplements.v
@@ -610,3 +610,64 @@ Proof using ltc_ll_ul ltc_lr_ur.
 Qed.
 
 End lt_pred_mul_ul_ur.
+
+
+Lemma INR_injection: forall n1 n2,
+    INR n1 = INR n2 -> 
+    n1 = n2.
+Proof.
+  induction n1 using nat_ind2;intros.
+  - destruct n2;auto.
+    destruct n2.
+    + exfalso.
+      cbn in H.
+      lra.
+    + exfalso.
+      repeat rewrite S_INR in H.
+      assert (INR n2 >= 0)%R.
+      { apply Rle_ge, pos_INR. }
+      cbn in H.
+      lra.
+  - destruct n2.
+    + exfalso.
+      cbn in H.
+      lra.
+    + destruct n2;auto.
+      repeat rewrite S_INR in H.
+      assert (INR n2 >= 0)%R.
+      { apply Rle_ge, pos_INR. }
+      cbn in H.
+      lra.
+  - assert (INR n1 >= 0)%R.
+    { apply Rle_ge, pos_INR. }
+    repeat rewrite S_INR in H.
+    destruct n2.
+    { exfalso.
+      cbn in H.
+      lra. }
+    destruct n2.
+    + exfalso.
+      cbn in H.
+      lra.
+    + rewrite (IHn1 n2);auto.
+      repeat rewrite S_INR in H.
+      lra.
+Qed.
+
+Lemma INR_add_le: forall x y: nat, x < y -> (((INR x) + 1 <= (INR y))%R).
+Proof.
+  intros x y h_Rlt.
+  change (IZR (Zpos xH)) with (INR 1).
+  rewrite <- plus_INR.
+  apply le_INR.
+  lia.
+Qed.
+
+Lemma INR_add_le_INR: forall x y : nat,
+    ((INR x) < (INR y))%R ->
+    (((INR x) + 1 <= (INR y))%R) .
+Proof.
+  intros x y h_Rlt.
+  apply INR_add_le.
+  now apply INR_lt.
+Qed.
diff --git a/Util/Preliminary.v b/Util/Preliminary.v
index 40b02a496a3b6ee83ee55586626b16ed2282b8e1..91855d36490ac9edf2ffe86e754180f3f886ff66 100644
--- a/Util/Preliminary.v
+++ b/Util/Preliminary.v
@@ -76,6 +76,89 @@ Ltac revert_all :=
     | |- ?B => repeat revert_one B
   end.
 
+
+(* Dealing with hyps generated by "remember" *)
+Ltac reremember_in h :=
+  repeat match goal with
+      H: ?X = _ |- _ =>
+        is_var X;
+        rewrite <- H in h
+    end.
+
+Ltac reremember_ :=
+  repeat match goal with
+      H: ?X = _ |- _ =>
+        is_var X;
+        rewrite <- H
+    end.
+
+Ltac unremember_in h :=
+  repeat match goal with
+      H: ?X = _ |- _ =>
+        is_var X;
+        rewrite H in h
+    end.
+
+Ltac unremember_ :=
+  repeat match goal with
+      H: ?X = _ |- _ =>
+        is_var X;
+        rewrite H
+    end.
+
+Ltac remember_all_move_after t name ath :=
+  remember t as name in *;
+  match goal with
+    H: name = t |- _ =>
+      move H after ath;
+      move name at top           
+  end.
+
+Ltac remember_all_move_before t name ath :=
+  remember t as name in *;
+  match goal with
+    H: name = t |- _ =>
+      move H before ath;
+      move name at top           
+  end.
+
+Ltac remember_in_move_after t name inlh ath :=
+  remember t as name in inlh ;
+  match goal with
+    H: name = t |- _ =>
+      move H after ath;
+      move name at top           
+  end.
+
+Ltac remember_in_move_before t name inlh ath :=
+  remember t as name in inlh ;
+  match goal with
+    H: name = t |- _ =>
+      move H before ath;
+      move name at top           
+  end.
+
+(* alias is like remember but allows to move around the var and hyp
+   created . *)
+Tactic Notation "alias" constr(t) "as" ident(name) "after" hyp(h) :=
+  remember_all_move_after t name h.
+Tactic Notation "alias" constr(t) "as" ident(name) "before" hyp(h) :=
+  remember_all_move_before t name h.
+
+Tactic Notation "alias" constr(t) "as" ident(name) "in" hyp_list(lh) "after" hyp(h) :=
+  remember_in_move_after t name lh h.
+Tactic Notation "alias" constr(t) "as" ident(name) "in" hyp_list(lh) "before" hyp(h) :=
+  remember_in_move_before t name lh h.
+
+(* unremember is equivalent to applying subst to all hyps of the form
+id = term (not the symmetric though, as would subst so). *)
+Tactic Notation "unremember" := unremember_.
+Tactic Notation "unremember" "in" hyp(h) := unremember_in h.
+(* reremember tries to identify new occurrences of remembered
+   expressions and replace them by the name. *)
+Tactic Notation "reremember" := reremember_.
+Tactic Notation "reremember" "in" hyp(h) := reremember_in h.
+
 Definition injective {A B : Type} (eqA : relation A) (eqB : relation B) (f : A -> B) : Prop :=
   forall x y, eqB (f x) (f y) -> eqA x y.
 
diff --git a/Util/Ratio.v b/Util/Ratio.v
index 496d3492c9ba76aab8379d5422cb7373733fd8f7..f59fd8917c9b52254ee03f8569257e5099a1b944 100644
--- a/Util/Ratio.v
+++ b/Util/Ratio.v
@@ -81,6 +81,20 @@ Defined.
 Global Instance add_ratio_compat : Proper (equiv ==> equiv ==> equiv) add_ratio.
 Proof. intros [] [] ? [] [] ?. unfold add_ratio. simpl in *. subst. destruct_match; reflexivity. Qed.
 
+Lemma add_ratio_ge_left (r1 r2 : ratio) : (r1 <= add_ratio r1 r2)%R. 
+Proof.
+unfold add_ratio. destruct_match.
++ apply ratio_bounds.
++ cbn. generalize (ratio_bounds r2). lra.
+Qed.
+
+Lemma add_ratio_ge_right (r1 r2 : ratio) : (r2 <= add_ratio r1 r2)%R. 
+Proof.
+unfold add_ratio. destruct_match.
++ apply ratio_bounds.
++ cbn. generalize (ratio_bounds r1). lra.
+Qed. 
+
 (** A strict ratio is a [ratio] that is neither [0] nor [1]. *)
 Definition strict_ratio := {x : R | 0 < x < 1}%R.
 
diff --git a/Util/SetoidDefs.v b/Util/SetoidDefs.v
index 00f85237bba26440716690c679bc772c18a92c6f..2a3bcc20e452198b7dfae48b305511c7b83adad1 100644
--- a/Util/SetoidDefs.v
+++ b/Util/SetoidDefs.v
@@ -50,6 +50,10 @@ 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 auto with *. Qed.
 
+Lemma equiv_dec_other {T U : Type} {S : Setoid T} {E : EqDec S}:
+  forall (e e': T) (A B : U), e=/= e' -> (if equiv_dec e e' then A else B) = B.
+Proof using . intros. destruct_match; intuition auto with *. Qed.
+
 Definition equiv_decb_refl [T] [S : Setoid T] (E : EqDec S) : forall x : T, x ==b x = true :=
   fun x => equiv_dec_refl x true false.
 
diff --git a/_CoqProject b/_CoqProject
index 5a8175545707fa670391b671187e977c2bf090bb..ba2193a3cd3a3aff7f6ba2c5ff6186c2f93e627f 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -20,7 +20,7 @@
 Pactole_all.v
 
 ## External libraries
-#Util/FSets/OrderedType.v
+# # Util/FSets/OrderedType.v
 Util/FSets/FSetInterface.v
 Util/FSets/FSetFacts.v
 Util/FSets/FSetList.v
@@ -118,6 +118,19 @@ CaseStudies/Gathering/InR2/Algorithm_withLight.v
 CaseStudies/Gathering/InR2/Algorithm_withLight_Assumptions.v
 
 
+## Case Study : Weber
+CaseStudies/Gathering/InR2/Weber/Utils.v
+CaseStudies/Gathering/InR2/Weber/Segment.v
+CaseStudies/Gathering/InR2/Weber/Line.v
+CaseStudies/Gathering/InR2/Weber/Weber_point.v
+CaseStudies/Gathering/InR2/Weber/Align_rigid_ssync.v
+CaseStudies/Gathering/InR2/Weber/Align_flex_ssync.v
+CaseStudies/Gathering/InR2/Weber/Align_flex_async.v
+CaseStudies/Gathering/InR2/Weber/Gather_flex_async.v
+# Seg Fault?
+CaseStudies/Gathering/InR2/Weber/Gather_flex_async_Assumptions.v
+
+
 ## Case Study: Ring Exploration
 CaseStudies/Exploration/ExplorationDefs.v
 CaseStudies/Exploration/ImpossibilityKDividesN.v
diff --git a/minipactole/minipactole.v b/minipactole/minipactole.v
index 15fca270a42984f0cb57c5530e0cbf2e0afe95b9..80c277b77fb8da2bdcb14a925464d8e5c289fda3 100644
--- a/minipactole/minipactole.v
+++ b/minipactole/minipactole.v
@@ -4,6 +4,7 @@ Require Import SetoidDec SetoidList SetoidClass.
 Require  Pactole.Util.Stream.
 Require Import Pactole.Util.SetoidDefs Pactole.Util.NumberComplements.
 
+Unset Automatic Proposition Inductives.
 Open Scope R_scope.
 
 Inductive ident:Type := A | B. (* Nous avons seulement 2 robots. *)
diff --git a/q.log b/q.log
new file mode 100644
index 0000000000000000000000000000000000000000..aff7df568952a15062ad3943ce7f3426f368985c
--- /dev/null
+++ b/q.log
@@ -0,0 +1,24 @@
+This is pdfTeX, Version 3.141592653-2.6-1.40.23 (TeX Live 2021) (preloaded format=pdflatex 2021.7.29)  22 AUG 2022 12:06
+entering extended mode
+ restricted \write18 enabled.
+ %&-line parsing enabled.
+**weber_main.tex
+(/usr/local/texlive/2021/texmf-dist/tex/latex/tools/q.tex
+LaTeX2e <2021-06-01> patch level 1
+L3 programming layer <2021-07-12> File ignored
+)
+! Emergency stop.
+<*> weber_main.tex
+                  
+*** (job aborted, no legal \end found)
+
+ 
+Here is how much of TeX's memory you used:
+ 18 strings out of 478542
+ 520 string characters out of 5850401
+ 291678 words of memory out of 5000000
+ 18055 multiletter control sequences out of 15000+600000
+ 403430 words of font info for 27 fonts, out of 8000000 for 9000
+ 1141 hyphenation exceptions out of 8191
+ 13i,0n,12p,94b,10s stack positions out of 5000i,500n,10000p,200000b,80000s
+!  ==> Fatal error occurred, no output PDF file produced!