diff --git a/FEM/Algebra/AffineSpace_new.v b/FEM/Algebra/AffineSpace_new.v
new file mode 100644
index 0000000000000000000000000000000000000000..4525b81e497d89e0d6476ead4c0cfa8430d8eba0
--- /dev/null
+++ b/FEM/Algebra/AffineSpace_new.v
@@ -0,0 +1,3388 @@
+(**
+This file is part of the Elfic library
+
+Copyright (C) Boldo, Clément, Martin, Mayero, Mouhcine
+
+This library is free software; you can redistribute it and/or
+modify it under the terms of the GNU Lesser General Public
+License as published by the Free Software Foundation; either
+version 3 of the License, or (at your option) any later version.
+
+This library is distributed in the hope that it will be useful,
+but WITHOUT ANY WARRANTY; without even the implied warranty of
+MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+COPYING file for more details.
+*)
+
+From FEM.Compl Require Import stdlib.
+
+Set Warnings "-notation-overridden".
+From FEM.Compl Require Import ssr.
+
+(* Next import provides Reals, Coquelicot.Hierarchy, functions to module spaces,
+ and linear mappings. *)
+From LM Require Import linear_map.
+Close Scope R_scope.
+Set Warnings "notation-overridden".
+
+From Lebesgue Require Import Subset Subset_dec Subset_any Function.
+
+From FEM Require Import Compl Sub_struct Monoid_compl Group_compl Ring_compl.
+From FEM Require Import ModuleSpace_compl ModuleSpace_R_compl.
+
+
+(** Results needing a commutative Ring are only stated for
+ the ring of real numbers R_Ring. *)
+
+
+Local Open Scope Monoid_scope.
+Local Open Scope Group_scope.
+Local Open Scope Ring_scope.
+
+
+Module AffineSpace.
+
+(* FIXME: essai avec un existentiel fort (ax2). *)
+Record mixin_of {K : Ring} (V : ModuleSpace K) (E : Type) := Mixin {
+  ax0 : E;
+  vect : E -> E -> V;
+  ax1 : forall A B C, vect B A + vect A C = vect B C;
+(*  ax2 : forall A u, exists! B, vect A B = u;*)
+  ax2 : forall A u, { B | unique (fun X => vect A X = u) B };
+}.
+
+Notation class_of := mixin_of (only parsing).
+
+Section ClassDef.
+
+Context {K : Ring}.
+Variable V : ModuleSpace K.
+
+Structure type := Pack { sort; _ : class_of V sort; _ : Type; }.
+Local Coercion sort : type >-> Sortclass.
+Definition class (cT : type) := let: Pack _ c _ := cT return class_of V cT in c.
+
+End ClassDef.
+
+Module Exports.
+
+Coercion sort : type >-> Sortclass.
+Notation AffineSpace := type.
+
+End Exports.
+
+End AffineSpace.
+
+
+Export AffineSpace.Exports.
+
+
+Section AffineSpace_Def1.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Definition vect := AffineSpace.vect _ _ (AffineSpace.class V E).
+
+End AffineSpace_Def1.
+
+
+Declare Scope AffineSpace_scope.
+Delimit Scope AffineSpace_scope with AS.
+Notation "A --> B" := (vect A B) (at level 55) : AffineSpace_scope.
+
+Local Open Scope AffineSpace_scope.
+
+
+Section AffineSpace_Facts0.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Variable E : AffineSpace V.
+
+Definition point_of_as : E := AffineSpace.ax0 _ _ (AffineSpace.class V _).
+
+Lemma inhabited_as : inhabited E.
+Proof. apply (inhabits point_of_as). Qed.
+
+Lemma inhabited_fct_as : forall {T : Type}, inhabited (T -> E).
+Proof. intros; apply fun_to_nonempty_compat, inhabited_as. Qed.
+
+End AffineSpace_Facts0.
+
+
+Section AffineSpace_Facts1.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma vect_chasles : forall (A B C : E), (B --> A) + (A --> C) = B --> C.
+Proof. intros; apply AffineSpace.ax1. Qed.
+
+Lemma vect_l_bij_EX :
+  forall (A : E) u, { B | unique (fun X => A --> X = u) B }.
+Proof. apply AffineSpace.ax2. Qed.
+
+Lemma vect_l_bij_ex : forall (A : E) u, exists! B, A --> B = u.
+Proof. intros A u; destruct (vect_l_bij_EX A u) as [B HB]; exists B; easy. Qed.
+
+(* FIXME: attention, bijective est un inductif, ~ un existentiel faible... *)
+Lemma vect_l_bij : forall (A : E), bijective (vect A).
+Proof. intros; apply bij_ex_uniq_rev, vect_l_bij_ex. Qed.
+
+Lemma vect_zero : forall (A : E), A --> A = 0.
+Proof.
+intros A; apply (plus_reg_l (A --> A)); rewrite vect_chasles plus_zero_r; easy.
+Qed.
+
+Lemma vect_zero_equiv : forall (A B : E), A --> B = 0 <-> B = A.
+Proof.
+intros A B; split; intros H.
+(* *)
+apply eq_sym; destruct (vect_l_bij_EX A 0) as [C [_ HC]].
+rewrite -(HC A (vect_zero _)); apply HC; easy.
+(* *)
+rewrite H; apply vect_zero.
+Qed.
+
+Lemma vect_sym : forall (A B : E), A --> B = - (B --> A).
+Proof.
+intros; apply plus_is_zero_l_equiv; rewrite vect_chasles; apply vect_zero.
+Qed.
+
+Lemma vect_r_bij_EX :
+  forall (B : E) u, { A | unique (fun X => X --> B = u) A }.
+Proof.
+intros B u; destruct (vect_l_bij_EX B (- u)) as [A [HA1 HA2]].
+exists A; split. rewrite vect_sym HA1 opp_opp; easy.
+move=> A' /opp_eq; rewrite -vect_sym; apply HA2.
+Qed.
+
+Lemma vect_r_bij_ex : forall (B : E) u, exists! A, A --> B = u.
+Proof. intros B u; destruct (vect_r_bij_EX B u) as [A HA]; exists A; easy. Qed.
+
+Lemma vect_r_bij : forall (B : E), bijective (vect^~ B).
+Proof. intros; apply bij_ex_uniq_rev, vect_r_bij_ex. Qed.
+
+Lemma vect_l_eq : forall (A B1 B2 : E), B1 = B2 -> A --> B1 = A --> B2.
+Proof. intros; subst; easy. Qed.
+
+Lemma vect_r_eq : forall (B A1 A2 : E), A1 = A2 -> A1 --> B = A2 --> B.
+Proof. intros; subst; easy. Qed.
+
+Lemma vect_l_inj : forall (A B1 B2 : E), A --> B1 = A --> B2 -> B1 = B2.
+Proof. move=>>; eapply (bij_inj (vect_l_bij _)). Qed.
+
+Lemma vect_r_inj : forall (B A1 A2 : E), A1 --> B = A2 --> B -> A1 = A2.
+Proof. move=>>; eapply (bij_inj (vect_r_bij _)). Qed.
+
+Lemma vect_l_inj_equiv : forall (A B1 B2 : E), A --> B1 = A --> B2 <-> B1 = B2.
+Proof. intros; split. apply vect_l_inj. apply vect_l_eq. Qed.
+
+Lemma vect_r_inj_equiv : forall (B A1 A2 : E), A1 --> B = A2 --> B <-> A1 = A2.
+Proof. intros; split. apply vect_r_inj. apply vect_r_eq. Qed.
+
+End AffineSpace_Facts1.
+
+
+Section AffineSpace_Facts2.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Variable E : AffineSpace V.
+
+Lemma as_w_zero_struct_ms : zero_struct V -> is_unit_type E.
+Proof.
+destruct (inhabited_as E) as [A]; intros HV; exists A; intros B.
+apply (vect_l_inj A); rewrite !(HV (_ --> _)); easy.
+Qed.
+
+Lemma as_w_zero_struct_r : zero_struct K -> is_unit_type E.
+Proof. intros; apply as_w_zero_struct_ms, ms_w_zero_struct_r; easy. Qed.
+
+End AffineSpace_Facts2.
+
+
+Section AffineSpace_Def2.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+(* FIXME: comment se passer de ex_EX puisque vect_l_bij est un inductif ? *)
+Lemma transl_EX :
+  { f | forall (A : E), cancel (vect A) (f A) /\ cancel (f A) (vect A) }.
+Proof.
+assert (H : forall A : E, exists ! f, cancel (vect A) f /\ cancel f (vect A)).
+  intros A; destruct (vect_l_bij A) as [f Hf1 Hf2]; exists f; split; [easy |].
+  intros f' [Hf'1 Hf'2]; apply fun_ext; intros u.
+  apply (vect_l_inj A); rewrite Hf2 Hf'2; easy.
+apply ex_EX; destruct (unique_choice _ H) as [f Hf]; exists f; easy.
+Qed.
+
+Definition transl : E -> V -> E := proj1_sig transl_EX.
+
+Lemma transl_correct_l : forall (A : E), cancel (vect A) (transl A).
+                      (* forall (A B : E), A +++ (A --> B) = B. *)
+Proof. apply (proj2_sig transl_EX). Qed.
+
+Lemma transl_correct_r : forall (A : E), cancel (transl A) (vect A).
+                      (* forall (A : E) u, A --> (A +++ u) = u. *)
+Proof. apply (proj2_sig transl_EX). Qed.
+
+End AffineSpace_Def2.
+
+
+Notation "A +++ u" := (transl A u) (at level 50) : AffineSpace_scope.
+Notation "A ++- u" := (transl A (- u)) (at level 50) : AffineSpace_scope.
+
+Local Open Scope AffineSpace_scope.
+
+
+Section AffineSpace_Facts3.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma vect_l_full :
+  forall (PE : E -> Prop) (O : E), full PE -> full (image (vect O) PE).
+Proof. intros PE O HPE u; rewrite -(transl_correct_r O u); easy. Qed.
+
+Lemma vect_l_fullset : forall (O : E), image (vect O) fullset = fullset.
+Proof. intros; rewrite -full_fullset; apply vect_l_full; easy. Qed.
+
+Lemma transl_l_full : forall PV (O : E), full PV -> full (image (transl O) PV).
+Proof. intros PV O HPV A; rewrite -(transl_correct_l O A); easy. Qed.
+
+Lemma transl_l_fullset : forall (O : E), image (transl O) fullset = fullset.
+Proof. intros; rewrite -full_fullset; apply transl_l_full; easy. Qed.
+
+Lemma transl_assoc : forall (A : E) u v, (A +++ u) +++ v = A +++ (u + v).
+Proof.
+intros A u v; rewrite -(transl_correct_l A ((A +++ u) +++ v))
+    -(vect_chasles (A +++ u)) 2!transl_correct_r; easy.
+Qed.
+
+Lemma transl_zero : forall (A : E), A +++ 0 = A.
+Proof. intros; erewrite <- vect_zero; apply transl_correct_l. Qed.
+
+Lemma transl_zero_equiv : forall (A : E) u, A +++ u = A <-> u = 0.
+Proof. intros; rewrite -vect_zero_equiv transl_correct_r; easy. Qed.
+
+Lemma transl_l_eq : forall (A : E) u1 u2, u1 = u2 -> A +++ u1 = A +++ u2.
+Proof. intros; subst; easy. Qed.
+
+Lemma transl_r_eq : forall u (A1 A2 : E), A1 = A2 -> A1 +++ u = A2 +++ u.
+Proof. intros; subst; easy. Qed.
+
+Lemma transl_l_inj : forall (A : E) u1 u2, A +++ u1 = A +++ u2 -> u1 = u2.
+Proof.
+intros A u1 u2 H; rewrite -(transl_correct_r A u1) -(transl_correct_r A u2).
+apply vect_l_eq; easy.
+Qed.
+
+Lemma transl_r_inj : forall u (A1 A2 : E), A1 +++ u = A2 +++ u -> A1 = A2.
+Proof.
+intros u A1 A2 H.
+rewrite -(transl_zero A1) -(transl_zero A2) -(plus_opp_r u) -2!transl_assoc.
+apply transl_r_eq; easy.
+Qed.
+
+Lemma transl_l_inj_equiv : forall (A : E) u1 u2, A +++ u1 = A +++ u2 <-> u1 = u2.
+Proof. intros; split. apply transl_l_inj. apply transl_l_eq. Qed.
+
+Lemma transl_r_inj_equiv : forall u (A1 A2 : E), A1 +++ u = A2 +++ u <-> A1 = A2.
+Proof. intros; split. apply transl_r_inj. apply transl_r_eq. Qed.
+
+Lemma transl_opp_equiv : forall (A B : E) u, A +++ u = B <-> A = B ++- u.
+Proof.
+intros A B u.
+rewrite -(transl_r_inj_equiv (- u)) transl_assoc plus_opp_r transl_zero; easy.
+Qed.
+
+Lemma transl_vect_equiv : forall (A B : E) u, B = A +++ u <-> u = A --> B.
+Proof. intros; erewrite <- vect_l_inj_equiv, transl_correct_r; easy. Qed.
+
+Lemma transl_l_bij : forall (A : E), bijective (transl A).
+Proof.
+intros A; apply (bij_can_bij (vect_l_bij A)); apply transl_correct_l.
+Qed.
+
+Lemma transl_l_bij_ex : forall (A B : E), exists! u, A +++ u = B.
+Proof. intros; apply bij_ex_uniq, transl_l_bij. Qed.
+
+Lemma transl_r_bij_ex : forall u (B : E), exists! A, A +++ u = B.
+Proof.
+intros u B; exists (B ++- u); split.
+apply transl_opp_equiv; easy.
+intros A; rewrite transl_opp_equiv; easy.
+Qed.
+
+Lemma transl_r_bij : forall u, bijective (transl^~ u : E -> E).
+Proof. intros; apply bij_ex_uniq_rev, transl_r_bij_ex. Qed.
+
+Lemma vect_transl : forall (O : E) u v, (O +++ u) --> (O +++ v) = v - u.
+Proof.
+intros O u v; rewrite -{1}(plus_minus_l v u) plus_comm -transl_assoc.
+apply transl_correct_r.
+Qed.
+
+Lemma vect_transl_assoc : forall (O A : E) u, O --> (A +++ u) = (O --> A) + u.
+Proof. intros; rewrite -(vect_chasles A) transl_correct_r; easy. Qed.
+
+Lemma vect_transl_plus :
+  forall (O : E) u v,
+    O --> (O +++ (u + v)) = (O --> (O +++ u)) + (O --> (O +++ v)).
+Proof. intros; rewrite !transl_correct_r; easy. Qed.
+
+Lemma vect_transl_scal :
+  forall (O : E) a u, O --> (O +++ (scal a u)) = scal a (O --> (O +++ u)).
+Proof. intros; rewrite !transl_correct_r; easy. Qed.
+
+End AffineSpace_Facts3.
+
+
+Section Fct_AffineSpace.
+
+Context {T : Type}.
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Definition fct_point_of_as : T -> E := fun _ => point_of_as E.
+
+Definition fct_vect (fA fB : T -> E) : T -> V := fun x => fA x --> fB x.
+
+Lemma fct_vect_chasles :
+  forall (fA fB fC : T -> E), fct_vect fB fA + fct_vect fA fC = fct_vect fB fC.
+Proof. intros; apply fun_ext; intro; apply vect_chasles. Qed.
+
+(* FIXME: il a fallu ajouter un ex_EX ici ! *)
+Lemma fct_vect_l_bij_EX :
+  forall (fA : T -> E) (fu : T -> V),
+    { fB | unique (fun fX => fct_vect fA fX = fu) fB }.
+Proof.
+intros fA fu; apply ex_EX.
+destruct (unique_choice (fun x B => fA x --> B = fu x)) as [fB HfB].
+intros x; apply (vect_l_bij_ex (fA x) (fu x)).
+exists fB; split. apply fun_ext_equiv; easy.
+move=> fB' /fun_ext_equiv HfB'; apply fun_ext; intros x.
+apply (vect_l_inj (fA x)); rewrite HfB -HfB'; easy.
+Qed.
+
+Definition fct_AffineSpace_mixin :=
+  AffineSpace.Mixin _ _ _ fct_point_of_as _ fct_vect_chasles fct_vect_l_bij_EX.
+
+Canonical Structure fct_AffineSpace :=
+  AffineSpace.Pack _ _ fct_AffineSpace_mixin (T -> E).
+
+Lemma fct_vect_eq : forall (fA fB : T -> E) t, (fA --> fB) t = fA t --> fB t.
+Proof. easy. Qed.
+
+Lemma fct_transl_eq :
+  forall (fA : T -> E) (fu : T -> V) t, (fA +++ fu) t = fA t +++ fu t.
+Proof.
+intros fA fu t; apply (vect_l_inj (fA t)).
+rewrite -fct_vect_eq 2!transl_correct_r; easy.
+Qed.
+
+End Fct_AffineSpace.
+
+
+Section prod_AffineSpace.
+
+Context {K : Ring}.
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Definition prod_point_of_as : E1 * E2 := (point_of_as E1, point_of_as E2).
+
+Definition prod_vect (A B : E1 * E2) : V1 * V2 := (A.1 --> B.1, A.2 --> B.2).
+
+Lemma prod_vect_chasles :
+  forall A B C, prod_vect B A + prod_vect A C = prod_vect B C.
+Proof. intros; apply (f_equal2 Datatypes.pair); apply vect_chasles. Qed.
+
+Lemma prod_vect_l_bij_EX :
+  forall A u, { B | unique (fun X => prod_vect A X = u) B }.
+Proof.
+intros [A1 A2] [u1 u2]; destruct (vect_l_bij_EX A1 u1) as [B1 [HB1a HB1b]],
+    (vect_l_bij_EX A2 u2) as [B2 [HB2a HB2b]].
+exists (B1, B2); split. rewrite -HB1a -HB2a; easy.
+move=> [C1 C2] /pair_equal_spec /= [HC1 HC2]; apply (f_equal2 Datatypes.pair);
+    [apply HB1b | apply HB2b]; easy.
+Qed.
+
+Definition prod_AffineSpace_mixin :=
+  AffineSpace.Mixin _ _ _ prod_point_of_as _
+    prod_vect_chasles prod_vect_l_bij_EX.
+
+Canonical Structure prod_AffineSpace :=
+  AffineSpace.Pack _ _ prod_AffineSpace_mixin (E1 * E2).
+
+Lemma prod_vect_eq : forall (A B : E1 * E2), A --> B = (A.1 --> B.1, A.2 --> B.2).
+Proof. easy. Qed.
+
+Lemma prod_transl_eq :
+  forall (A : E1 * E2) (u : V1 * V2), A +++ u = (A.1 +++ u.1, A.2 +++ u.2).
+Proof.
+intros A [u1 u2]; apply (vect_l_inj A).
+rewrite (prod_vect_eq _ (A.1 +++ u1, A.2 +++ u2)) 3!transl_correct_r; easy.
+Qed.
+
+End prod_AffineSpace.
+
+
+Section ModuleSpace_Is_AffineSpace.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+
+(* This could be abstracted. *)
+Definition ms_point_of_as : V := 0.
+
+Definition ms_vect (A B : V) : V := B - A.
+
+Lemma ms_vect_chasles : forall A B C, ms_vect B A + ms_vect A C = ms_vect B C.
+Proof. intros; rewrite plus_comm -minus_trans; easy. Qed.
+
+Lemma ms_vect_l_bij_EX :
+  forall A u, { B | unique (fun X => ms_vect A X = u) B }.
+Proof.
+intros A u; exists (A + u); split; [apply: minus_plus_l |].
+intros v Hv; rewrite -Hv plus_comm; apply: plus_minus_l.
+Qed.
+
+Definition ModuleSpace_AffineSpace_mixin :=
+  AffineSpace.Mixin _ _ _ ms_point_of_as _ ms_vect_chasles ms_vect_l_bij_EX.
+
+Canonical Structure ModuleSpace_AffineSpace :=
+  AffineSpace.Pack _ _ ModuleSpace_AffineSpace_mixin V.
+
+Lemma ms_vect_eq : forall (A B : V), A --> B = B - A.
+Proof. easy. Qed.
+
+Lemma ms_transl_eq : forall (A u : V), A +++ u = A + u.
+Proof.
+intros A u; apply (vect_l_inj A); unfold vect at 2; simpl; unfold ms_vect.
+rewrite transl_correct_r minus_plus_l; easy.
+Qed.
+
+End ModuleSpace_Is_AffineSpace.
+
+
+(* FIXME: TO BE REMOVED FOR PUBLICATION!
+ Unfortunately, the following seems to generate problems when mixing
+ R_AffineSpace and ModuleSpace_AffineSpace stuff...
+Section R_AffineSpace.
+
+(* Don't use any ModuleSpace stuff to define the canonical structure R_AffineSpace.
+ Otherwise, it is redundant with the previous ModuleSpace_AffineSpace, and simply ignored! *)
+
+Definition R_point_of_as : R := 0%R.
+
+Definition R_vect (A B : R) : R := (B - A)%R.
+
+Lemma R_vect_chasles : forall A B C, (R_vect B A + R_vect A C)%R = R_vect B C.
+Proof. unfold R_vect; intros; field. Qed.
+
+Lemma R_vect_l_bij_EX : forall A u, { B | unique (fun X => R_vect A X = u) B }.
+Proof.
+unfold R_vect; intros A u; exists (A + u)%R; split; try field.
+intros B HB; subst; field.
+Qed.
+
+Definition R_AffineSpace_mixin :=
+  AffineSpace.Mixin _ _ _ R_point_of_as _ R_vect_chasles R_vect_l_bij_EX.
+
+Canonical Structure R_AffineSpace :=
+  AffineSpace.Pack _ _ R_AffineSpace_mixin R.
+
+End R_AffineSpace.
+*)
+
+
+(* FIXME: TO BE REMOVED FOR PUBLICATION!
+ Just for checking it is possible, but it generates typing problems in the sequel...
+ Note that this was attempted before the introduction of monoids!
+ This should be updated...
+Section AffineSpace_Is_ModuleSpace.
+
+(* From Gostiaux T4, Th 1.8 p. 3. *)
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Variable O : E.
+
+Definition as_plus (A B : E) : E := O +++ ((O --> A) + (O --> B)).
+Definition as_opp (A : E) : E := O +++ (- (O --> A)).
+Definition as_zero : E := O.
+
+Lemma as_plus_comm : forall A B, as_plus A B = as_plus B A.
+Proof. intros; unfold as_plus; rewrite plus_comm; easy. Qed.
+
+Lemma as_plus_assoc :
+  forall A B C, as_plus A (as_plus B C) = as_plus (as_plus A B) C.
+Proof.
+intros; unfold as_plus; apply (vect_l_inj O); rewrite !transl_correct_r.
+rewrite plus_assoc; easy.
+Qed.
+
+Lemma as_plus_zero_r : forall A, as_plus A as_zero = A.
+Proof.
+intros; unfold as_plus, as_zero.
+rewrite vect_zero plus_zero_r transl_correct_l; easy.
+Qed.
+
+Lemma as_plus_opp_r : forall A, as_plus A (as_opp A) = as_zero.
+Proof.
+intros; unfold as_plus, as_opp, as_zero.
+rewrite transl_correct_r plus_opp_r transl_zero; easy.
+Qed.
+
+Definition AffineSpace_AbelianGroup_mixin :=
+  AbelianGroup.Mixin _ _ _ _
+    as_plus_comm as_plus_assoc as_plus_zero_r as_plus_opp_r.
+
+Canonical Structure AffineSpace_AbelianGroup :=
+  AbelianGroup.Pack _ AffineSpace_AbelianGroup_mixin E.
+
+Lemma as_plus_eq : forall A B : E, A + B = O +++ ((O --> A) + (O --> B)).
+Proof. easy. Qed.
+
+Lemma as_opp_eq : forall A : E, - A = O +++ (- (O --> A)).
+Proof. easy. Qed.
+
+Lemma as_zero_eq : (0 : E) = O.
+Proof. easy. Qed.
+
+Lemma as_minus_eq : forall A B : E, A - B = O +++ ((O --> A) - (O --> B)).
+Proof.
+intros; unfold minus; rewrite as_opp_eq as_plus_eq transl_correct_r; easy.
+Qed.
+
+Definition as_scal (x : K) (A : E) : E := O +++ (scal x (O --> A)).
+
+Lemma as_scal_assoc :
+  forall x y A, as_scal x (as_scal y A) = as_scal (x * y) A.
+Proof.
+intros; unfold as_scal; rewrite transl_correct_r scal_assoc; easy.
+Qed.
+
+Lemma as_scal_one_l : forall A, as_scal 1 A = A.
+Proof. intros; unfold as_scal; rewrite scal_one transl_correct_l; easy. Qed.
+
+Lemma as_scal_distr_l :
+  forall x A B, as_scal x (A + B) = as_scal x A + as_scal x B.
+Proof.
+intros; unfold as_scal; rewrite !as_plus_eq.
+rewrite !transl_correct_r scal_distr_l; easy.
+Qed.
+
+Lemma as_scal_distr_r :
+  forall x y A, as_scal (x + y) A = as_scal x A + as_scal y A.
+Proof.
+intros; unfold as_scal; rewrite as_plus_eq.
+rewrite !transl_correct_r scal_distr_r; easy.
+Qed.
+
+Definition AffineSpace_ModuleSpace_mixin :=
+  ModuleSpace.Mixin _ _ _
+    as_scal_assoc as_scal_one_l as_scal_distr_l as_scal_distr_r.
+
+Canonical Structure AffineSpace_ModuleSpace :=
+  ModuleSpace.Pack _ _
+    (ModuleSpace.Class _ _ _ AffineSpace_ModuleSpace_mixin) E.
+
+Lemma as_scal_eq : forall x A, scal x A = O +++ (scal x (O --> A)).
+Proof. easy. Qed.
+
+Lemma vect_l_is_linear_mapping : is_linear_mapping (vect O).
+Proof.
+split.
+intros; rewrite as_plus_eq transl_correct_r; easy.
+intros; rewrite as_scal_eq transl_correct_r; easy.
+Qed.
+
+Lemma transl_l_is_linear_mapping :
+  is_linear_mapping (transl O : V -> AffineSpace_ModuleSpace).
+Proof.
+split.
+intros; rewrite as_plus_eq !transl_correct_r; easy.
+intros; rewrite as_scal_eq transl_correct_r; easy.
+Qed.
+
+End AffineSpace_Is_ModuleSpace.
+*)
+
+
+Section FF_AffineSpace_Def.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Definition vectF {n} O (A : 'E^n) : 'V^n := mapF (vect O) A.
+Definition translF {n} O (u : 'V^n) : 'E^n := mapF (transl O) u.
+
+Definition frameF {n} (A : 'E^n.+1) i0 : 'V^n := vectF (A i0) (skipF A i0).
+Definition inv_frameF {n} O (u : 'V^n) i0 : 'E^n.+1 := translF O (insertF u 0 i0).
+
+(** Correctness lemmas. *)
+
+Lemma vectF_correct : forall {n} O (A : 'E^n) i, vectF O A i = O --> A i.
+Proof. easy. Qed.
+
+Lemma translF_correct : forall {n} O (u : 'V^n) i, translF O u i = O +++ u i.
+Proof. easy. Qed.
+
+Lemma frameF_correct :
+  forall {n} (A : 'E^n.+1) i0 i (H : i <> i0),
+    frameF A i0 (insert_ord H) = A i0 --> A i.
+Proof. intros; unfold frameF; rewrite vectF_correct skipF_correct; easy. Qed.
+
+Lemma inv_frameF_correct :
+  forall {n} (O : E) (u : 'V^n) i0 j,
+    inv_frameF O u i0 (skip_ord i0 j) = O +++ u j.
+Proof.
+intros; unfold inv_frameF; rewrite translF_correct insertF_correct; easy.
+Qed.
+
+End FF_AffineSpace_Def.
+
+
+Section FF_AffineSpace_Facts.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma vectF_change_orig :
+  forall {n} O1 O2 (A : 'E^n), vectF O2 A = constF n (O2 --> O1) + vectF O1 A.
+Proof.
+intros; apply extF; intro; rewrite fct_plus_eq !vectF_correct constF_correct.
+apply eq_sym, vect_chasles.
+Qed.
+
+Lemma vectF_chasles :
+  forall {n} O (A B : 'E^n), vectF O A + (A --> B) = vectF O B.
+Proof.
+intros; apply extF; intro; rewrite fct_plus_eq !vectF_correct fct_vect_eq.
+apply vect_chasles.
+Qed.
+
+(* FIXME: il faut ajouter ici une utilisation de ex_EX ! *)
+Lemma vectF_l_bij_EX :
+  forall {n} O u, { A : 'E^n | unique (fun X => vectF O X = u) A }.
+Proof.
+intros n O u; apply ex_EX; destruct (unique_choice _ (vect_l_bij_ex O))
+    as [A HA]; exists (mapF A u); split.
+apply extF; intro; rewrite vectF_correct mapF_correct; easy.
+move=> B /extF_rev HB; apply extF; intro.
+apply (vect_l_inj O); rewrite mapF_correct -vectF_correct HB; easy.
+Qed.
+
+Lemma vectF_l_bij_ex : forall {n} O u, exists! (A : 'E^n), vectF O A = u.
+Proof.
+intros n O u; destruct (vectF_l_bij_EX O u) as [A HA]; exists A; easy.
+Qed.
+
+Lemma vectF_l_bij : forall {n} (O : E), bijective (@vectF _ _ _ n O).
+Proof. intros; apply bij_ex_uniq_rev, vectF_l_bij_ex. Qed.
+
+Lemma vectF_zero : forall {n} (O : E), vectF O (constF n O) = 0.
+Proof. intros; apply extF; intro; apply vect_zero. Qed.
+
+Lemma vectF_zero_equiv :
+  forall {n} O (A : 'E^n), vectF O A = 0 <-> A = constF n O.
+Proof.
+intros; split.
+move=> /extF_rev H; apply extF; intro; apply vect_zero_equiv; apply H.
+intros; subst; apply vectF_zero.
+Qed.
+
+Lemma vectF_zero_alt : forall {n} (A : 'E^n) i, vectF (A i) A i = 0.
+Proof. intros; apply vect_zero. Qed.
+
+Lemma vectF_l_eq :
+  forall {n} O (A1 A2 : 'E^n), A1 = A2 -> vectF O A1 = vectF O A2.
+Proof. intros; subst; easy. Qed.
+
+Lemma vectF_r_eq :
+  forall {n} (A : 'E^n) O1 O2, O1 = O2 -> vectF O1 A = vectF O2 A.
+Proof. intros; subst; easy. Qed.
+
+Lemma vectF_l_inj :
+  forall {n} O (A1 A2 : 'E^n), vectF O A1 = vectF O A2 -> A1 = A2.
+Proof. move=>>; eapply (bij_inj (vectF_l_bij _)). Qed.
+
+Lemma vectF_r_inj :
+  forall {n} (A : 'E^n.+1) O1 O2, vectF O1 A = vectF O2 A -> O1 = O2.
+Proof.
+move=> n A O1 O2 /extF_rev H; specialize (H ord0).
+apply (vect_r_inj (A ord0)); easy.
+Qed.
+
+Lemma vectF_l_inj_equiv :
+  forall {n} O (A1 A2 : 'E^n), vectF O A1 = vectF O A2 <-> A1 = A2.
+Proof. intros; split. apply vectF_l_inj. apply vectF_l_eq. Qed.
+
+Lemma vectF_r_inj_equiv :
+  forall {n} (A : 'E^n.+1) O1 O2, vectF O1 A = vectF O2 A <-> O1 = O2.
+Proof. intros; split. apply vectF_r_inj. apply vectF_r_eq. Qed.
+
+Lemma vectF_inj_equiv :
+  forall {n} O (A : 'E^n), injective (vectF O A) <-> injective A.
+Proof.
+intros n O A; split; intros HA i1 i2 H; apply HA;
+    [rewrite vectF_correct H | apply (vect_l_inj O)]; easy.
+Qed.
+
+Lemma vectF_constF :
+  forall {n} (O A : E), vectF O (constF n A) = constF n (O --> A).
+Proof.
+intros; apply extF; intro; rewrite vectF_correct !constF_correct; easy.
+Qed.
+
+Lemma vectF_w_zero_struct_r :
+  forall {n} O (A : 'E^n), zero_struct K -> vectF O A = 0.
+Proof.
+move=>> HK; move: (@ms_w_zero_struct_r _ V HK) => HV.
+apply extF; intro; rewrite vectF_correct; apply HV.
+Qed.
+
+Lemma vectF_singleF :
+  forall (O A0 : E), vectF O (singleF A0) = singleF (O --> A0).
+Proof.
+intros; apply extF; intro; rewrite ord1 vectF_correct !singleF_0; easy.
+Qed.
+
+Lemma vectF_coupleF :
+  forall (O A0 A1 : E), vectF O (coupleF A0 A1) = coupleF (O --> A0) (O --> A1).
+Proof.
+intros; apply extF; intros i; destruct (ord2_dec i) as [Hi | Hi]; rewrite Hi.
+rewrite vectF_correct !coupleF_0; easy.
+rewrite vectF_correct !coupleF_1; easy.
+Qed.
+
+Lemma vectF_tripleF :
+  forall (O A0 A1 A2 : E),
+    vectF O (tripleF A0 A1 A2) = tripleF (O --> A0) (O --> A1) (O --> A2).
+Proof.
+intros; apply extF; intros i;
+    destruct (ord3_dec i) as [[Hi | Hi] | Hi]; rewrite Hi vectF_correct.
+rewrite !tripleF_0; easy.
+rewrite !tripleF_1; easy.
+rewrite !tripleF_2; easy.
+Qed.
+
+Lemma vectF_invalF :
+  forall {n1 n2} O (A1 : 'E^n1) (A2 : 'E^n2),
+    invalF (vectF O A1) (vectF O A2) <-> invalF A1 A2.
+Proof.
+intros; split; intros HA i1; destruct (HA i1) as [i2 Hi2]; exists i2.
+apply (vect_l_inj O); easy.
+rewrite !vectF_correct Hi2; easy.
+Qed.
+
+Lemma vectF_insertF :
+  forall {n} O (A : 'E^n) Ai0 i0,
+    vectF O (insertF A Ai0 i0) = insertF (vectF O A) (O --> Ai0) i0.
+Proof.
+intros n O A Ai0 i0; apply extF; intros i;
+    destruct (ord_eq_dec i i0) as [Hi | Hi].
+rewrite vectF_correct !insertF_correct_l; easy.
+rewrite vectF_correct !insertF_correct_r; easy.
+Qed.
+
+Lemma vectF_skipF :
+  forall {n} O (A : 'E^n.+1) i0,
+    vectF O (skipF A i0) = skipF (vectF O A) i0.
+Proof.
+intros n O A i0; apply extF; intros j; destruct (lt_dec j i0) as [Hj | Hj].
+rewrite vectF_correct !skipF_correct_l; easy.
+rewrite vectF_correct !skipF_correct_r; easy.
+Qed.
+
+Lemma vectF_permutF :
+  forall {n} p O (A : 'E^n), vectF O (permutF p A) = permutF p (vectF O A).
+Proof. easy. Qed.
+
+Lemma vectF_revF :
+  forall {n} O (A : 'E^n), vectF O (revF A) = revF (vectF O A).
+Proof. easy. Qed.
+
+Lemma vectF_moveF :
+  forall {n} O (A : 'E^n.+1) i0 i1,
+    vectF O (moveF A i0 i1) = moveF (vectF O A) i0 i1.
+Proof. easy. Qed.
+
+Lemma vectF_transpF :
+  forall {n} O (A : 'E^n) i0 i1,
+    vectF O (transpF A i0 i1) = transpF (vectF O A) i0 i1.
+Proof. easy. Qed.
+
+Lemma vectF_filterPF :
+  forall {n} P O (A : 'E^n), vectF O (filterPF P A) = filterPF P (vectF O A).
+Proof. easy. Qed.
+
+Lemma vectF_concatnF :
+  forall {n} {b : 'nat^n} (O : E) (A : forall i, 'E^(b i)),
+    vectF O (concatnF A) = concatnF (fun i => vectF O (A i)).
+Proof.
+intros; apply extF; intros k; rewrite (splitn_ord k).
+rewrite vectF_correct !concatn_ord_correct; easy.
+Qed.
+
+Lemma translF_vectF :
+  forall {n} (O : E), cancel (@vectF _ _ _ n O) (translF O).
+  (* forall {n} O (A : 'E^n), translF O (vectF O A) = A. *)
+Proof.
+move=>>; apply extF; intro; rewrite translF_correct vectF_correct.
+apply transl_correct_l.
+Qed.
+
+Lemma vectF_translF :
+  forall {n} (O : E), cancel (@translF _ _ _ n O) (vectF O).
+  (* forall {n} (O : E) (u : 'V^n), vectF O (translF O u) = u. *)
+Proof.
+move=>>; apply extF; intro; rewrite vectF_correct translF_correct.
+apply transl_correct_r.
+Qed.
+
+Lemma translF_assoc :
+  forall {n} (O : E) (u v : 'V^n), (translF O u) +++ v = translF O (u + v).
+Proof.
+intros n O u v; rewrite -(translF_vectF O (_ +++ _)).
+rewrite -(vectF_chasles _ (translF O u)) vectF_translF transl_correct_r; easy.
+Qed.
+
+Lemma translF_zero : forall {n} (O : E), translF O 0 = constF n O.
+Proof. intros; apply extF; intro; rewrite translF_correct transl_zero //. Qed.
+
+Lemma translF_zero_equiv :
+  forall {n} (O : E) u, translF O u = constF n O <-> u = 0.
+Proof. intros; rewrite -vectF_zero_equiv vectF_translF; easy. Qed.
+
+Lemma translF_l_eq :
+  forall {n} (O : E) (u1 u2 : 'V^n), u1 = u2 -> translF O u1 = translF O u2.
+Proof. intros; subst; easy. Qed.
+
+Lemma translF_r_eq :
+  forall {n} (u : 'V^n) (O1 O2 : E), O1 = O2 -> translF O1 u = translF O2 u.
+Proof. intros; subst; easy. Qed.
+
+Lemma translF_l_inj :
+  forall {n} (O : E) (u1 u2 : 'V^n), translF O u1 = translF O u2 -> u1 = u2.
+Proof.
+intros n O u1 u2 H; rewrite -(vectF_translF O u1) -(vectF_translF O u2).
+apply vectF_l_eq; easy.
+Qed.
+
+Lemma translF_r_inj :
+  forall {n} (u : 'V^n.+1) (O1 O2 : E), translF O1 u = translF O2 u -> O1 = O2.
+Proof.
+intros n u O1 O2 H; apply constF_inj with n.
+rewrite -(translF_zero O1) -(translF_zero O2) -(plus_opp_r u) -!translF_assoc.
+rewrite H; easy.
+Qed.
+
+Lemma translF_l_inj_equiv :
+  forall {n} (O : E) (u1 u2 : 'V^n), translF O u1 = translF O u2 <-> u1 = u2.
+Proof. intros; split. apply translF_l_inj. apply translF_l_eq. Qed.
+
+Lemma translF_r_inj_equiv :
+  forall {n} (u : 'V^n.+1) (O1 O2 : E),
+    translF O1 u = translF O2 u <-> O1 = O2.
+Proof. intros; split. apply translF_r_inj. apply translF_r_eq. Qed.
+
+Lemma translF_vectF_equiv :
+  forall {n} O (A : 'E^n) u, A = translF O u <-> u = vectF O A.
+Proof. intros; erewrite <- vectF_l_inj_equiv, vectF_translF; easy. Qed.
+
+Lemma translF_l_bij : forall {n} (O : E), bijective (@translF _ _ _ n O).
+Proof.
+intros n O; apply (bij_can_bij (vectF_l_bij O)); apply translF_vectF.
+Qed.
+
+Lemma translF_l_bij_ex : forall {n} O (A : 'E^n), exists! u, translF O u = A.
+Proof. intros; apply bij_ex_uniq, translF_l_bij. Qed.
+
+Lemma translF_inj_equiv :
+  forall {n} (O : E) (u : 'V^n), injective (translF O u) <-> injective u.
+Proof.
+intros n O u; split; intros Hu i1 i2 H; apply Hu;
+    [rewrite translF_correct H | apply (transl_l_inj O)]; easy.
+Qed.
+
+Lemma translF_constF :
+  forall {n} (O : E) u, translF O (constF n u) = constF n (O +++ u).
+Proof.
+intros; apply extF; intro; rewrite translF_correct !constF_correct; easy.
+Qed.
+
+Lemma translF_w_zero_struct_r :
+  forall {n} (O O' : E) (u : 'V^n), zero_struct K -> translF O u = constF n O'.
+Proof.
+intros n O O' u HK; move: (@ms_w_zero_struct_r _ V HK) => HV.
+apply eq_sym, translF_vectF_equiv, extF; intros i.
+rewrite vectF_w_zero_struct_r; easy.
+Qed.
+
+Lemma translF_singleF :
+  forall (O : E) u0, translF O (singleF u0) = singleF (O +++ u0).
+Proof.
+intros; apply extF; intro; rewrite ord1 translF_correct !singleF_0; easy.
+Qed.
+
+Lemma translF_coupleF :
+  forall (O : E) u0 u1, translF O (coupleF u0 u1) = coupleF (O +++ u0) (O +++ u1).
+Proof.
+intros; apply extF; intros i; destruct (ord2_dec i) as [Hi | Hi]; rewrite Hi.
+rewrite translF_correct !coupleF_0; easy.
+rewrite translF_correct !coupleF_1; easy.
+Qed.
+
+Lemma translF_tripleF :
+  forall (O : E) u0 u1 u2,
+    translF O (tripleF u0 u1 u2) = tripleF (O +++ u0) (O +++ u1) (O +++ u2).
+Proof.
+intros; apply extF; intros i;
+    destruct (ord3_dec i) as [[Hi | Hi] | Hi]; rewrite Hi translF_correct.
+rewrite !tripleF_0; easy.
+rewrite !tripleF_1; easy.
+rewrite !tripleF_2; easy.
+Qed.
+
+Lemma translF_invalF :
+  forall {n1 n2} (O : E) (u1 : 'V^n1) (u2 : 'V^n2),
+    invalF (translF O u1) (translF O u2) <-> invalF u1 u2.
+Proof.
+intros; split; intros Hu i1; destruct (Hu i1) as [i2 Hi2]; exists i2.
+apply (transl_l_inj O); easy.
+rewrite !translF_correct Hi2; easy.
+Qed.
+
+Lemma translF_insertF :
+  forall {n} (O : E) (u : 'V^n) u0 i0,
+    translF O (insertF u u0 i0) = insertF (translF O u) (O +++ u0) i0.
+Proof.
+intros n O u u0 i0; apply extF; intros i; destruct (ord_eq_dec i i0) as [Hi | Hi].
+rewrite translF_correct !insertF_correct_l; easy.
+rewrite translF_correct !insertF_correct_r; easy.
+Qed.
+
+Lemma translF_skipF :
+  forall {n} (O : E) (u : 'V^n.+1) i0,
+    translF O (skipF u i0) = skipF (translF O u) i0.
+Proof.
+intros n O A i0; apply extF; intros j; destruct (lt_dec j i0) as [Hj | Hj].
+rewrite translF_correct !skipF_correct_l; easy.
+rewrite translF_correct !skipF_correct_r; easy.
+Qed.
+
+Lemma translF_permutF :
+  forall {n} p (O : E) (u : 'V^n),
+    translF O (permutF p u) = permutF p (translF O u).
+Proof. easy. Qed.
+
+Lemma translF_revF :
+  forall {n} (O : E) (u : 'V^n), translF O (revF u) = revF (translF O u).
+Proof. easy. Qed.
+
+Lemma translF_moveF :
+  forall {n} (O : E) (u : 'V^n.+1) i0 i1,
+    translF O (moveF u i0 i1) = moveF (translF O u) i0 i1.
+Proof. easy. Qed.
+
+Lemma translF_transpF :
+  forall {n} (O : E) (u : 'V^n) i0 i1,
+    translF O (transpF u i0 i1) = transpF (translF O u) i0 i1.
+Proof. easy. Qed.
+
+Lemma translF_filterPF :
+  forall {n} P (O : E) (u : 'V^n),
+    translF O (filterPF P u) = filterPF P (translF O u).
+Proof. easy. Qed.
+
+Lemma translF_concatnF :
+  forall {n} {b : 'nat^n} (O : E) (u : forall i, 'V^(b i)),
+    translF O (concatnF u) = concatnF (fun i => translF O (u i)).
+Proof.
+intros; apply extF; intros k; rewrite (splitn_ord k).
+rewrite translF_correct !concatn_ord_correct; easy.
+Qed.
+
+Lemma frameF_inv_frameF :
+  forall {n} (O : E) (u : 'V^n) i0, frameF (inv_frameF O u i0) i0 = u.
+Proof.
+intros; unfold frameF, inv_frameF; apply eq_sym, translF_vectF_equiv.
+rewrite translF_correct (insertF_correct_l _ _ (erefl _)).
+rewrite -{1}(transl_zero O) translF_insertF skipF_insertF; easy.
+Qed.
+
+Lemma inv_frameF_frameF :
+  forall {n} (A : 'E^n.+1) i0, inv_frameF (A i0) (frameF A i0) i0 = A.
+Proof.
+intros; unfold frameF, inv_frameF; apply eq_sym, translF_vectF_equiv.
+rewrite -(vect_zero (A i0)) -vectF_insertF insertF_skipF; easy.
+Qed.
+
+Lemma frameF_inj_equiv :
+  forall {n} (A : 'E^n.+1) i0,
+    injective (frameF A i0) <-> injective (skipF A i0).
+Proof. move=>>; apply vectF_inj_equiv. Qed.
+
+Lemma inv_frameF_inj_equiv :
+  forall {n} (O : E) (u : 'V^n) i0,
+    injective (skipF (inv_frameF O u i0) i0) <-> injective u.
+Proof. move=>>; rewrite -frameF_inj_equiv frameF_inv_frameF; easy. Qed.
+
+(* (preimage (transl (A i0)) PE) will be (vectP PE (A i0)) in Sub_struct. *)
+Lemma frameF_inclF :
+  forall {PE : E -> Prop} {n} {A : 'E^n.+1} i0,
+    inclF A PE -> inclF (frameF A i0) (preimage (transl (A i0)) PE).
+Proof.
+move=>> HA i; unfold preimage, frameF.
+rewrite vectF_correct transl_correct_l.
+apply: (inclF_monot_l _ _ _ _ HA); try easy.
+apply skipF_monot_l, invalF_refl.
+Qed.
+
+Lemma frameF_invalF :
+  forall {n1 n2} (A1 : 'E^n1.+1) (A2 : 'E^n2.+1) i1 i2,
+    A1 i1 = A2 i2 ->
+    invalF (skipF A1 i1) (skipF A2 i2) -> invalF (frameF A1 i1) (frameF A2 i2).
+Proof. move=>> H; unfold frameF; rewrite H; apply vectF_invalF. Qed.
+
+Lemma frameF_invalF_rev :
+  forall {n1 n2} (A1 : 'E^n1.+1) (A2 : 'E^n2.+1) i1 i2,
+    A1 i1 = A2 i2 ->
+    invalF (frameF A1 i1) (frameF A2 i2) -> invalF (skipF A1 i1) (skipF A2 i2).
+Proof. move=>> H; unfold frameF; rewrite H; apply vectF_invalF. Qed.
+
+Lemma frameF_invalF_equiv :
+  forall {n1 n2} (A1 : 'E^n1.+1) (A2 : 'E^n2.+1) i1 i2,
+    A1 i1 = A2 i2 ->
+    invalF (frameF A1 i1) (frameF A2 i2) <->
+    invalF (skipF A1 i1) (skipF A2 i2).
+Proof. move=>> H; unfold frameF; rewrite H; apply vectF_invalF. Qed.
+
+Lemma frameF_skipF :
+  forall {n} (A : 'E^n.+2) {i0 i1} (H10 : i1 <> i0) (H01 : i0 <> i1),
+    frameF (skipF A i0) (insert_ord H10) =
+      skipF (frameF A i1) (insert_ord H01).
+Proof.
+intros n A i0 i1 H10 H01; unfold frameF.
+rewrite -vectF_skipF (skipF_correct_alt H10); [| apply skip_insert_ord].
+rewrite -skip2F_correct skip2F_equiv_def; easy.
+Qed.
+
+Lemma frameF_skipF_alt :
+  forall {n} (A : 'E^n.+2) {i0 i1} {H10 : i1 <> i0},
+    frameF (skipF A i0) (insert_ord H10) =
+      skipF (frameF A i1) (insert_ord (not_eq_sym H10)).
+Proof. intros; apply frameF_skipF. Qed.
+
+Lemma frameF_permutF :
+  forall {n} {A : 'E^n.+1} {i0} {p} (Hp : injective p),
+    frameF (permutF p A) i0 = permutF (skip_f_ord Hp i0) (frameF A (p i0)).
+Proof. intros; unfold frameF; rewrite skipF_permutF; easy. Qed.
+
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Lemma vectF_mapF :
+  forall {n} O1 (A1 : 'E1^n) (f : E1 -> E2),
+    vectF O1 (mapF f A1) = mapF (fun A1i => O1 --> f A1i) A1.
+Proof. easy. Qed.
+
+Lemma translF_mapF :
+  forall {n} (O2 : E2) (u1 : 'V1^n) (f : V1 -> V2),
+    translF O2 (mapF f u1) = mapF (fun u1i => O2 +++ f u1i) u1.
+Proof. easy. Qed.
+
+Lemma fct_vectF_eq :
+  forall {T : Type} {n} (fO : T -> E) (fA : '(T -> E)^n) t,
+    (vectF fO fA)^~ t = vectF (fO t) (fA^~ t).
+Proof. easy. Qed.
+
+Lemma fct_translF_eq :
+  forall {T : Type} {n} (fO : T -> E) (fu : '(T -> V)^n) t,
+    (translF fO fu)^~ t = translF (fO t) (fu^~ t).
+Proof.
+intros; apply extF; intro; rewrite !translF_correct; apply fct_transl_eq.
+Qed.
+
+End FF_AffineSpace_Facts.
+
+
+Section Barycenter_Def.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+(* Gostiaux T4, p. 4. *)
+Lemma comb_lin_vectF_uniq :
+  forall {n} L (A : 'E^n) O1 O2,
+    sum L = 0 -> comb_lin L (vectF O1 A) = comb_lin L (vectF O2 A).
+Proof.
+intros n L A O1 O2 HL; rewrite (vectF_change_orig O1 O2).
+rewrite comb_lin_plus_r -scal_sum_l HL scal_zero_l plus_zero_l; easy.
+Qed.
+
+Definition is_comb_aff {n} L (A : 'E^n) G : Prop := comb_lin L (vectF G A) = 0.
+
+Lemma is_comb_aff_w_zero_struct_r :
+  forall {n} L (A : 'E^n) G, zero_struct K -> is_comb_aff L A G.
+Proof.
+move=>> HK; move: (@ms_w_zero_struct_r _ V HK) => HV.
+unfold is_comb_aff; rewrite vectF_w_zero_struct_r; easy.
+Qed.
+
+Lemma is_comb_aff_compat_l :
+  forall {n} L1 L2 (A : 'E^n) G,
+    L1 = L2 -> is_comb_aff L1 A G <-> is_comb_aff L2 A G.
+Proof. intros; subst; easy. Qed.
+
+Lemma is_comb_aff_compat_r :
+  forall {n} L (A1 A2 : 'E^n) G,
+    A1 = A2 -> is_comb_aff L A1 G <-> is_comb_aff L A2 G.
+Proof. intros; subst; easy. Qed.
+
+Lemma is_comb_aff_uniq :
+  forall {n} L (A : 'E^n) G1 G2,
+    invertible (sum L) -> is_comb_aff L A G1 -> is_comb_aff L A G2 -> G1 = G2.
+Proof.
+unfold is_comb_aff; intros n L A G1 G2 HL HG1 HG2.
+apply eq_sym, vect_zero_equiv, (scal_zero_reg_r (sum L)); try easy.
+rewrite scal_sum_l -(minus_eq_zero 0) -{1}HG1 -HG2 -comb_lin_minus_r; f_equal.
+rewrite -plus_minus_r_equiv plus_comm; apply vectF_change_orig.
+Qed.
+
+Lemma is_comb_aff_homogeneous :
+  forall {n} L a (A : 'E^n) G,
+    invertible a -> is_comb_aff L A G <-> is_comb_aff (scal a L) A G.
+Proof.
+intros; unfold is_comb_aff; rewrite comb_lin_scal_l; split; intros HG.
+rewrite HG scal_zero_r; easy.
+destruct (scal_zero_rev _ _ HG); easy.
+Qed.
+
+Lemma is_comb_aff_orig_compat :
+  forall {n} L (A : 'E^n) G O,
+    scal (sum L) (O --> G) = comb_lin L (vectF O A) -> is_comb_aff L A G.
+Proof.
+unfold is_comb_aff; intros n L A G O HG.
+rewrite (vectF_change_orig O) comb_lin_plus_r -scal_sum_l.
+rewrite -HG vect_sym scal_opp_r plus_opp_l; easy.
+Qed.
+
+Lemma is_comb_aff_orig_reg :
+  forall {n} L (A : 'E^n) G,
+    is_comb_aff L A G ->
+    forall O, scal (sum L) (O --> G) = comb_lin L (vectF O A).
+Proof.
+unfold is_comb_aff; move=>> HG O; generalize HG; clear HG.
+rewrite (vectF_change_orig O) comb_lin_plus_r scal_sum_l.
+rewrite vect_sym constF_opp comb_lin_opp_r plus_is_zero_r_equiv opp_opp; easy.
+Qed.
+
+(* FIXME: pfff, il faut aussi un existentiel fort dans invertible... *)
+Lemma is_comb_aff_EX :
+  forall {n L} (A : 'E^n), invertible (sum L) -> { G | is_comb_aff L A G }.
+Proof.
+intros n L A [sLm1 [_ HL]]; destruct (inhabited_as E) as [O].
+pose (u := scal sLm1 (comb_lin L (vectF O A))).
+destruct (vect_l_bij_ex O u) as [G [HG _]]; exists G.
+apply (is_comb_aff_orig_compat _ _ _ O); rewrite HG; unfold u.
+rewrite scal_assoc HL scal_one; easy.
+Qed.
+
+Lemma is_comb_aff_permutF :
+  forall {n} p {L} (A : 'E^n) G,
+    injective p -> invertible (sum L) ->
+    is_comb_aff L A G -> is_comb_aff (permutF p L) (permutF p A) G.
+Proof.
+intros; unfold is_comb_aff; rewrite vectF_permutF comb_lin_permutF//.
+Qed.
+
+Lemma is_comb_aff_revF :
+  forall {n} {L} (A : 'E^n) G,
+    invertible (sum L) -> is_comb_aff L A G -> is_comb_aff (revF L) (revF A) G.
+Proof. move=>>; apply is_comb_aff_permutF, rev_ord_inj. Qed.
+
+Lemma is_comb_aff_moveF :
+  forall {n} {L} (A : 'E^n.+1) i0 i1 G,
+    invertible (sum L) ->
+    is_comb_aff L A G -> is_comb_aff (moveF L i0 i1) (moveF A i0 i1) G.
+Proof. move=>>; apply is_comb_aff_permutF, move_ord_inj. Qed.
+
+Lemma is_comb_aff_transpF :
+  forall {n} {L} (A : 'E^n) i0 i1 G,
+    invertible (sum L) ->
+    is_comb_aff L A G -> is_comb_aff (transpF L i0 i1) (transpF A i0 i1) G.
+Proof. move=>>; apply is_comb_aff_permutF, transp_ord_inj. Qed.
+
+Lemma is_comb_aff_squeezeF :
+  forall {n} {L} {A : 'E^n.+1} {i0 i1} G,
+    i1 <> i0 -> A i1 = A i0 ->
+    is_comb_aff L A G -> is_comb_aff (squeezeF L i0 i1) (skipF A i1) G.
+Proof.
+move=>> Hi HA HG; unfold is_comb_aff.
+rewrite vectF_skipF comb_lin_squeezeF// !vectF_correct HA; easy.
+Qed.
+
+Lemma barycenter_EX :
+  forall {n L} (A : 'E^n), invertible (sum L) -> { G | is_comb_aff L A G }.
+Proof. intros; apply ex_EX, is_comb_aff_ex; easy. Qed.
+
+Definition barycenter {n} L (A : 'E^n) : E :=
+  match invertible_dec (sum L) with
+  | left HL => proj1_sig (barycenter_EX A HL)
+  | right _ => point_of_as E
+  end.
+
+Definition isobarycenter {n} (A : 'E^n) : E := barycenter ones A.
+
+Definition middle (A B : E) : E := isobarycenter (coupleF A B).
+
+Lemma barycenter_correct :
+  forall {n L} (A : 'E^n),
+    invertible (sum L) -> is_comb_aff L A (barycenter L A).
+Proof.
+intros n L A HL; unfold barycenter;
+    destruct (invertible_dec (sum L)) as [HL' | ]; try easy.
+apply (proj2_sig (barycenter_EX A HL')).
+Qed.
+
+Lemma barycenter_correct_equiv :
+  forall {n} {L} (A : 'E^n) {G},
+    invertible (sum L) -> G = barycenter L A <-> is_comb_aff L A G.
+Proof.
+intros n L A G HL; split.
+intros; subst; apply barycenter_correct; easy.
+intros HG; apply (is_comb_aff_uniq L A); try apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_correct_orig :
+  forall O {n} {L} (A : 'E^n),
+    invertible (sum L) ->
+    scal (sum L) (O --> barycenter L A) = comb_lin L (vectF O A).
+Proof. intros; apply is_comb_aff_orig_reg, barycenter_correct; easy. Qed.
+
+Lemma barycenter_correct_orig_equiv :
+  forall O {n} {L} (A : 'E^n) {G},
+    invertible (sum L) ->
+    G = barycenter L A <-> scal (sum L) (O --> G) = comb_lin L (vectF O A).
+Proof.
+intros O n L A G HL; split.
+intros; subst; apply barycenter_correct_orig; easy.
+intros HG; apply (is_comb_aff_uniq L A); try easy.
+apply (is_comb_aff_orig_compat _ _ _ _ HG).
+apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_eq_l :
+  forall {n} {L} M {A : 'E^n}, L = M -> barycenter L A = barycenter M A.
+Proof. intros; f_equal; easy. Qed.
+
+Lemma barycenter_eq_r :
+  forall {n} {L A} (B : 'E^n), A = B -> barycenter L A = barycenter L B.
+Proof. intros; f_equal; easy. Qed.
+
+Lemma barycenter_eq :
+  forall {n} {L} M {A} (B : 'E^n),
+    L = M -> A = B -> barycenter L A = barycenter M B.
+Proof. intros; f_equal; easy. Qed.
+
+End Barycenter_Def.
+
+
+Section Barycenter_Facts.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma barycenter_w_zero_struct_r :
+  forall {n} L (A : 'E^n) G, zero_struct K -> G = barycenter L A.
+Proof.
+intros; apply barycenter_correct_equiv.
+apply invertible_zero_struct; easy.
+apply is_comb_aff_w_zero_struct_r; easy.
+Qed.
+
+Lemma transl_comb_lin :
+  forall {n} i0 (A : E) L (u : 'V^n),
+    A +++ comb_lin L u =
+      barycenter (insertF L (1 - sum L) i0) (insertF (translF A u) A i0).
+Proof.
+intros n i0 A L u; apply (barycenter_correct_orig_equiv A);
+  rewrite sum_insertF_baryc; try apply invertible_one.
+rewrite scal_one transl_correct_r vectF_insertF comb_lin_insertF.
+rewrite vect_zero scal_zero_r plus_zero_l vectF_translF; easy.
+Qed.
+
+Lemma barycenter_insertF :
+  forall {n} i0 O (A : 'E^n) L,
+    barycenter (insertF L (1 - sum L) i0) (insertF A O i0) =
+      O +++ comb_lin L (vectF O A).
+Proof. intros; rewrite (transl_comb_lin i0) translF_vectF; easy. Qed.
+
+Lemma comb_lin_vectF :
+  forall {n} i0 O (A : 'E^n) L,
+    comb_lin L (vectF O A) =
+      O --> barycenter (insertF L (1 - sum L) i0) (insertF A O i0).
+Proof.
+intros n i0 O A L; apply (transl_l_inj O).
+rewrite -(barycenter_insertF i0) transl_correct_l; easy.
+Qed.
+
+Lemma comb_lin_vectF_skipF :
+  forall {n} i0 (A : 'E^n.+1) L,
+    comb_lin L (vectF (A i0) A) =
+      comb_lin (skipF L i0) (vectF (A i0) (skipF A i0)).
+Proof.
+intros n i0 A L; rewrite (comb_lin_skipF i0) vectF_correct vect_zero
+    scal_zero_r plus_zero_l; easy.
+Qed.
+
+Lemma comb_lin_vectF_0 :
+  forall {n} (A : 'E^n.+1) L,
+    comb_lin L (vectF (A ord0) A) =
+      comb_lin (liftF_S L) (vectF (A ord0) (liftF_S A)).
+Proof. intros; rewrite -!skipF_first; apply comb_lin_vectF_skipF. Qed.
+
+Lemma comb_lin_vectF_max :
+  forall {n} (A : 'E^n.+1) L,
+    comb_lin L (vectF (A ord_max) A) =
+      comb_lin (widenF_S L) (vectF (A ord_max) (widenF_S A)).
+Proof. intros; rewrite -!skipF_last; apply comb_lin_vectF_skipF. Qed.
+
+Lemma barycenter2_r :
+  forall {n1 n2} L1 (A1 : 'E^n1) M12 (A2 : 'E^n2),
+    invertible (sum L1) -> (forall i1, sum (M12 i1) = 1) ->
+    (forall i1, A1 i1 = barycenter (M12 i1) A2) ->
+    barycenter L1 A1 = barycenter (fun i2 => comb_lin L1 (M12^~ i2)) A2.
+Proof.
+intros n1 n2 L1 A1 M12 A2 HL1 HM12 HA1; pose (O := point_of_as E).
+assert (HLM : sum (fun i2 => comb_lin L1 (M12^~ i2)) = sum L1).
+  rewrite -comb_lin_sum_r -(comb_lin_ext_r _ ones).
+  apply comb_lin_ones_r.
+  intros i1; rewrite (HM12 i1); easy.
+apply (barycenter_correct_orig_equiv O); rewrite HLM; try easy.
+assert (HA1' : forall i1, O --> A1 i1 = comb_lin (M12 i1) (vectF O A2)).
+  intros i1; rewrite -(scal_one (O --> _)) -(HM12 i1).
+  apply barycenter_correct_orig_equiv; try easy.
+  rewrite HM12; apply invertible_one.
+apply (comb_lin2_alt L1) in HA1'; rewrite -HA1'.
+apply barycenter_correct_orig; easy.
+Qed.
+
+Lemma barycenter_kron_l :
+  forall {n} (A : 'E^n) (j : 'I_n), barycenter ((kron K)^~ j) A = A j.
+Proof.
+intros n A j; destruct n as [|n]; try now destruct j.
+apply eq_sym, (barycenter_correct_orig_equiv (A ord0)); rewrite sum_kron_l.
+apply invertible_one.
+rewrite scal_one comb_lin_kron_in_l; easy.
+Qed.
+
+Lemma barycenter_kron_r :
+  forall {n} (A : 'E^n) (i : 'I_n), barycenter (kron K i) A = A i.
+Proof.
+intros n A i; destruct n as [|n]; try now destruct i.
+apply eq_sym, (barycenter_correct_orig_equiv (A ord0)); rewrite sum_kron_r.
+apply invertible_one.
+rewrite scal_one comb_lin_kron_in_r; easy.
+Qed.
+
+Lemma barycenter_singleF_equiv :
+  forall L0 (A0 : E) G,
+    invertible (sum (singleF L0)) ->
+    G = barycenter (singleF L0) (singleF A0) <-> scal L0 (G --> A0) = 0.
+Proof.
+intros; rewrite barycenter_correct_equiv; unfold is_comb_aff; try easy.
+rewrite comb_lin_1; easy.
+Qed.
+
+Lemma barycenter_coupleF_equiv :
+  forall L0 L1 (A0 A1 : E) G,
+    invertible (sum (coupleF L0 L1)) ->
+    G = barycenter (coupleF L0 L1) (coupleF A0 A1) <->
+    scal L0 (G --> A0) + scal L1 (G --> A1) = 0.
+Proof.
+intros; rewrite barycenter_correct_equiv; unfold is_comb_aff; try easy.
+rewrite comb_lin_2 2!vectF_correct 2!coupleF_0 2!coupleF_1; easy.
+Qed.
+
+Lemma barycenter_tripleF_equiv :
+  forall L0 L1 L2 (A0 A1 A2 : E) G,
+    invertible (sum (tripleF L0 L1 L2)) ->
+    G = barycenter (tripleF L0 L1 L2) (tripleF A0 A1 A2) <->
+    scal L0 (G --> A0) + scal L1 (G --> A1) + scal L2 (G --> A2) = 0.
+Proof.
+intros; rewrite barycenter_correct_equiv; unfold is_comb_aff; try easy.
+rewrite comb_lin_3 3!vectF_correct 2!tripleF_0 2!tripleF_1 2!tripleF_2; easy.
+Qed.
+
+Lemma isobarycenter_correct :
+  forall {n} (A : 'E^n),
+    invertible (plusn1 K n) -> is_comb_aff ones A (isobarycenter A).
+Proof. intros; apply barycenter_correct; easy. Qed.
+
+Lemma isobarycenter_correct_orig :
+  forall {n} (A : 'E^n) O,
+    invertible (plusn1 K n) ->
+    scal (plusn1 K n) (O --> isobarycenter A) = sum (vectF O A).
+Proof.
+intros; rewrite -comb_lin_ones_l; apply barycenter_correct_orig; easy.
+Qed.
+
+Lemma isobarycenter_correct_equiv :
+  forall {n} (A : 'E^n) G,
+    invertible (plusn1 K n) ->
+    G = isobarycenter A <-> is_comb_aff ones A G.
+Proof. intros; apply barycenter_correct_equiv; easy. Qed.
+
+Lemma middle_correct :
+  forall (A B : E), let M := middle A B in
+    invertible (2 : K) -> (M --> A) + (M --> B) = 0.
+Proof.
+rewrite -plusn1_2; intros A B M H2;
+    generalize (isobarycenter_correct (coupleF A B) H2).
+unfold is_comb_aff; rewrite comb_lin_2 !scal_one
+    !vectF_correct coupleF_0 coupleF_1; easy.
+Qed.
+
+Lemma middle_correct_orig :
+  forall (A B O : E), let M := middle A B in
+    invertible (2 : K) -> scal 2 (O --> M) = (O --> A) + (O --> B).
+Proof.
+rewrite -plusn1_2; intros A B O M H2;
+    generalize (isobarycenter_correct_orig (coupleF A B) O H2).
+rewrite sum_2 !vectF_correct coupleF_0 coupleF_1; easy.
+Qed.
+
+Lemma middle_correct_equiv :
+  forall (A B M : E), invertible (2 : K) ->
+    M = middle A B <-> (M --> A) + (M --> B) = 0.
+Proof.
+intros A B M H2; rewrite -plusn1_2 in H2;
+    generalize (isobarycenter_correct_equiv (coupleF A B) M H2).
+unfold is_comb_aff; rewrite comb_lin_2 !scal_one
+    !vectF_correct coupleF_0 coupleF_1; easy.
+Qed.
+
+Lemma barycenter_homogeneous :
+  forall {n a L} (A : 'E^n),
+    invertible a -> invertible (sum L) ->
+    barycenter L A = barycenter (scal a L) A.
+Proof.
+intros n a L A Ha HL; apply (is_comb_aff_uniq L A); try easy.
+apply barycenter_correct; easy.
+apply (is_comb_aff_homogeneous _ a); try easy.
+apply barycenter_correct, sum_scal_invertible_compat; easy.
+Qed.
+
+Lemma barycenter_normalized :
+  forall {n} L (A : 'E^n) G, let L1 := scal (inv (sum L)) L in
+    invertible (sum L) -> G = barycenter L A ->
+    sum L1 = 1 /\ G = barycenter L1 A.
+Proof.
+intros n L A G L1 HL; generalize HL; intros [a Ha];
+    generalize Ha; intros [Ha1 Ha2] HG; rewrite HG.
+unfold L1; rewrite (inv_correct Ha); split.
+rewrite -scal_sum_distr_l -Ha1; easy.
+apply barycenter_homogeneous; try easy.
+apply (is_inverse_invertible_r (sum L)); easy.
+Qed.
+
+Lemma barycenter_normal :
+  forall {n} L1 (A : 'E^n) G,
+    sum L1 = 1 -> is_comb_aff L1 A G -> G = barycenter L1 A.
+Proof.
+intros n L1 A G HL1 HG; apply (is_comb_aff_uniq L1 A); try easy.
+2: apply barycenter_correct.
+1,2: apply invertible_eq_one; easy.
+Qed.
+
+Lemma barycenter_skip_zero :
+  forall {n} L (A : 'E^n.+1) i0,
+    invertible (sum L) -> L i0 = 0 ->
+    barycenter L A = barycenter (skipF L i0) (skipF A i0).
+Proof.
+intros n L A i0 HL Hi0; apply barycenter_correct_equiv; unfold is_comb_aff.
+rewrite -(plus_zero_l (sum _)) -Hi0 -sum_skipF; easy.
+rewrite -(plus_zero_l (comb_lin _ _))
+    -{1}(scal_zero_l (barycenter L A --> A i0))
+    -{1}Hi0 vectF_skipF -comb_lin_skipF.
+fold (is_comb_aff L A (barycenter L A)).
+apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_squeezeF :
+  forall {n} {L} {A : 'E^n.+1} {i0 i1},
+    invertible (sum L) -> i1 <> i0 -> A i1 = A i0 ->
+    barycenter L A = barycenter (squeezeF L i0 i1) (skipF A i1).
+Proof.
+intros n L A i0 i1 HL Hia Hib; apply barycenter_correct_equiv.
+apply invertible_sum_squeezeF; easy.
+apply is_comb_aff_squeezeF; try easy.
+apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_injF_ex :
+  forall {n} L (A : 'E^n), invertible (sum L) ->
+    exists m M (B : 'E^m), invertible (sum M) /\
+      injective B /\ invalF B A /\ barycenter L A = barycenter M B.
+Proof.
+intros n L A HL; induction n as [|n Hn].
+(* *)
+rewrite sum_nil in HL; move: (invertible_zero HL) => HK.
+move: (@as_w_zero_struct_r _ _ E HK) => [O HE].
+exists 0, 0, (constF 0 O); repeat split.
+apply invertible_zero_struct; easy.
+1,2: intros [i Hi]; easy.
+apply barycenter_w_zero_struct_r; easy.
+(* *)
+destruct (classic (injective A)) as [HA | HA].
+exists n.+1, L, A; repeat split; try apply invalF_refl; easy.
+move: HA => /not_all_ex_not [i1 /not_all_ex_not [i0 Hi]].
+move: Hi => /not_imp_and_equiv [Hia Hib].
+destruct (Hn (squeezeF L i0 i1) (skipF A i1)) as [m [M [B [H1 [H2 [H3 H4]]]]]];
+    try now apply invertible_sum_squeezeF.
+exists m, M, B; repeat split; try easy.
+apply (skipF_monot_r _ _ i1); easy.
+rewrite (barycenter_squeezeF _ Hib); easy.
+Qed.
+
+Lemma barycenter_filter_n0F :
+  forall {n} L (A : 'E^n),
+    invertible (sum L) ->
+    barycenter L A = barycenter (filter_n0F L) (filter_n0F_gen L A).
+Proof.
+intros n L A HL. pose (G := barycenter L A); fold G.
+apply barycenter_correct_equiv; try rewrite -invertible_sum_equiv//.
+unfold is_comb_aff; rewrite vectF_filterPF comb_lin_filter_n0F_l.
+apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_normalized_n0_ex :
+  forall {n} L (A : 'E^n), invertible (sum L) ->
+    exists n1 L1 (A1 : 'E^n1), sum L1 = 1 /\ (forall i, L1 i <> 0) /\
+      invalF A1 A /\ barycenter L A = barycenter L1 A1.
+Proof.
+intros n L A HL.
+destruct (barycenter_normalized
+    (filter_n0F L) (filter_n0F_gen L A) (barycenter L A)) as [HL1 HA1].
+rewrite -invertible_sum_equiv//.
+apply barycenter_filter_n0F; easy.
+eexists; exists (scal (/ sum (filter_n0F L)) (filter_n0F L)),
+    (filter_n0F_gen L A); repeat split; try easy.
+2: apply filterPF_invalF.
+intros; rewrite fct_scal_eq; unfold scal; simpl; apply mult_not_zero_l.
+apply inv_invertible; rewrite -invertible_sum_equiv//.
+apply filter_n0F_correct.
+Qed.
+
+Lemma barycenter_constF :
+  forall {n} A L (B : 'E^n),
+    invertible (sum L) -> B = constF n A -> barycenter L B = A.
+Proof.
+intros n A L B HL HB; subst; symmetry.
+apply barycenter_correct_equiv, (is_comb_aff_orig_compat _ _ _ A); try easy.
+rewrite vectF_constF scal_sum_l; easy.
+Qed.
+
+Lemma barycenter_singleF :
+  forall a A L (B : 'E^1),
+    invertible a -> L = singleF a -> B = singleF A -> barycenter L B = A.
+Proof.
+intros; apply barycenter_constF; subst; try now apply sum_singleF_invertible.
+apply eq_sym, constF_1.
+Qed.
+
+Lemma barycenter_permutF :
+  forall {n} p L (A : 'E^n),
+    injective p -> invertible (sum L) ->
+    barycenter (permutF p L) (permutF p A) = barycenter L A.
+Proof.
+intros n p L A Hp HL; apply eq_sym, barycenter_correct_equiv.
+rewrite sum_permutF; easy.
+apply is_comb_aff_permutF; try easy.
+apply barycenter_correct; easy.
+Qed.
+
+Lemma barycenter_revF :
+  forall {n} L (A : 'E^n),
+    invertible (sum L) -> barycenter (revF L) (revF A) = barycenter L A.
+Proof. intros; apply barycenter_permutF; [apply rev_ord_inj | easy]. Qed.
+
+Lemma barycenter_moveF :
+  forall {n} L (A : 'E^n.+1) i0 i1,
+    invertible (sum L) ->
+    barycenter (moveF L i0 i1) (moveF A i0 i1) = barycenter L A.
+Proof. intros; apply barycenter_permutF; [apply move_ord_inj | easy]. Qed.
+
+Lemma barycenter_transpF :
+  forall {n} L (A : 'E^n) i0 i1,
+    invertible (sum L) ->
+    barycenter (transpF L i0 i1) (transpF A i0 i1) = barycenter L A.
+Proof. intros; apply barycenter_permutF; [apply transp_ord_inj | easy]. Qed.
+
+Lemma barycenter_unfun0F :
+  forall {n1 n2} L1 (A1 : 'E^n1) (A2 : 'E^n2),
+    invertible (sum L1) -> injective A1 -> invalF A1 A2 ->
+    exists L2, invalF L1 L2 /\
+      invertible (sum L2) /\ barycenter L1 A1 = barycenter L2 A2.
+Proof.
+move=> n1 n2 L1 A1 A2 HL1 HA1 HA.
+move: (barycenter_correct A1 HL1) => HG1; unfold is_comb_aff in HG1.
+pose (G1 := barycenter L1 A1); fold G1 in HG1; fold G1.
+destruct (invalF_fun HA) as [f Hf]; move: (invalF_fun_inj _ HA1 HA Hf) => Hf'.
+assert (HL2 : invertible (sum (unfun0F f L1))) by now rewrite sum_unfun0F.
+exists (unfun0F f L1); repeat split; try easy.
+apply unfun0F_ub; easy.
+apply barycenter_correct_equiv; try easy; unfold is_comb_aff.
+rewrite -comb_lin_filter_n0F_l filter_n0F_unfun0F.
+pose (q1 := proj1_sig (injF_restr_bij_EX Hf')); fold q1.
+pose (Hq1a := proj1 (proj2_sig (injF_restr_bij_EX Hf'))); fold q1 in Hq1a.
+pose (Hq1b := bij_inj Hq1a).
+rewrite 2!comb_lin_castF_l.
+rewrite -(comb_lin_permutF Hq1b) -comb_lin_filter_n0F_l in HG1.
+rewrite -HG1; f_equal; rewrite -castF_sym_equiv.
+apply eq_sym; rewrite -castF_sym_equiv.
+rewrite filter_n0F_gen_unfun0F_l; repeat f_equal.
+apply extF; intros i1; rewrite !vectF_correct Hf; easy.
+Qed.
+
+(* Gostiaux T4, Th 1.18 pp. 6-7. *)
+Lemma barycenter_assoc :
+  forall {n} (b : 'nat^n) L (A : forall i, 'E^(b i)),
+    (forall i, invertible (sum (L i))) ->
+    invertible (sum (concatnF L)) ->
+    barycenter (concatnF L) (concatnF A) =
+      barycenter (fun i => sum (L i)) (fun i => barycenter (L i) (A i)).
+Proof.
+intros n b L A HLi HL; destruct (inhabited_as E) as [B].
+apply barycenter_correct_equiv; try now rewrite -sum_assoc.
+apply (is_comb_aff_orig_compat _ _ _ B).
+rewrite -sum_assoc barycenter_correct_orig; try easy.
+rewrite vectF_concatnF comb_lin_concatnF.
+unfold comb_lin; f_equal; apply extF; intros i.
+rewrite scalF_correct barycenter_correct_orig; easy.
+Qed.
+
+Lemma fct_barycenter_eq :
+  forall {T : Type} {n} L (fA : '(T -> E)^n) t,
+    invertible (sum L) -> barycenter L fA t = barycenter L (fA^~ t).
+Proof.
+intros T n L fA t HL; pose (G := barycenter L fA); fold G.
+assert (HG : G = barycenter L fA) by easy.
+apply barycenter_correct_equiv in HG; try easy.
+apply barycenter_correct_equiv; try easy; unfold is_comb_aff in *.
+rewrite -fct_vectF_eq -fct_comb_lin_r_eq HG; easy.
+Qed.
+
+End Barycenter_Facts.
+
+
+Section Barycenter_R_Facts.
+
+Context {V : ModuleSpace R_Ring}.
+Context {E : AffineSpace V}.
+
+Lemma barycenter_comm_R :
+  forall {n1 n2} L1 L2 (A : 'E^{n1,n2}),
+    sum L1 <> 0 -> sum L2 <> 0 ->
+    barycenter L2 (barycenter L1 A) = barycenter L1 (mapF (barycenter L2) A).
+Proof.
+intros n1 n2 L1 L2 A HL1 HL2; generalize HL1 HL2.
+move=> /invertible_equiv_R HL1' /invertible_equiv_R HL2'.
+pose (G1 := mapF (barycenter L2) A); fold G1.
+assert (HG1' : forall i1, is_comb_aff L2 (A i1) (G1 i1))
+    by now intros; apply barycenter_correct_equiv.
+unfold is_comb_aff in HG1'; pose (G1A := fun i1 => vectF (G1 i1) (A i1)).
+assert (HG1 : mapF (comb_lin L2) G1A = 0)
+    by now apply extF; intro; rewrite mapF_correct; unfold G1A. clear HG1'.
+pose (G2 := barycenter L1 A).
+generalize (barycenter_correct A HL1'); fold G2; intros HG2.
+pose (G := barycenter L2 G2).
+generalize (barycenter_correct G2 HL2'); fold G; intros HG.
+apply barycenter_correct_equiv; try easy; unfold is_comb_aff in *.
+pose (G1s := fun i1 => constF n2 (G1 i1)).
+assert (HGa :
+      (fun i1 => comb_lin L2 (vectF G (G1s i1))) +
+      (fun i1 => comb_lin L2 ((G1s i1) --> G2)) = 0)
+    by now apply extF; intro; rewrite fct_plus_eq -comb_lin_plus_r vectF_chasles.
+apply (comb_lin_eq_r L1) in HGa.
+rewrite comb_lin_plus_r comb_lin_zero_r in HGa.
+apply (scal_zero_reg_r_R (sum L2)); try easy.
+rewrite -(plus_zero_r (scal _ _)) -{2}HGa; f_equal; clear HGa HG.
+rewrite -comb_lin_scal_r; f_equal; apply extF; intro.
+rewrite scal_sum_l fct_comb_lin_r_eq; easy.
+symmetry; rewrite (comb_lin_eq_r _ _
+    (fun i1 => comb_lin L2 (vectF (G1 i1) (A i1) +
+                            (fun i2 => (A i1 i2) --> (G2 i2))))).
+2: apply extF; intros i1; f_equal; unfold G1s; rewrite vectF_chasles; easy.
+rewrite (comb_lin_eq_r _ _
+    ((fun i1 => comb_lin L2 (vectF (G1 i1) (A i1))) +
+     (fun i1 => comb_lin L2 (fun i2 => A i1 i2 --> G2 i2)))).
+2: apply extF; intros i1; rewrite fct_plus_eq; apply comb_lin_plus_r.
+rewrite comb_lin_plus_r -(plus_zero_r 0); f_equal.
+apply comb_lin_zero_compat_r; rewrite -HG1; apply extF; intros i1.
+rewrite mapF_correct; easy.
+generalize (comb_linT_sym_R L1 L2 (fun i1 => A i1 --> G2));
+    rewrite comb_lin_flipT_r; move=>> <-.
+apply comb_lin_zero_compat_r.
+rewrite opp_zero_equiv -comb_lin_opp_r (comb_lin_eq_r _ _ (vectF G2 A)); try easy.
+apply extT; intros i1 i2; rewrite !fct_opp_eq -vect_sym; easy.
+Qed.
+
+End Barycenter_R_Facts.
+
+
+Section Parallelogram.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Definition parallelogram (A B C D : E) : Prop := A --> B = D --> C.
+
+Lemma parallelogram_shift :
+  forall A B C D, parallelogram A B C D <-> parallelogram B C D A.
+                     (* A --> B = D --> C <-> B --> C = A --> D *)
+Proof.
+intros A B C D; unfold parallelogram.
+rewrite -(vect_chasles D B C) -(vect_chasles B A D) (plus_comm (B --> D)).
+rewrite -plus_compat_r_equiv; easy.
+Qed.
+
+Lemma parallelogram_rev :
+  forall A B C D, parallelogram A B C D <-> parallelogram D C B A.
+                     (* A --> B = D --> C <-> D --> C = A --> B *)
+Proof. easy. Qed.
+
+Lemma parallelogram_relation :
+  forall A B C D, parallelogram A B C D <-> parallelogram A D C B.
+                     (* A --> B = D --> C <-> A --> D = B --> C *)
+Proof. intros; rewrite parallelogram_rev -parallelogram_shift; easy. Qed.
+
+Lemma parallelogram_transl_vect :
+  forall O A B, parallelogram O A B (O +++ (A --> B)).
+Proof.
+intros; apply parallelogram_relation.
+unfold parallelogram; rewrite transl_correct_r; easy.
+Qed.
+
+Lemma parallelogram_equiv_def :
+  forall A B C D,
+    invertible (2 : K) -> parallelogram A B C D <-> middle A C = middle B D.
+Proof.
+intros A B C D H; unfold parallelogram; rewrite (middle_correct_equiv _ _ _ H).
+rewrite -(vect_chasles A (middle _ _) B) -(vect_chasles C (middle _ _) D).
+rewrite plus_assoc (plus_comm (_ --> A)) -(plus_assoc (A --> B)).
+rewrite (middle_correct _ _ H) plus_zero_r (vect_sym D) plus_is_zero_l_equiv; easy.
+Qed.
+
+Lemma parallelogram_equiv_def_barycenter :
+  forall A B C D,
+    parallelogram A B C D <->
+    D = barycenter (tripleF 1 (- (1 : K)) 1) (tripleF A B C).
+Proof.
+intros A B C D; unfold parallelogram.
+rewrite barycenter_tripleF_equiv; try apply sum_alt_ones_3_invertible.
+rewrite scal_opp_l !scal_one -vect_sym (plus_comm (D --> A)) vect_chasles
+    plus_comm (vect_sym B) minus_zero_equiv; easy.
+Qed.
+
+Lemma transl_vect_barycenter :
+  forall (O A B : E),
+    O +++ (A --> B) = barycenter (tripleF 1 (- (1 : K)) 1) (tripleF O A B).
+Proof.
+intros; apply parallelogram_equiv_def_barycenter, parallelogram_transl_vect.
+Qed.
+
+Lemma transl_vect_barycenter_alt :
+  forall (O A B : E),
+    B = barycenter (tripleF 1 (- (1 : K)) 1) (tripleF A O (O +++ (A --> B))).
+Proof.
+intros; apply parallelogram_equiv_def_barycenter, parallelogram_shift.
+unfold parallelogram; rewrite transl_correct_r; easy.
+Qed.
+
+Lemma transl_plus_barycenter :
+  forall (O : E) u v,
+    O +++ (u + v) =
+      barycenter (tripleF 1 (- (1 : K)) 1) (tripleF (O +++ v) O (O +++ u)).
+Proof.
+intros; apply parallelogram_equiv_def_barycenter, parallelogram_shift.
+unfold parallelogram; rewrite plus_comm -transl_assoc 2!transl_correct_r; easy.
+Qed.
+
+Lemma transl_plus_barycenter_alt :
+  forall (O : E) u v,
+    O +++ u =
+      barycenter (tripleF 1 (- (1 : K)) 1) (tripleF O (O +++ v) (O +++ (u + v))).
+Proof.
+intros; apply parallelogram_equiv_def_barycenter, parallelogram_shift.
+unfold parallelogram. rewrite plus_comm -transl_assoc 2!transl_correct_r; easy.
+Qed.
+
+Lemma transl_scal_barycenter :
+  forall (O : E) a u,
+    O +++ scal a u = barycenter (coupleF (1 - a) a) (coupleF O (O +++ u)).
+Proof.
+intros O a u; apply (barycenter_correct_orig_equiv O);
+    rewrite sum_coupleF plus_minus_l.
+apply invertible_one.
+rewrite scal_one vectF_coupleF comb_lin_coupleF
+    vect_zero scal_zero_r plus_zero_l.
+apply vect_transl_scal.
+Qed.
+
+End Parallelogram.
+
+
+Section AffineSpace_Morphism_Def.
+
+(** Morphisms between affine spaces are usually called affine mappings. *)
+
+Context {K : Ring}.
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Definition fct_lm (f : E1 -> E2) (O1 : E1) : V1 -> V2 :=
+  fun u1 => f O1 --> f (O1 +++ u1).
+
+Definition f_vect_compat_gen (f : E1 -> E2) (lf : V1 -> V2) : Prop :=
+  forall A1 B1, lf (A1 --> B1) = f A1 --> f B1.
+
+Definition f_transl_compat_gen (f : E1 -> E2) (lf : V1 -> V2) : Prop :=
+  forall A1 u1, f (A1 +++ u1) = f A1 +++ lf u1.
+
+Definition f_vect_compat f O1 := f_vect_compat_gen f (fct_lm f O1).
+Definition f_transl_compat f O1 := f_transl_compat_gen f (fct_lm f O1).
+
+Definition f_vectF_compat_gen (f : E1 -> E2) (lf : V1 -> V2) : Prop :=
+  forall n O (A1 : 'E1^n), mapF lf (vectF O A1) = vectF (f O) (mapF f A1).
+
+Definition f_translF_compat_gen (f : E1 -> E2) (lf : V1 -> V2) : Prop :=
+  forall n (A1 : E1) (u1 : 'V1^n),
+    mapF f (translF A1 u1) = translF (f A1) (mapF lf u1).
+
+Definition f_vectF_compat f O1 := f_vectF_compat_gen f (fct_lm f O1).
+Definition f_translF_compat f O1 := f_translF_compat_gen f (fct_lm f O1).
+
+(* The compatibility of functions with barycenters is denoted 'is_affine_mapping'
+ instead of something as 'f_baryc_compat'. *)
+Definition is_affine_mapping (f : E1 -> E2) : Prop :=
+  forall n L (A1 : 'E1^n),
+    invertible (sum L) -> f (barycenter L A1) = barycenter L (mapF f A1).
+
+Lemma fct_lm_ext :
+  forall (f g : E1 -> E2), same_fun f g -> fct_lm f = fct_lm g.
+Proof. move=>> /fun_ext H; subst; easy. Qed.
+
+Lemma is_affine_mapping_ext :
+  forall (f g : E1 -> E2),
+    same_fun f g -> is_affine_mapping f -> is_affine_mapping g.
+Proof. move=>> /fun_ext H; subst; easy. Qed.
+
+End AffineSpace_Morphism_Def.
+
+
+Section AffineSpace_Morphism_Facts0.
+
+Context {V : ModuleSpace R_Ring}.
+Context {E : AffineSpace V}.
+
+Lemma barycenter_is_affine_mapping :
+  forall {n} L,
+    invertible (sum L) -> is_affine_mapping (fun B : 'E^n => barycenter L B).
+Proof.
+move=>> HL; move=>> HM; apply barycenter_comm_R; apply invertible_equiv_R; easy.
+Qed.
+
+End AffineSpace_Morphism_Facts0.
+
+
+Section AffineSpace_Morphism_Facts1.
+
+Context {K : Ring}.
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Lemma f_vect_transl_compat_gen_equiv :
+  forall {f : E1 -> E2} {lf},
+    f_vect_compat_gen f lf <-> f_transl_compat_gen f lf.
+Proof.
+intros; split; intros Hf; move=>>.
+rewrite transl_vect_equiv -Hf transl_correct_r; easy.
+rewrite -transl_vect_equiv -Hf transl_correct_l; easy.
+Qed.
+
+Lemma f_vect_transl_compat_equiv :
+  forall {f : E1 -> E2} {O1}, f_vect_compat f O1 <-> f_transl_compat f O1.
+Proof. intros; apply f_vect_transl_compat_gen_equiv. Qed.
+
+Lemma f_plus_transl_compat :
+  forall {f : E1 -> E2} {O1},
+    f_plus_compat (fct_lm f O1) -> f_transl_compat f O1.
+Proof.
+intros f O1 Hf A1 u1; apply (vect_l_inj (f O1)).
+rewrite -(transl_correct_l O1 A1) transl_assoc vect_transl_assoc; apply Hf.
+Qed.
+
+Lemma f_transl_plus_compat :
+  forall {f : E1 -> E2} O1,
+    f_transl_compat f O1 -> f_plus_compat (fct_lm f O1).
+Proof.
+intros f O1 Hf u1 v1; unfold fct_lm; apply (transl_l_inj (f O1)).
+rewrite -transl_assoc -(transl_assoc (f O1)) !transl_correct_l; apply Hf.
+Qed.
+
+Lemma f_transl_plus_compat_equiv :
+  forall {f : E1 -> E2} O1,
+    f_transl_compat f O1 <-> f_plus_compat (fct_lm f O1).
+Proof.
+intros; split; [apply f_transl_plus_compat | apply f_plus_transl_compat].
+Qed.
+
+Lemma f_plus_vect_compat :
+  forall {f : E1 -> E2} {O1},
+    f_plus_compat (fct_lm f O1) -> f_vect_compat f O1.
+Proof.
+move=>>; rewrite f_vect_transl_compat_equiv; apply f_plus_transl_compat.
+Qed.
+
+Lemma f_vect_plus_compat :
+  forall {f : E1 -> E2} O1, f_vect_compat f O1 -> f_plus_compat (fct_lm f O1).
+Proof. move=>> /f_vect_transl_compat_equiv; apply f_transl_plus_compat. Qed.
+
+Lemma f_vect_plus_compat_equiv :
+  forall {f : E1 -> E2} O1, f_vect_compat f O1 <-> f_plus_compat (fct_lm f O1).
+Proof.
+intros; rewrite f_vect_transl_compat_equiv; apply f_transl_plus_compat_equiv.
+Qed.
+
+Lemma fct_lm_orig_indep :
+  forall {f : E1 -> E2} O1 O1',
+    f_plus_compat (fct_lm f O1) -> fct_lm f O1 = fct_lm f O1'.
+Proof.
+unfold fct_lm; intros f O1 O1' Hlf; apply fun_ext; intros u1.
+rewrite -(vect_chasles (f O1) (f O1')) -{2}(transl_correct_l O1 O1').
+rewrite transl_assoc Hlf transl_correct_l plus_assoc.
+rewrite (vect_sym (f O1')) plus_opp_l plus_zero_l; easy.
+Qed.
+
+Lemma f_vect_compat_orig_indep :
+  forall {f : E1 -> E2} O1 O1', f_vect_compat f O1 -> f_vect_compat f O1'.
+Proof.
+intros f O1 O1' Hf; generalize (f_vect_plus_compat O1 Hf); intros Hlf A1 B1.
+rewrite -(fct_lm_orig_indep O1); easy.
+Qed.
+
+Lemma f_transl_compat_orig_indep :
+  forall {f : E1 -> E2} O1 O1', f_transl_compat f O1 -> f_transl_compat f O1'.
+Proof.
+intros f O1 O1' Hf; generalize (f_transl_plus_compat O1 Hf); intros Hlf A1 u1.
+rewrite -(fct_lm_orig_indep O1); easy.
+Qed.
+
+Lemma f_transl_compat_gen_flm_uniq :
+  forall {f : E1 -> E2} {lf} {O1 : E1},
+    f_plus_compat lf -> f_transl_compat_gen f lf -> lf = fct_lm f O1.
+Proof.
+move=>> Hlf Hf; apply fun_ext; intro; unfold fct_lm.
+rewrite Hf transl_correct_r; easy.
+Qed.
+
+Lemma f_vect_compat_gen_flm_uniq :
+  forall {f : E1 -> E2} {lf} {O1 : E1},
+    f_plus_compat lf -> f_vect_compat_gen f lf -> lf = fct_lm f O1.
+Proof.
+move=>>; rewrite f_vect_transl_compat_gen_equiv.
+apply f_transl_compat_gen_flm_uniq.
+Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (27/02/2024): add properties of f_vectF_compat*/f_translF_compat*. *)
+
+Lemma is_affine_mapping_is_linear_mapping :
+  forall {f : E1 -> E2} O1,
+    is_affine_mapping f -> is_linear_mapping (fct_lm f O1).
+Proof.
+intros f O1 Hf; unfold fct_lm; split.
+(* *)
+intros u1 v1; generalize (@sum_alt_ones_3_invertible K); intros H3.
+rewrite transl_plus_barycenter Hf; try easy.
+rewrite mapF_tripleF -(scal_one (_ --> barycenter _ _)) -{1}(sum_alt_ones_3).
+rewrite barycenter_correct_orig; try easy.
+rewrite vectF_tripleF comb_lin_tripleF.
+rewrite vect_zero scal_zero_r plus_zero_r 2!scal_one plus_comm; easy.
+(* *)
+intros u1 a; generalize (sum_unit_partition_1_invertible a); intros Ha.
+rewrite transl_scal_barycenter Hf; try easy.
+rewrite mapF_coupleF -(scal_one (_ --> barycenter _ _))
+    -{1}(sum_unit_partition_1 a).
+rewrite barycenter_correct_orig; try easy.
+rewrite vectF_coupleF comb_lin_coupleF vect_zero scal_zero_r plus_zero_l; easy.
+Qed.
+
+Lemma is_affine_mapping_is_linear_mapping_rev :
+  forall {f : E1 -> E2} O1,
+    is_linear_mapping (fct_lm f O1) -> is_affine_mapping f.
+Proof.
+intros f O1 Hf.
+generalize (is_linear_mapping_comb_lin _ Hf); intros Hf'.
+intros n L A1 HL; apply (barycenter_correct_orig_equiv (f O1)); try easy.
+rewrite -(transl_correct_l O1 (barycenter _ _)).
+fold (fct_lm f O1 (O1 --> barycenter L A1)).
+rewrite -(proj2 Hf) (barycenter_correct_orig _ _ HL) Hf' vectF_mapF.
+apply comb_lin_ext_r; intros i.
+unfold fct_lm; rewrite !mapF_correct transl_correct_l; easy.
+Qed.
+
+Lemma is_affine_mapping_equiv :
+  forall {f : E1 -> E2} O1,
+    is_affine_mapping f <-> is_linear_mapping (fct_lm f O1).
+Proof.
+intros; split.
+apply is_affine_mapping_is_linear_mapping.
+apply is_affine_mapping_is_linear_mapping_rev.
+Qed.
+
+Lemma is_affine_mapping_flm_orig_indep :
+  forall {f : E1 -> E2} {O1 O1'},
+    is_affine_mapping f -> fct_lm f O1 = fct_lm f O1'.
+Proof.
+intros; apply fct_lm_orig_indep, is_linear_mapping_plus,
+    is_affine_mapping_equiv; easy.
+Qed.
+
+Lemma is_affine_mapping_vect :
+  forall {f : E1 -> E2} O1, is_affine_mapping f -> f_vect_compat f O1.
+Proof.
+move=> f O1 /(is_affine_mapping_equiv O1) Hf A1 B1.
+rewrite (fct_lm_orig_indep _ A1).
+unfold fct_lm; rewrite transl_correct_l; easy.
+apply is_linear_mapping_plus; easy.
+Qed.
+
+Lemma is_affine_mapping_transl :
+  forall {f : E1 -> E2} O1, is_affine_mapping f -> f_transl_compat f O1.
+Proof.
+intros; apply f_vect_transl_compat_equiv, is_affine_mapping_vect; easy.
+Qed.
+
+Lemma fct_lm_bijective_compat :
+  forall {f : E1 -> E2} O1, bijective f -> bijective (fct_lm f O1).
+Proof.
+intros f O1 [g Hg1 Hg2]; apply Bijective with (fct_lm g (f O1)); unfold fct_lm.
+intro; rewrite transl_correct_l !Hg1 transl_correct_r; easy.
+intro; rewrite Hg1 transl_correct_l Hg2 transl_correct_r; easy.
+Qed.
+
+Lemma is_affine_mapping_bijective_equiv :
+  forall {f : E1 -> E2} O1,
+    is_affine_mapping f -> bijective f <-> bijective (fct_lm f O1).
+Proof.
+intros f O1 Hf1; split; try apply fct_lm_bijective_compat.
+intros [lg Hlg1 Hlg2];
+    apply Bijective with (fun A2 => O1 +++ lg (f O1 --> A2)).
+intro; rewrite -(is_affine_mapping_vect O1 Hf1) Hlg1 transl_correct_l; easy.
+intro; rewrite (is_affine_mapping_transl O1 Hf1) Hlg2 transl_correct_l; easy.
+Qed.
+
+End AffineSpace_Morphism_Facts1.
+
+
+Section AffineSpace_Morphism_Facts2.
+
+Context {K : Ring}.
+Context {V1 V2 V3 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+Context {E3 : AffineSpace V3}.
+
+Lemma fct_lm_comp :
+  forall {f12 : E1 -> E2} {f23 : E2 -> E3} O1,
+    fct_lm (f23 \o f12) O1 = (fct_lm f23 (f12 O1)) \o (fct_lm f12 O1).
+Proof.
+intros; apply fun_ext; intro.
+unfold fct_lm; rewrite !comp_correct transl_correct_l; easy.
+Qed.
+
+Lemma is_affine_comp :
+  forall {f12 : E1 -> E2} {f23 : E2 -> E3},
+    is_affine_mapping f12 -> is_affine_mapping f23 ->
+    is_affine_mapping (f23 \o f12).
+Proof.
+intros f12 f23; pose (O1 := point_of_as E1).
+move=> /(is_affine_mapping_equiv O1) H12
+    /(is_affine_mapping_equiv (f12 O1)) H23.
+apply (is_affine_mapping_equiv O1).
+rewrite fct_lm_comp; apply is_linear_mapping_comp; easy.
+Qed.
+
+End AffineSpace_Morphism_Facts2.
+
+
+Section AffineSpace_Morphism_Facts3.
+
+Context {K : Ring}.
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Lemma fct_lm_inv :
+  forall {f : E1 -> E2} (Hf : bijective f) O1,
+    fct_lm (f_inv Hf) (f O1) = f_inv (fct_lm_bijective_compat O1 Hf).
+Proof.
+intros f Hf O1; apply fun_ext; intros u2.
+rewrite -{1}(f_inv_can_r (fct_lm_bijective_compat O1 Hf) u2); unfold fct_lm.
+rewrite transl_correct_l !f_inv_can_l transl_correct_r; easy.
+Qed.
+
+Lemma is_affine_mapping_bij_compat :
+  forall {f : E1 -> E2} (Hf : bijective f),
+    is_affine_mapping f -> is_affine_mapping (f_inv Hf).
+Proof.
+intros f Hf; pose (O1 := point_of_as E1).
+rewrite (is_affine_mapping_equiv O1) (is_affine_mapping_equiv (f O1)) fct_lm_inv.
+apply is_linear_mapping_bij_compat.
+Qed.
+
+End AffineSpace_Morphism_Facts3.
+
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (02/10/2023): define and add results about convexity and convex hull. *)
+
+
+Section ModuleSpace_AffineSpace.
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Lemma vectF_ms_eq : forall {n} (O : E) (A : 'E^n), vectF O A = A - constF n O.
+Proof. easy. Qed.
+
+Lemma translF_ms_eq :
+  forall {n} (O : E) (u : 'E^n), translF O u = constF n O + u.
+Proof.
+intros; apply extF; intro; rewrite translF_correct ms_transl_eq; easy.
+Qed.
+
+Lemma barycenter_ms_eq :
+  forall {n} L (A : 'E^n), sum L = 1 -> barycenter L A = comb_lin L A.
+Proof.
+intros n L A HL; apply eq_sym, barycenter_correct_equiv.
+rewrite HL; apply invertible_one.
+unfold is_comb_aff; rewrite vectF_ms_eq comb_lin_minus_r.
+rewrite minus_zero_equiv -scal_sum_l HL scal_one; easy.
+Qed.
+
+End ModuleSpace_AffineSpace.
+
+
+Section ModuleSpace_AffineSpace_Morphism.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Lemma fct_lm_ms_eq :
+  forall {f : E1 -> E2} {O1 u1}, fct_lm f O1 u1 = f (O1 + u1) - f O1.
+Proof.
+intros f O1 u1; unfold fct_lm; rewrite ms_vect_eq ms_transl_eq; easy.
+Qed.
+
+Lemma fct_lm_ms_eq_ex:
+  forall {f lf : E1 -> E2},
+    is_linear_mapping lf ->
+    lf = fct_lm f 0 <-> exists c2, f = lf + (fun=> c2).
+Proof.
+intros f lf Hlf; split;
+    [intros Hf; exists (f 0) | intros [c2 Hf]]; apply fun_ext; intro.
+rewrite Hf fct_plus_eq fct_lm_ms_eq plus_zero_l plus_minus_l; easy.
+rewrite fct_lm_ms_eq Hf !fct_plus_eq (is_linear_mapping_zero lf Hlf)
+    !plus_zero_l minus_plus_r; easy.
+Qed.
+
+Lemma fct_lm_ms_cst_reg :
+  forall (f : E1 -> E2) (c2 : E2) (O1 : E1),
+    fct_lm (f + (fun=> c2)) O1 = fct_lm f O1.
+Proof.
+intros; apply fun_ext; intro.
+rewrite !fct_lm_ms_eq !fct_plus_eq -minus_plus_r_eq; easy.
+Qed.
+
+Lemma fct_lm_ms_lin :
+  forall {f : E1 -> E2} (O1 : E1), is_linear_mapping f -> fct_lm f O1 = f.
+Proof.
+move=>> [Hf _]; apply fun_ext; intro.
+rewrite fct_lm_ms_eq Hf minus_plus_l; easy.
+Qed.
+
+Lemma is_affine_mapping_ms_equiv :
+  forall {f : E1 -> E2},
+    is_affine_mapping f <->
+    exists lf c2, is_linear_mapping lf /\ f = lf + (fun=> c2).
+Proof.
+intros f; split.
+(* *)
+intros Hf; exists (fct_lm f 0), (f 0); split.
+apply is_affine_mapping_equiv; easy.
+apply fun_ext; intros u; rewrite fct_plus_eq plus_minus_l_equiv
+    -{2}(plus_zero_l u) -ms_transl_eq -ms_vect_eq; easy.
+(* *)
+intros [lf [c2 [Hlf1 Hlf2]]]; apply (is_affine_mapping_equiv (0 : E1)).
+apply (is_linear_mapping_ext lf); try easy.
+intros u; unfold fct_lm; rewrite Hlf2 ms_transl_eq ms_vect_eq !fct_plus_eq.
+rewrite (is_linear_mapping_zero lf Hlf1) !plus_zero_l minus_plus_r; easy.
+Qed.
+
+Lemma is_linear_affine_mapping_ms :
+  forall {lf : E1 -> E2} c2,
+    is_linear_mapping lf -> is_affine_mapping (lf + (fun=> c2)).
+Proof.
+intros lf c2 Hf; rewrite is_affine_mapping_ms_equiv; exists lf, c2; easy.
+Qed.
+
+Lemma is_affine_linear_mapping_ms :
+  forall {f : E1 -> E2},
+    is_affine_mapping f -> is_linear_mapping (f - (fun=> f 0)).
+Proof.
+move=>> /is_affine_mapping_ms_equiv [lf [c2 [Hlf1 Hlf2]]].
+apply is_linear_mapping_ext with lf; try easy.
+intro; rewrite Hlf2 fct_minus_eq !fct_plus_eq (is_linear_mapping_zero _ Hlf1)
+    plus_zero_l minus_plus_r; easy.
+Qed.
+
+Lemma is_affine_mapping_comb_aff_compat :
+  forall {n} {f : E1 -> E2} {L} {B : 'E1^n},
+    is_affine_mapping f -> sum L = 1 ->
+    f (comb_lin L B) = comb_lin L (fun i => f (B i)).
+Proof.
+move=>> /is_affine_mapping_ms_equiv [lf [c [Hlf1 Hlf2]]] HL.
+rewrite Hlf2 fct_plus_eq linear_mapping_comb_lin_compat; try easy.
+rewrite comb_lin_plus_r; f_equal.
+rewrite -{1}(scal_one c) -HL scal_sum_l; easy.
+Qed.
+
+End ModuleSpace_AffineSpace_Morphism.
+
+
+Section Compatible_AffineSpace_Def.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Definition vectP (PE : E -> Prop) (O : E) : V -> Prop :=
+  preimage (transl O) PE. (* = fun u => PE (O +++ u). *)
+
+Definition translP (PV : V -> Prop) (O : E) : E -> Prop :=
+  preimage (vect O) PV. (* = fun A => PV (O --> A). *)
+
+Definition vect_closed_gen (PE : E -> Prop) (PV : V -> Prop) : Prop :=
+  forall O A, PE O -> PE A <-> translP PV O A. (* <-> PV (O --> A). *)
+
+Definition transl_closed_gen (PE : E -> Prop) (PV : V -> Prop) : Prop :=
+  forall O u, PE O -> PV u <-> vectP PE O u. (* <-> PE (O +++ u). *)
+
+Definition vect_closed PE O := vect_closed_gen PE (vectP PE O).
+Definition transl_closed PE O := transl_closed_gen PE (vectP PE O).
+
+(* Gostiaux T4, Def. 1.23 p. 9.
+  compatible_as is actually barycenter_closed. *)
+Definition compatible_as (PE : E -> Prop) : Prop :=
+  forall n L (A : 'E^n),
+    invertible (sum L) -> inclF A PE -> PE (barycenter L A).
+
+Definition span_as (gen : E -> Prop) : E -> Prop :=
+  span_struct compatible_as gen.
+
+Inductive barycenter_closure (gen : E -> Prop) : E -> Prop :=
+  | Barycenter_closure :
+    forall n L (A : 'E^n),
+      invertible (sum L) -> inclF A gen ->
+      barycenter_closure gen (barycenter L A).
+
+End Compatible_AffineSpace_Def.
+
+
+Section Compatible_AffineSpace_Facts.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma vectP_correct :
+  forall (PE : E -> Prop) O, PE O -> PE = image (transl O) (vectP PE O).
+Proof. intros; rewrite image_of_preimage transl_l_fullset inter_full_r //. Qed.
+
+Lemma vectP_eq : forall (PE : E -> Prop) O, vectP PE O = image (vect O) PE.
+Proof.
+intros; apply preimage_eq; [apply transl_correct_r | apply transl_correct_l].
+Qed.
+
+Lemma vectP_inj :
+  forall (PE QE : E -> Prop) O, vectP PE O = vectP QE O -> PE = QE.
+Proof. move=>>; eapply preimage_inj, transl_correct_l. Qed.
+
+Lemma vectP_full :
+  forall (PE : E -> Prop) O, full PE -> full (vectP PE O).
+Proof. intros; apply preimage_full_equiv; easy. Qed.
+
+Lemma vectP_fullset :
+  forall (O : E), vectP fullset O = fullset.
+Proof. easy. Qed.
+
+Lemma vectP_inter :
+  forall (PE1 PE2 : E -> Prop) O,
+    vectP (inter PE1 PE2) O = inter (vectP PE1 O) (vectP PE2 O).
+Proof. easy. Qed.
+
+Lemma vectP_inter_any :
+  forall {Idx : Type} (PE : Idx -> E -> Prop) O,
+    vectP (inter_any PE) O = inter_any (fun idx => vectP (PE idx) O).
+Proof. easy. Qed.
+
+Lemma vectP_zero_closed_equiv :
+  forall (PE : E -> Prop) O, zero_closed (vectP PE O) <-> PE O.
+Proof. intros PE O; rewrite -{2}(transl_zero O); easy. Qed.
+
+Lemma translP_correct :
+  forall (PV : V -> Prop) (O : E), PV = image (vect O) (translP PV O).
+Proof. intros; rewrite image_of_preimage vect_l_fullset inter_full_r //. Qed.
+
+Lemma translP_eq :
+  forall (PV : V -> Prop) (O : E), translP PV O = image (transl O) PV.
+Proof.
+intros; apply preimage_eq; [apply transl_correct_l | apply transl_correct_r].
+Qed.
+
+Lemma translP_inj :
+  forall (PV QV : V -> Prop) (O : E), translP PV O = translP QV O -> PV = QV.
+Proof. move=>>; eapply preimage_inj, transl_correct_r. Qed.
+
+Lemma translP_full :
+  forall (PV : V -> Prop) (O : E), full PV -> full (translP PV O).
+Proof. intros; apply preimage_full_equiv; easy. Qed.
+
+Lemma translP_fullset :
+  forall (O : E), translP fullset O = fullset.
+Proof. easy. Qed.
+
+Lemma translP_inter :
+  forall (PV1 PV2 : V -> Prop) (O : E),
+    translP (inter PV1 PV2) O = inter (translP PV1 O) (translP PV2 O).
+Proof. easy. Qed.
+
+Lemma translP_inter_any :
+  forall {Idx : Type} (PV : Idx -> V -> Prop) (O : E),
+    translP (inter_any PV) O = inter_any (fun idx => translP (PV idx) O).
+Proof. easy. Qed.
+
+Lemma translP_zero_closed_equiv :
+  forall (PV : V -> Prop) (O : E), zero_closed PV <-> translP PV O O.
+Proof. intros; unfold translP, preimage; rewrite vect_zero; easy. Qed.
+
+Lemma vectP_translP :
+  forall (PV : V -> Prop) (O : E), vectP (translP PV O) O = PV.
+Proof. intros; apply preimage_cancel, transl_correct_r. Qed.
+
+Lemma translP_vectP :
+  forall (PE : E -> Prop) O, translP (vectP PE O) O = PE.
+Proof. intros; apply preimage_cancel, transl_correct_l. Qed.
+
+Lemma vect_transl_closed_gen_equiv :
+  forall {PE : E -> Prop} {PV}, vect_closed_gen PE PV <-> transl_closed_gen PE PV.
+Proof.
+intros; split; intros HPE A; unfold vectP, translP, preimage.
+intros u HA; rewrite -{1}(transl_correct_r A u) (HPE A); easy.
+intros B HA; rewrite -{1}(transl_correct_l A B) (HPE A); easy.
+Qed.
+
+Lemma vect_transl_closed_equiv :
+  forall {PE : E -> Prop} {O}, vect_closed PE O <-> transl_closed PE O.
+Proof. intros; apply vect_transl_closed_gen_equiv. Qed.
+
+Lemma transl_plus_closed :
+  forall {PE : E -> Prop} {O}, transl_closed PE O -> plus_closed (vectP PE O).
+Proof.
+move=>> HPE u v; unfold vectP, preimage; rewrite -transl_assoc; apply HPE; easy.
+Qed.
+
+Lemma vect_plus_closed :
+  forall {PE : E -> Prop} {O}, vect_closed PE O -> plus_closed (vectP PE O).
+Proof. move=>> /vect_transl_closed_equiv; apply transl_plus_closed. Qed.
+
+Lemma compatible_ms_transl :
+  forall {PE : E -> Prop} {O}, PE O ->
+    compatible_ms (vectP PE O) -> transl_closed PE O.
+Proof.
+unfold vectP; intros PE O HO HPE A u; unfold vectP, preimage.
+rewrite -(transl_correct_l O A) transl_assoc iff_sym_equiv.
+apply (compatible_ms_plus_equiv HPE).
+Qed.
+
+Lemma compatible_ms_vect :
+  forall {PE : E -> Prop} {O}, PE O ->
+    compatible_ms (vectP PE O) -> vect_closed PE O.
+Proof.
+unfold vect_closed.
+move=>>; rewrite vect_transl_closed_gen_equiv; apply compatible_ms_transl.
+Qed.
+
+Lemma vectP_orig_indep :
+  forall {PE : E -> Prop} {O} (O' : E), PE O ->
+    compatible_ms (vectP PE O') -> vectP PE O = vectP PE O'.
+Proof.
+intros; apply subset_ext; intros u.
+apply iff_sym, compatible_ms_transl; try easy.
+apply vectP_zero_closed_equiv, compatible_ms_zero; easy.
+Qed.
+
+Lemma translP_orig_indep :
+  forall {PV : V -> Prop} {O} (O' : E), translP PV O O' ->
+    compatible_ms PV -> translP PV O = translP PV O'.
+Proof.
+intros PV O O' HO' HPV; apply subset_ext; intros A.
+rewrite -{2}(vectP_translP PV O); apply compatible_ms_vect; try easy.
+apply translP_zero_closed_equiv, compatible_ms_zero; easy.
+rewrite vectP_translP; easy.
+Qed.
+
+Lemma vect_closed_orig_indep :
+  forall {PE : E -> Prop} O O',
+    PE O -> compatible_ms (vectP PE O') ->
+    vect_closed PE O -> vect_closed PE O'.
+Proof.
+move=>> HO HPV HPE; move=>>; rewrite -(vectP_orig_indep _ HO); auto.
+Qed.
+
+Lemma transl_closed_orig_indep :
+  forall {PE : E -> Prop} O O',
+    PE O -> compatible_ms (vectP PE O') ->
+    transl_closed PE O -> transl_closed PE O'.
+Proof.
+move=>> HO HPV HPE; move=>>; rewrite -(vectP_orig_indep _ HO); auto.
+Qed.
+
+Lemma transl_closed_gen_sms_uniq :
+  forall {PE : E -> Prop} {PV} {O : E},
+    PE O -> compatible_ms PV -> transl_closed_gen PE PV -> PV = vectP PE O.
+Proof.
+move=>> HO HPV HPE; apply subset_ext_equiv; split; intro;
+    [apply HPE | apply (HPE _ _ HO)]; easy.
+Qed.
+
+Lemma vect_closed_gen_sms_uniq :
+  forall {PE : E -> Prop} {PV} {O : E},
+    PE O -> compatible_ms PV -> vect_closed_gen PE PV -> PV = vectP PE O.
+Proof.
+move=>>; rewrite vect_transl_closed_gen_equiv; apply transl_closed_gen_sms_uniq.
+Qed.
+
+Lemma compatible_as_empty :
+  forall {PE : E -> Prop}, nonzero_struct K -> empty PE -> compatible_as PE.
+Proof.
+intros PE HK HPE n L A HL HA; destruct n.
+contradict HL; apply sum_nil_noninvertible; easy.
+contradict HA; rewrite not_all_ex_not_equiv; exists ord0.
+intros HA0; apply (HPE _ HA0).
+Qed.
+
+Lemma compatible_as_emptyset :
+  forall {PE : E -> Prop}, nonzero_struct K -> compatible_as (@emptyset E).
+Proof. intros; apply compatible_as_empty; easy. Qed.
+
+Lemma compatible_as_singleton : forall (O : E), compatible_as (singleton O).
+Proof. move=>> HL /extF ->; apply barycenter_constF; easy. Qed.
+
+Lemma compatible_as_unit :
+  forall {PE : E -> Prop},
+    nonzero_struct K -> is_unit_type E -> compatible_as PE.
+Proof.
+intros PE HK [O HE].
+destruct (empty_dec PE) as [HPE | HPE].
+apply compatible_as_empty; easy.
+rewrite (unit_subset_is_singleton PE HE HPE).
+apply compatible_as_singleton.
+Qed.
+
+Lemma compatible_as_full : forall {PE : E -> Prop}, full PE -> compatible_as PE.
+Proof. easy. Qed.
+
+Lemma compatible_as_fullset : compatible_as (@fullset E).
+Proof. easy. Qed.
+
+Lemma compatible_ms_as :
+  forall {PE : E -> Prop} {O},
+    PE O -> nonzero_struct K -> invertible (2 : K) ->
+    compatible_ms (vectP PE O) -> compatible_as PE.
+Proof.
+move=>> HO HK1 HK2 HPE n L A HL HA.
+destruct n. contradict HL; apply sum_nil_noninvertible; easy.
+generalize (compatible_ms_vect HO HPE (A ord0)); intros HPE'.
+unfold translP, vectP, preimage in HPE'.
+apply HPE', (compatible_ms_scal_rev HPE (sum L) _ HL); try easy.
+rewrite barycenter_correct_orig; try easy.
+apply compatible_ms_comb_lin; try easy.
+unfold vectP, preimage; intro; rewrite vectF_correct -HPE'; easy.
+Qed.
+
+Lemma compatible_as_scal :
+  forall {PE : E -> Prop} {O},
+    PE O -> compatible_as PE -> scal_closed (vectP PE O).
+Proof.
+intros PE O HO HPE a u Hu; unfold vectP, preimage in *.
+assert (H0 : invertible (sum (coupleF (1 - a) a))).
+  rewrite sum_coupleF -plus_assoc plus_opp_l plus_zero_r.
+  apply invertible_one.
+assert (H : O +++ scal a u =
+    barycenter (coupleF (1 - a) a) (coupleF O (O +++ u))).
+  apply barycenter_coupleF_equiv; try easy.
+  rewrite -{2}(transl_zero O) 2!vect_transl.
+  rewrite 2!scal_minus_distr_l scal_zero_r scal_minus_distr_r scal_one.
+  rewrite minus_zero_l plus_opp_l; easy.
+unfold vectP; rewrite H; apply HPE; try easy.
+intros i; destruct (ord2_dec i) as [Hi | Hi]; rewrite Hi;
+    [rewrite coupleF_0 | rewrite coupleF_1]; easy.
+Qed.
+
+Lemma compatible_as_scal_rev :
+  forall {PE : E -> Prop} {O},
+    PE O -> compatible_as PE -> scal_rev_closed (vectP PE O).
+Proof. intros; apply scal_scal_rev_closed, compatible_as_scal; easy. Qed.
+
+Lemma compatible_as_plus :
+  forall {PE : E -> Prop} {O},
+    PE O -> invertible (2 : K) ->
+    compatible_as PE -> plus_closed (vectP PE O).
+Proof.
+intros PE O HO HK HPE; unfold vectP, preimage.
+unfold two in HK; rewrite -(sum_2 ones) in HK.
+intros u v Hu Hv; pose (A := coupleF (O +++ u) (O +++ v)).
+rewrite (transl_l_eq _ _ (sum (vectF O A))).
+(* *)
+rewrite -comb_lin_ones_l -barycenter_correct_orig; try easy.
+apply (compatible_as_scal_rev HO HPE (inv (@sum _ 2 ones)));
+    try now apply inv_invertible. unfold vectP, preimage.
+rewrite -> scal_assoc, mult_inv_l, scal_one, transl_correct_l; try easy.
+apply HPE; try easy; unfold A;
+    intros i; destruct (ord2_dec i) as [Hi | Hi]; rewrite Hi;
+    [rewrite coupleF_0 | rewrite coupleF_1]; easy.
+(* *)
+unfold A; rewrite vectF_coupleF sum_2 coupleF_0 coupleF_1 !transl_correct_r //.
+Qed.
+
+Lemma compatible_as_ms :
+  forall {PE : E -> Prop} {O},
+    PE O -> invertible (2 : K) ->
+    compatible_as PE -> compatible_ms (vectP PE O).
+Proof.
+intros; apply plus_scal_closed_compatible_ms.
+unfold vectP, preimage; exists 0; rewrite transl_zero; easy.
+apply compatible_as_plus; easy.
+apply compatible_as_scal; easy.
+Qed.
+
+Lemma compatible_as_ms_equiv :
+  forall {PE : E -> Prop} {O},
+    PE O -> nonzero_struct K -> invertible (2 : K) ->
+    compatible_as PE <-> compatible_ms (vectP PE O).
+Proof.
+intros; split; [apply compatible_as_ms | apply compatible_ms_as]; easy.
+Qed.
+
+Lemma compatible_as_sms_orig_indep :
+  forall {PE : E -> Prop} {O O'},
+    PE O -> PE O' -> invertible (2 : K) ->
+    compatible_as PE -> vectP PE O = vectP PE O'.
+Proof. intros; apply vectP_orig_indep; try apply compatible_as_ms; easy. Qed.
+
+Lemma compatible_as_vect :
+  forall {PE : E -> Prop} {O},
+    PE O -> invertible (2 : K) ->
+    compatible_as PE -> vect_closed PE O.
+Proof. intros; apply compatible_ms_vect, compatible_as_ms; easy. Qed.
+
+Lemma compatible_as_transl :
+  forall {PE : E -> Prop} {O},
+    PE O -> invertible (2 : K) ->
+    compatible_as PE -> transl_closed PE O.
+Proof.
+intros; apply vect_transl_closed_gen_equiv, compatible_as_vect; easy.
+Qed.
+
+Lemma compatible_as_inter :
+  forall {PE1 PE2 : E -> Prop},
+    compatible_as PE1 -> compatible_as PE2 -> compatible_as (inter PE1 PE2).
+Proof.
+move=>> HPE1 HPE2; move=>> HL HA; split; [apply HPE1 | apply HPE2];
+    try easy; intro; apply HA.
+Qed.
+
+Lemma compatible_as_inter_any :
+  forall {Idx : Type} {PE : Idx -> E -> Prop},
+    (forall idx, compatible_as (PE idx)) ->
+    compatible_as (inter_any PE).
+Proof.
+move=>> HPE; move=>> HL HA idx; apply HPE; try easy; intro; apply HA.
+Qed.
+
+Lemma span_as_compatible_as :
+  forall (gen : E -> Prop), compatible_as (span_as gen).
+Proof.
+apply span_struct_compatible; move=>>; apply compatible_as_inter_any.
+Qed.
+
+Lemma span_as_incl : forall (gen : E -> Prop), incl gen (span_as gen).
+Proof. apply span_struct_incl. Qed.
+
+Lemma span_as_lub :
+  forall (gen PE : E -> Prop),
+    compatible_as PE -> incl gen PE -> incl (span_as gen) PE.
+Proof. apply span_struct_lub. Qed.
+
+Lemma span_as_full :
+  forall (PE : E -> Prop), compatible_as PE -> span_as PE = PE.
+Proof. apply span_struct_full. Qed.
+
+Lemma barycenter_closure_ex :
+  forall (gen : E -> Prop) A,
+    (exists n L (B : 'E^n),
+      invertible (sum L) /\ inclF B gen /\ A = barycenter L B) ->
+    barycenter_closure gen A.
+Proof.
+move=>> [n [L [B [HL [HB HA]]]]]; rewrite HA; apply Barycenter_closure; easy.
+Qed.
+
+Lemma barycenter_closure_ex_rev :
+  forall (gen : E -> Prop) A,
+    barycenter_closure gen A ->
+    exists n L (B : 'E^n),
+      sum L = 1 /\ (forall i, L i <> 0) /\ inclF B gen /\ A = barycenter L B.
+Proof.
+intros gen A HA; induction HA as [n L B HL HB].
+destruct (barycenter_normalized_n0_ex L B HL)
+    as [n1 [L1 [A1 [HL1a [HL1b [HA1a HA1b]]]]]].
+exists n1, L1, A1; repeat split; try easy.
+apply (inclF_monot_l _ B); easy.
+Qed.
+
+Lemma bc_len_EX :
+  forall (gen : E -> Prop) A,
+    { n | barycenter_closure gen A ->
+      exists L (B : 'E^n),
+        sum L = 1 /\ (forall i, L i <> 0) /\
+        inclF B gen /\ A = barycenter L B }.
+Proof.
+intros gen A; apply ex_EX.
+destruct (classic_dec (barycenter_closure gen A)) as [HA | HA].
+destruct (barycenter_closure_ex_rev _ _ HA) as [n [L [B HLB]]].
+exists n; intros _; exists L, B; easy.
+exists 0; intros HA'; easy.
+Qed.
+
+Definition bc_len (gen : E -> Prop) A := proj1_sig (bc_len_EX gen A).
+
+Lemma bc_EX :
+  forall (gen : E -> Prop) A,
+    { LB : 'K^(bc_len gen A) * 'E^(bc_len gen A) |
+      barycenter_closure gen A ->
+      sum LB.1 = 1 /\ (forall i, LB.1 i <> 0) /\
+      inclF LB.2 gen /\ A = barycenter LB.1 LB.2 }.
+Proof.
+intros gen A; apply ex_EX.
+destruct (classic_dec (barycenter_closure gen A)) as [HA | HA].
+destruct (proj2_sig (bc_len_EX gen A)) as [L [B HLB]]; try easy.
+exists (L, B); intros _; easy.
+destruct (inhabited_as E) as [O].
+exists ((fun _ => 0), (fun _ => O)); intros HA'; easy.
+Qed.
+
+Definition bc_coef (gen : E -> Prop) A := fst (proj1_sig (bc_EX gen A)).
+Definition bc_point (gen : E -> Prop) A := snd (proj1_sig (bc_EX gen A)).
+
+Lemma bc_nLB_correct :
+  forall (gen : E -> Prop) A,
+    barycenter_closure gen A ->
+    sum (bc_coef gen A) = 1 /\
+    (forall i, bc_coef gen A i <> 0) /\
+    inclF (bc_point gen A) gen /\
+    A = barycenter (bc_coef gen A) (bc_point gen A).
+Proof. intros gen A HA; destruct (proj2_sig (bc_EX gen A)); easy. Qed.
+
+Lemma barycenter_closure_exs :
+  forall {gen : E -> Prop} {n} (A : 'E^n),
+    inclF A (barycenter_closure gen) ->
+    exists (b : 'nat^n) L (B : forall i, 'E^(b i)),
+      (forall i, sum (L i) = 1) /\ (forall i j, L i j <> 0) /\
+      (forall i, inclF (B i) gen) /\ A = fun i => barycenter (L i) (B i).
+Proof.
+intros gen n A HA; destruct n.
+(* *)
+destruct (inhabited_as E) as [O]; exists 0, (fun _ => 0), (fun _ _ => O).
+repeat split; try apply extF; intros [i Hi]; easy.
+(* *)
+exists (fun i => bc_len gen (A i)), (fun i => bc_coef gen (A i)),
+    (fun i => bc_point gen (A i)).
+repeat split; try apply extF; intro; apply bc_nLB_correct; easy.
+Qed.
+
+Lemma compatible_as_equiv :
+  forall (PE : E -> Prop), compatible_as PE <-> PE = barycenter_closure PE.
+Proof.
+intros PE; split; intros HPE.
+(* *)
+apply subset_ext_equiv; split; intros A HA; try induction HA; auto.
+rewrite -(barycenter_singleF 1 A (singleF 1) (singleF A)); try easy.
+apply Barycenter_closure; try easy.
+apply sum_singleF_invertible.
+1,2: apply invertible_one.
+(* *)
+move=>> HL HB; rewrite HPE; apply Barycenter_closure; easy.
+Qed.
+
+Lemma barycenter_closure_incl :
+  forall (gen : E -> Prop), incl gen (barycenter_closure gen).
+Proof.
+intros gen A HA.
+rewrite -(barycenter_singleF 1 A (singleF 1) (singleF A));
+    try apply invertible_one; try easy.
+apply Barycenter_closure; try easy.
+rewrite sum_1 singleF_0; apply invertible_one.
+Qed.
+
+End Compatible_AffineSpace_Facts.
+
+
+Section Compatible_AffineSpace_Affine_Mapping_Facts1.
+
+Context {K : Ring}.
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : is_affine_mapping f.
+
+Lemma compatible_as_image : compatible_as PE1 -> compatible_as (image f PE1).
+Proof.
+intro; move=>> HL /inclF_image_equiv [A [HA1 HA2]].
+rewrite HA2 -Hf; try easy; apply Im; auto.
+Qed.
+
+Lemma compatible_as_preimage :
+  compatible_as PE2 -> compatible_as (preimage f PE2).
+Proof. intro; move=>> Hl HA; unfold preimage; rewrite Hf; auto. Qed.
+
+End Compatible_AffineSpace_Affine_Mapping_Facts1.
+
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (27/02/2024):
+ add missing sections similar to Compatible_ModuleSpace_Linear_Mapping_{Def,Facts*}. *)
+
+
+Section Sub_AffineSpace_Def.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_as PE.
+
+Variable O : E.
+Hypothesis HO : PE O.
+
+Definition sub_point_of_as : sub_struct HPE := mk_sub HO.
+
+Let PV := vectP PE O.
+Let HPV : compatible_ms PV := compatible_as_ms HO HK HPE.
+
+Lemma sub_vect_aux : forall {A B}, PE A -> PE B -> PV (A --> B).
+Proof. move=>>; apply compatible_as_vect; easy. Qed.
+
+Lemma sub_vectF_aux :
+  forall {n} {A} {B : 'E^n}, PE A -> inclF B PE -> inclF (vectF A B) PV.
+Proof. intros; intro; apply sub_vect_aux; easy. Qed.
+
+Definition sub_vect (A B : sub_struct HPE) : sub_ms HPV :=
+  mk_sub_ms (sub_vect_aux (in_sub A) (in_sub B)).
+
+Lemma sub_vect_chasles :
+  forall A B C, sub_vect B A + sub_vect A C = sub_vect B C.
+Proof. intros; apply val_inj, vect_chasles. Qed.
+
+Lemma sub_vect_l_bij_ex : forall A u, exists! B, sub_vect A B = u.
+Proof.
+intros [A HA] [u Hu]; unfold PV in Hu.
+assert (Hu' : vectP PE A u) by now rewrite -(compatible_as_transl HO).
+exists (mk_sub_ HPE (A +++ u) Hu'); split.
+apply val_inj, transl_correct_r.
+move=> [B HB1] /(f_equal val) /= HB2.
+apply val_inj; simpl; rewrite -HB2; apply transl_correct_l.
+Qed.
+
+Definition sub_AffineSpace_mixin :=
+  AffineSpace.Mixin _ _ _ sub_point_of_as _ sub_vect_chasles sub_vect_l_bij_ex.
+
+Canonical Structure sub_AffineSpace :=
+  AffineSpace.Pack _ _ sub_AffineSpace_mixin (sub_struct HPE).
+
+Lemma sub_transl_aux : forall {A u}, PE A -> PV u -> PE (A +++ u).
+Proof. move=>>; apply compatible_as_transl; easy. Qed.
+
+Lemma sub_translF_aux :
+  forall {n} {A} {u : 'V^n}, PE A -> inclF u PV -> inclF (translF A u) PE.
+Proof. intros; intro; apply sub_transl_aux; easy. Qed.
+
+Definition sub_transl (A : sub_struct HPE) (u : sub_ms HPV) : sub_struct HPE :=
+  mk_sub (sub_transl_aux (in_sub A) (in_sub u)).
+
+Lemma val_vect : f_vect_compat_gen val val.
+Proof. easy. Qed.
+
+Lemma val_vectF : f_vectF_compat_gen val val.
+Proof. easy. Qed.
+
+Lemma val_transl : f_transl_compat_gen val val.
+Proof.
+intros A u; pose (B := A +++ u).
+assert (HB : u = A --> B) by now unfold B; rewrite transl_correct_r.
+rewrite HB val_vect 2!transl_correct_l; easy.
+Qed.
+
+Lemma val_translF : f_translF_compat_gen val val.
+Proof. move=>>; apply extF; intro; apply val_transl. Qed.
+
+Lemma val_barycenter :
+  forall {n} L (A : '(sub_struct HPE)^n),
+    invertible (sum L) -> val (barycenter L A) = barycenter L (mapF val A).
+Proof.
+intros n L A HL; apply barycenter_correct_equiv; [easy |].
+unfold is_comb_aff; rewrite -val_vectF -val_comb_lin barycenter_correct; easy.
+Qed.
+
+Lemma val_aff_map : is_affine_mapping val.
+Proof. intro; apply val_barycenter. Qed.
+
+Lemma mk_sub_vect :
+  forall {A B : E} (HA : PE A) (HB : PE B),
+    mk_sub HA --> mk_sub HB = mk_sub_ms (sub_vect_aux HA HB).
+Proof. easy. Qed.
+
+Lemma mk_sub_vectF :
+  forall {n} O {A : 'E^n} (HO : PE O) (HA : inclF A PE),
+    vectF (mk_sub HO) (fun i => mk_sub (HA i)) =
+      fun i => mk_sub_ms (sub_vect_aux HO (HA i)).
+Proof. easy. Qed.
+
+Lemma mk_sub_transl :
+  forall {A u} (HA : PE A) (Hu : PV u),
+    mk_sub HA +++ mk_sub_ms Hu = mk_sub (sub_transl_aux HA Hu).
+Proof. intros; apply val_inj, val_transl. Qed.
+
+Lemma mk_sub_translF :
+  forall {n} {A} {u : 'V^n} (HA : PE A) (Hu : inclF u PV),
+    translF (mk_sub HA) (fun i => mk_sub_ms (Hu i)) =
+      fun i => mk_sub (sub_transl_aux HA (Hu i)).
+Proof. intros; apply extF; intro; apply mk_sub_transl. Qed.
+
+Lemma mk_sub_barycenter :
+  forall {n} {L} {A : 'E^n} (HL : invertible (sum L)) (HA : inclF A PE),
+    barycenter L (fun i => mk_sub (HA i)) = mk_sub (HPE _ _ _ HL HA).
+Proof. intros; apply val_inj, val_barycenter; easy. Qed.
+
+Lemma sub_vect_eq :
+  forall (A B : sub_struct HPE),
+    A --> B = mk_sub_ms (sub_vect_aux (in_sub A) (in_sub B)).
+Proof. easy. Qed.
+
+Lemma sub_vectF_eq :
+  forall {n} O (A : '(sub_struct HPE)^n),
+    vectF O A =
+      fun i => mk_sub_ms (sub_vect_aux (in_sub O) (in_sub (A i))).
+Proof. easy. Qed.
+
+Lemma sub_transl_eq :
+  forall (A : sub_struct HPE) (u : sub_ms HPV),
+    A +++ u = mk_sub (sub_transl_aux (in_sub A) (in_sub u)).
+Proof. intros; apply val_inj, val_transl. Qed.
+
+Lemma sub_translF_eq :
+  forall {n} (A : sub_struct HPE) (u : '(sub_ms HPV)^n),
+    translF A u =
+      fun i => mk_sub (sub_transl_aux (in_sub A) (in_sub (u i))).
+Proof. intros; apply extF; intro; apply sub_transl_eq. Qed.
+
+Lemma sub_barycenter_eq :
+  forall {n} {L} {A : '(sub_struct HPE)^n} (HL : invertible (sum L)),
+    barycenter L A = mk_sub (HPE _ _ _ HL (fun i => in_sub (A i))).
+Proof. intros; apply val_inj, val_barycenter; easy. Qed.
+
+End Sub_AffineSpace_Def.
+
+
+Section Sub_AffineSpace_Facts1.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+Context {PEa PEb : E -> Prop}.
+Hypothesis HPE : incl PEa PEb.
+Hypothesis HPEb : compatible_as PEb.
+Let PEb_as := sub_struct HPEb.
+Let PEa' : PEb_as -> Prop := preimage val PEa.
+
+Lemma RgS_val_as_eq : RgS PEa' val = PEa.
+Proof.
+apply RgS_preimage_equiv; intros x Hx; apply Rg_ex;
+    exists (mk_sub (HPE _ Hx)); easy.
+Qed.
+
+(*
+Lemma preimage_val_compatible_as : compatible_as PEa -> compatible_as PEa'.
+Proof. intros; apply compatible_as_preimage; easy. Qed.
+
+Lemma preimage_val_compatible_as_rev : compatible_as PEa' -> compatible_as PEa.
+Proof. intros; rewrite -RgS_val_as_eq; apply RgS_compatible_as; easy. Qed.
+
+Lemma preimage_val_compatible_as_equiv : compatible_as PEa' <-> compatible_as PEa.
+Proof.
+split; [apply preimage_val_compatible_as_rev | apply preimage_val_compatible_as].
+Qed.
+*)
+
+End Sub_AffineSpace_Facts1.
+
+
+Section Sub_AffineSpace_Facts2.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+Context {PEb : E -> Prop}.
+Hypothesis HPEb : compatible_as PEb.
+Let PEb_as := sub_struct HPEb.
+Context {PEa : PEb_as -> Prop}.
+Let PEa' := RgS PEa val.
+
+Lemma preimage_val_as_eq : preimage val PEa' = PEa.
+Proof.
+apply subset_ext_equiv; split; [| apply preimage_RgS].
+intros x Hx; inversion Hx as [y Hy1 Hy2]; rewrite -(val_inj _ _ Hy2); easy.
+Qed.
+
+(*
+Lemma RgS_val_compatible_as : compatible_as PEa -> compatible_as PEa'.
+Proof. intros; apply RgS_compatible_as; easy. Qed.
+
+Lemma RgS_val_compatible_as_rev : compatible_as PEa' -> compatible_as PEa.
+Proof.
+intros; rewrite -preimage_val_as_eq; apply compatible_as_preimage; easy.
+Qed.
+
+Lemma RgS_val_compatible_as_equiv : compatible_as PEa' <-> compatible_as PEa.
+Proof.
+split; [apply RgS_val_compatible_as_rev | apply RgS_val_compatible_as].
+Qed.
+*)
+
+End Sub_AffineSpace_Facts2.
+
+
+Section Sub_AffineSpace_Affine_Mapping_Facts1.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+Hypothesis HPE1 : compatible_as PE1.
+Hypothesis HPE2 : compatible_as PE2.
+
+Context {f : E1 -> E2}.
+Context {fS : sub_struct HPE1 -> sub_struct HPE2}.
+Hypothesis HfS : forall x, val (fS x) = f (val x).
+
+Lemma sub_as_fun_rev : funS PE1 PE2 f.
+Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
+
+Lemma sub_as_inj : injS PE1 f -> injective fS.
+Proof. apply sub_inj, HfS. Qed.
+
+Lemma sub_as_inj_rev : injective fS -> injS PE1 f.
+Proof. apply sub_inj_rev, HfS. Qed.
+
+Lemma sub_as_inj_equiv : injective fS <-> injS PE1 f.
+Proof. apply sub_inj_equiv, HfS. Qed.
+
+Lemma sub_as_surj : surjS PE1 PE2 f -> surjective fS.
+Proof. apply sub_surj, HfS. Qed.
+
+Lemma sub_as_surj_rev : surjective fS -> surjS PE1 PE2 f.
+Proof. apply sub_surj_rev, HfS. Qed.
+
+Lemma sub_as_surj_equiv : surjective fS <-> surjS PE1 PE2 f.
+Proof. apply sub_surj_equiv, HfS. Qed.
+
+Lemma sub_as_bij : bijS PE1 PE2 f -> bijective fS.
+Proof. apply sub_bij, HfS; apply inhabited_as. Qed.
+
+Lemma sub_as_bij_rev : bijective fS -> bijS PE1 PE2 f.
+Proof. apply sub_bij_rev, HfS; apply inhabited_as. Qed.
+
+Lemma sub_as_bij_equiv : bijective fS <-> bijS PE1 PE2 f.
+Proof. apply sub_bij_equiv, HfS; apply inhabited_as. Qed.
+
+Variable O1 : E1.
+Hypothesis HO1 : PE1 O1.
+Let PV1 := vectP PE1 O1.
+Let HPV1 : compatible_ms PV1 := compatible_as_ms HO1 HK HPE1.
+
+Let O2 := f O1.
+Hypothesis HO2 : PE2 O2.
+Let PV2 := vectP PE2 O2.
+Let HPV2 : compatible_ms PV2 := compatible_as_ms HO2 HK HPE2.
+
+Lemma sub_as_affine_mapping :
+  is_affine_mapping f ->
+  (@is_affine_mapping _ (sub_ModuleSpace HPV1) (sub_ModuleSpace HPV2) _ _ fS).
+Proof.
+intros Hf n; intros; apply val_inj; rewrite HfS !val_barycenter// Hf//.
+f_equal; apply extF; intro; rewrite !mapF_correct; easy.
+Qed.
+
+End Sub_AffineSpace_Affine_Mapping_Facts1.
+
+
+Section Sub_AffineSpace_Affine_Mapping_Facts2.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V1 V2 : ModuleSpace K}.
+Context {E1 : AffineSpace V1}.
+Context {E2 : AffineSpace V2}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+Hypothesis HPE1 : compatible_as PE1.
+Hypothesis HPE2 : compatible_as PE2.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : funS PE1 PE2 f.
+
+Definition fct_sub_as : sub_struct HPE1 -> sub_struct HPE2 :=
+  fct_sub HPE1 HPE2 Hf.
+
+Lemma fct_sub_as_inj : injS PE1 f -> injective fct_sub_as.
+Proof. apply fct_sub_inj. Qed.
+
+Lemma fct_sub_as_inj_rev : injective fct_sub_as -> injS PE1 f.
+Proof. apply fct_sub_inj_rev. Qed.
+
+Lemma fct_sub_as_inj_equiv : injective fct_sub_as <-> injS PE1 f.
+Proof. apply fct_sub_inj_equiv. Qed.
+
+Lemma fct_sub_as_surj : surjS PE1 PE2 f -> surjective fct_sub_as.
+Proof. apply fct_sub_surj. Qed.
+
+Lemma fct_sub_as_surj_rev : surjective fct_sub_as -> surjS PE1 PE2 f.
+Proof. apply fct_sub_surj_rev. Qed.
+
+Lemma fct_sub_as_surj_equiv : surjective fct_sub_as <-> surjS PE1 PE2 f.
+Proof. apply fct_sub_surj_equiv. Qed.
+
+Lemma fct_sub_as_bij : bijS PE1 PE2 f -> bijective fct_sub_as.
+Proof. apply fct_sub_bij, inhabited_as. Qed.
+
+Lemma fct_sub_as_bij_rev : bijective fct_sub_as -> bijS PE1 PE2 f.
+Proof. apply fct_sub_bij_rev, inhabited_as. Qed.
+
+Lemma fct_sub_as_bij_equiv : bijective fct_sub_as <-> bijS PE1 PE2 f.
+Proof. apply fct_sub_bij_equiv, inhabited_as. Qed.
+
+Variable O1 : E1.
+Hypothesis HO1 : PE1 O1.
+Let PV1 := vectP PE1 O1.
+Let HPV1 : compatible_ms PV1 := compatible_as_ms HO1 HK HPE1.
+
+Let O2 := f O1.
+Hypothesis HO2 : PE2 O2.
+Let PV2 := vectP PE2 O2.
+Let HPV2 : compatible_ms PV2 := compatible_as_ms HO2 HK HPE2.
+
+Lemma fct_sub_as_affine_mapping :
+  is_affine_mapping f ->
+  (@is_affine_mapping _
+    (sub_ModuleSpace HPV1) (sub_ModuleSpace HPV2) _ _ fct_sub_as).
+Proof. apply sub_as_affine_mapping, fct_sub_correct. Qed.
+
+End Sub_AffineSpace_Affine_Mapping_Facts2.
+
+
+Section Sub_AffineSpace_Affine_Mapping_Facts3.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V W : ModuleSpace K}.
+Context {E : AffineSpace V}.
+Context {F : AffineSpace W}.
+
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_as PE.
+
+Context {PF : F -> Prop}.
+Hypothesis HPF : compatible_as PF.
+
+Context {f : E -> F}.
+Hypothesis Hf : is_affine_mapping f.
+
+Variable O : E.
+Hypothesis HO : PE O.
+
+Let lf := fct_lm f O.
+Let PV := vectP PE O.
+Let PW := vectP PF (f O).
+
+Lemma funS_aff_lin_equiv : funS PE PF f <-> funS PV PW lf.
+Proof.
+rewrite !funS_equiv; split.
+intros Hf1 v Hv; unfold PW, vectP, preimage, lf; rewrite transl_correct_l; auto.
+intros Hlf A HA; rewrite -(transl_correct_l (f O) (f A)) -(transl_correct_l O A).
+apply Hlf; unfold PV, vectP, preimage; rewrite transl_correct_l; easy.
+Qed.
+
+Lemma injS_aff_lin_equiv : injS PE f <-> injS PV lf.
+Proof.
+split.
+(* *)
+intros Hf1 u1 u2 Hu1 Hu2 Hu; apply (transl_l_inj O), Hf1; auto.
+apply (vect_l_inj (f O)); auto.
+(* *)
+intros Hlf1 A1 A2 HA1 HA2 HA; apply (vect_l_inj O), Hlf1.
+1,2: unfold PV, vectP, preimage; rewrite transl_correct_l; easy.
+unfold lf, fct_lm; rewrite !transl_correct_l HA; easy.
+Qed.
+
+Lemma injS_aff_lin_equiv_alt : injS PE f <-> incl (KerS PV lf) zero_sub_struct.
+Proof.
+rewrite -> injS_aff_lin_equiv, <- KerS_zero_equiv, <- lmS_injS_equiv; [easy |..].
+1,3: apply compatible_as_ms; easy.
+1,2: apply is_affine_mapping_equiv; easy.
+Qed.
+
+Lemma surjS_gen_aff_lin_equiv : surjS PE PF f <-> surjS PV PW lf.
+Proof.
+split.
+(* *)
+intros Hf1 w Hw; destruct (Hf1 (f O +++ w)) as [A [HA1 HA2]]; try easy.
+exists (O --> A); split.
+unfold PV, vectP, preimage; rewrite transl_correct_l; easy.
+unfold lf, fct_lm; rewrite transl_correct_l HA2 transl_correct_r; easy.
+(* *)
+intros Hlf B HB; destruct (Hlf (f O --> B)) as [v [Hv1 Hv2]].
+unfold PW, vectP, preimage; rewrite transl_correct_l; easy.
+exists (O +++ v); split; try apply (vect_l_inj (f O)); easy.
+Qed.
+
+Lemma surjS_gen_aff_lin_equiv_alt : surjS PE PF f <-> incl PW (RgS PV lf).
+Proof.
+rewrite surjS_gen_aff_lin_equiv -RgS_gen_full_equiv; apply surjS_RgS_gen_equiv.
+Qed.
+
+Lemma bijS_gen_aff_lin_equiv : bijS PE PF f <-> bijS PV PW lf.
+Proof.
+rewrite !bijS_equiv; [| apply inhabited_ms | easy].
+rewrite funS_aff_lin_equiv injS_aff_lin_equiv surjS_gen_aff_lin_equiv; easy.
+Qed.
+
+Lemma bijS_gen_aff_lin_equiv_alt :
+  bijS PE PF f <->
+  funS PV PW lf /\ incl (KerS PV lf) zero_sub_struct /\ incl PW (RgS PV lf).
+Proof.
+rewrite bijS_equiv; [| apply inhabited_as].
+rewrite funS_aff_lin_equiv injS_aff_lin_equiv_alt surjS_gen_aff_lin_equiv_alt//.
+Qed.
+
+End Sub_AffineSpace_Affine_Mapping_Facts3.
+
+
+Section Sub_AffineSpace_Affine_Mapping_Facts4.
+
+Context {K : Ring}.
+Hypothesis HK : invertible (2 : K).
+Context {V W : ModuleSpace K}.
+Context {E : AffineSpace V}.
+Context {F : AffineSpace W}.
+
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_as PE.
+
+Context {f : E -> F}.
+Hypothesis Hf : is_affine_mapping f.
+
+Context {O : E}.
+Hypothesis HO : PE O.
+
+Let lf := fct_lm f O.
+Let PV := vectP PE O.
+
+Lemma surjS_aff_lin_equiv : surjS PE fullset f <-> surjS PV fullset lf.
+Proof. rewrite (surjS_gen_aff_lin_equiv O) vectP_fullset; easy. Qed.
+
+Lemma surjS_aff_lin_equiv_alt :
+  surjS PE fullset f <-> incl fullset (RgS PV lf).
+Proof. rewrite (surjS_gen_aff_lin_equiv_alt O) vectP_fullset; easy. Qed.
+
+Lemma bijS_aff_lin_equiv : bijS PE fullset f <-> bijS PV fullset lf.
+Proof. rewrite (bijS_gen_aff_lin_equiv O) vectP_fullset; easy. Qed.
+
+Lemma bijS_aff_lin_equiv_alt :
+  bijS PE fullset f <->
+  incl (KerS PV lf) zero_sub_struct /\ incl fullset (RgS PV lf).
+Proof.
+rewrite (bijS_gen_aff_lin_equiv_alt HK HPE Hf O)// vectP_fullset; easy.
+Qed.
+
+End Sub_AffineSpace_Affine_Mapping_Facts4.
+
+
+Section Compatible_AffineSpace_R_Facts.
+
+Context {V : ModuleSpace R_Ring}.
+Context {E : AffineSpace V}.
+
+Lemma compatible_as_ms_equiv_R :
+  forall {PE : E -> Prop} {O},
+    PE O -> compatible_as PE <-> compatible_ms (vectP PE O).
+Proof.
+move=>> HO; apply compatible_as_ms_equiv; try easy.
+apply nonzero_struct_R. apply invertible_2.
+Qed.
+
+Lemma compatible_as_barycenter_closure_R :
+  forall (gen : E -> Prop), compatible_as (barycenter_closure gen).
+Proof.
+intros gen n L A HL HA; destruct (barycenter_normalized_n0_ex _ A HL)
+    as [n1 [L1 [A1 [HL1a [HL1b [HA1a HA1b]]]]]]; rewrite HA1b.
+destruct (barycenter_closure_exs _ (inclF_monot_l _ _ _ HA1a HA))
+    as [b [M [B [HM1 [HM2 [HB HA']]]]]].
+assert (HL1' : L1 = fun i1 => sum (scal (L1 i1) (M i1))).
+  apply extF; intro; rewrite -scal_sum_distr_l HM1 scal_one_r; easy.
+rewrite (barycenter_eq_r (fun i1 => barycenter (scal (L1 i1) (M i1)) (B i1))).
+2: { rewrite HA'; apply extF; intro; apply barycenter_homogeneous.
+  apply invertible_equiv_R; easy.
+  rewrite HM1; apply invertible_one. }
+rewrite {1}HL1' -barycenter_assoc.
+apply Barycenter_closure.
+1,4: rewrite sum_assoc -HL1' HL1a; apply invertible_one.
+apply concatnF_inclF_equiv; easy.
+intros i; rewrite -(extF_rev _ _ HL1' i).
+apply invertible_equiv_R; easy.
+Qed.
+
+Lemma barycenter_closure_idem_R :
+  forall (gen : E -> Prop),
+    barycenter_closure (barycenter_closure gen) = barycenter_closure gen.
+Proof.
+intros; apply eq_sym, compatible_as_equiv, compatible_as_barycenter_closure_R.
+Qed.
+
+(* Gostiaux T4, Th 1.26 p. 10. *)
+Lemma span_as_eq_R :
+  forall {gen : E -> Prop}, span_as gen = barycenter_closure gen.
+Proof.
+intros gen; apply subset_ext_equiv; split.
+(* *)
+apply (span_as_lub gen); try easy.
+apply compatible_as_barycenter_closure_R.
+apply barycenter_closure_incl.
+(* *)
+intros A HA; induction HA; apply span_as_compatible_as; try easy.
+apply inclF_monot_r with gen; try easy.
+apply span_as_incl.
+Qed.
+
+End Compatible_AffineSpace_R_Facts.
+
+
+Section Sub_AffineSpace_Affine_Mapping_R_Facts.
+
+Context {V W : ModuleSpace R_Ring}.
+Context {E : AffineSpace V}.
+Context {F : AffineSpace W}.
+
+(* Proof path is inspired from that of barycenter_comm_R. *)
+Lemma compatible_as_affine_mapping :
+  compatible_as (@is_affine_mapping _ _ _ E F).
+Proof.
+intros n2 L2 f2 HL2 Hf2 n1 L1 A1 HL1; generalize HL1 HL2.
+move=> /invertible_equiv_R HL1' /invertible_equiv_R HL2'.
+rewrite fct_barycenter_eq; try easy.
+assert (Hf2a :
+  f2^~ (barycenter L1 A1) = fun i2 => barycenter L1 (mapF (f2 i2) A1))
+    by now apply extF; intros i2; apply (Hf2 i2).
+rewrite Hf2a.
+pose (f := barycenter L2 f2).
+generalize (barycenter_correct f2 HL2); fold f; intros Hf.
+pose (B1 := mapF f A1); fold B1.
+pose (B2 := fun i2 => barycenter L1 (mapF (f2 i2) A1)); fold B2.
+assert (HB2 : forall i2, is_comb_aff L1 (mapF (f2 i2) A1) (B2 i2))
+    by now intros; apply barycenter_correct_equiv.
+pose (B := barycenter L2 B2).
+generalize (barycenter_correct B2 HL2); fold B; intros HB.
+apply barycenter_correct_equiv; try easy; unfold is_comb_aff in *.
+pose (M i1 i2 := f2 i2 (A1 i1) --> B2 i2).
+assert (HM : comb_lin L1 (mapF (comb_lin L2 (vectF f f2)) A1) +
+    comb_lin L1 (mapF (comb_lin L2) M) = 0).
+  rewrite Hf mapF_zero_f comb_lin_zero_r plus_zero_l.
+  rewrite -comb_lin_flipT_r comb_linT_sym_R flipT_invol; unfold M.
+  apply comb_lin_zero_compat_r, extF; intros i2.
+  rewrite fct_comb_lin_r_eq fct_zero_eq opp_zero_equiv -comb_lin_opp_r -(HB2 i2).
+  apply comb_lin_eq_r, extF; intro; rewrite vectF_correct vect_sym; easy.
+apply (scal_zero_reg_r_R (sum L2)); try easy.
+rewrite -(plus_zero_r (scal _ _)) -{1}HM -comb_lin_plus_r; unfold M.
+rewrite (comb_lin_eq_r _ (_ + _) (comb_lin L2 (fun i2 i1 => B1 i1 --> B2 i2))).
+2: { apply extF; intro.
+    rewrite fct_plus_eq !mapF_correct !fct_comb_lin_r_eq -comb_lin_plus_r; f_equal.
+    rewrite vectF_chasles; easy. }
+rewrite -comb_lin_scal_r scal_sum_l -!comb_lin_plus_r.
+apply comb_lin_zero_compat_r, extF; intros i1.
+rewrite fct_comb_lin_r_eq fct_zero_eq -HB; f_equal; apply extF; intros i2.
+rewrite !fct_plus_eq constF_correct !vectF_correct; apply vect_chasles.
+Qed.
+
+(* Am is the affine subspace of affine mappings. *)
+Definition Am := sub_struct compatible_as_affine_mapping.
+
+End Sub_AffineSpace_Affine_Mapping_R_Facts.
diff --git a/FEM/Algebra/Finite_dim_new.v b/FEM/Algebra/Finite_dim_new.v
new file mode 100644
index 0000000000000000000000000000000000000000..a94d40e9f1f4fa14c4d72db970e81fe81f82e746
--- /dev/null
+++ b/FEM/Algebra/Finite_dim_new.v
@@ -0,0 +1,2051 @@
+(**
+This file is part of the Elfic library
+
+Copyright (C) Boldo, Clément, Martin, Mayero, Mouhcine
+
+This library is free software; you can redistribute it and/or
+modify it under the terms of the GNU Lesser General Public
+License as published by the Free Software Foundation; either
+version 3 of the License, or (at your option) any later version.
+
+This library is distributed in the hope that it will be useful,
+but WITHOUT ANY WARRANTY; without even the implied warranty of
+MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+COPYING file for more details.
+*)
+
+From FEM.Compl Require Import stdlib.
+
+Set Warnings "-notation-overridden".
+From FEM.Compl Require Import ssr.
+
+(* Next import provides Coquelicot.Hierarchy, functions to module spaces, and
+ linear mappings. *)
+From LM Require Import linear_map.
+Close Scope R_scope.
+Set Warnings "notation-overridden".
+
+From Lebesgue Require Import Subset Function.
+
+From FEM Require Import Compl Sub_struct Monoid_compl Group_compl Ring_compl.
+From FEM Require Import ModuleSpace_compl AffineSpace.
+
+
+(** Results needing a commutative Ring are only stated for
+ the ring of real numbers R_Ring in file 'Finite_dim_R.v'. *)
+
+
+Local Open Scope Monoid_scope.
+Local Open Scope Group_scope.
+Local Open Scope Ring_scope.
+Local Open Scope AffineSpace_scope.
+
+
+Section FD_Def.
+
+(** Definitions.
+
+  Let E be a real vector space.
+  Let PE be a subset of E.
+  Let B be an n-family of vectors in E.
+  Let x, y, z be vectors in E.
+
+  "span B" is the subset of the linear combinations of vectors of B.
+  Lemma span_EX provides a strong existential for spans.
+  Lemma span_ex provides some smoothness to build the linear combination.
+
+  "in_line x" is the subset of E generated by x.
+  It is a vector line iff x <> 0.
+
+  "in_plane x y" is the subset of E generated by x and y.
+  It is a vector plane iff x <> 0 /\ y <> scal lx x.
+
+  "in_3d_space x y z" is the subset of E generated by x, y and z.
+  It is a vector 3d-space iff x <> 0 /\ y <> scal lx x /\ z <> scal lx x + scal ly y .
+
+  "comb_lin_surjL PE B" means that each element of PE is a linear combination of B.
+  "comb_lin_injL B" means that linear combinations of B are unique.
+  "comb_lin_bijL PE B" means that each element of PE is a unique linear combination of B.
+
+  "is_generator PE B" means that the vectors of B span PE (i.e. PE = span B).
+
+  "is_free B" means that the trivial linear combination of B is the only decomposition of 0,
+    i.e. that the vectors of B are linearly independent.
+
+  "is_basis PE B" means that B is both generator of PE and free, a.k.a. a basis of PE.
+
+  "is_finite_dim PE" means the PE admits a finite generator.
+
+  "has_dim PE n" means that PE admits a basis of length n.
+
+  "dim PE" returns n when PE admits a basis of length n, and 0 otherwise.
+*)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (07/06/2023): try to use new span_struct. *)
+
+(* FIXME: explore the use of a strong existential. *)
+Definition span {n} (B : 'E^n) (x : E) : Type := { L | x = comb_lin L B }.
+
+Definition in_line (x0 : E) := span (singleF x0).
+Definition in_plane (x0 x1 : E) := span (coupleF x0 x1).
+Definition in_3d_space (x0 x1 x2 : E) := span (tripleF x0 x1 x2).
+
+Definition comb_lin_surjL (PE : E -> Prop) {n} (B : 'E^n) :=
+  forall x, PE x -> exists L, x = comb_lin L B.
+Definition comb_lin_injL {n} (B : 'E^n) := injective (comb_lin^~ B).
+(* ie, forall L L', comb_lin L B = comb_lin L' B -> L = L'.*)
+Definition comb_lin_bijL (PE : E -> Prop) {n} (B : 'E^n) :=
+  forall x, PE x -> exists! L, x = comb_lin L B.
+
+(* FIXME: ouin ! span B n'est plus un subset... *)
+Definition is_generator (PE : E -> Prop) {n} (B : 'E^n) := PE = span B.
+
+Lemma is_generator_ext :
+  forall {PE : E -> Prop} {n} {B1 B2 : 'E^n},
+    B1 = B2 -> is_generator PE B1 -> is_generator PE B2.
+Proof. intros; subst; easy. Qed.
+
+(* FIXME: pas sûr que ce soit plus lisible, si jamais ça typait... *)
+Definition is_finite_dim (PE : E -> Prop) :=
+  { n & { B : 'E^n | is_generator PE B } }.
+(*  exists n (B : 'E^n), is_generator PE B.*)
+
+Definition is_free {n} (B : 'E^n) :=
+  forall (L : 'K^n), comb_lin L B = 0 -> L = 0.
+
+Lemma is_free_ext :
+  forall {n} (B1 B2 : 'E^n), B1 = B2 -> is_free B1 -> is_free B2.
+Proof. intros; subst; easy. Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC: this could be a notation... *)
+Definition is_lin_dep {n} (B : 'E^n) := ~ is_free B.
+
+Lemma is_lin_dep_ext :
+  forall {n} (B1 B2 : 'E^n), B1 = B2 -> is_lin_dep B1 -> is_lin_dep B2.
+Proof. intros; subst; easy. Qed.
+
+Definition is_basis (PE : E -> Prop) {n} (B : 'E^n) :=
+  is_generator PE B /\ is_free B.
+
+Lemma is_basis_ext :
+  forall (PE : E -> Prop) {n} (B1 B2 : 'E^n),
+    B1 = B2 -> is_basis PE B1 -> is_basis PE B2.
+Proof. intros; subst; easy. Qed.
+
+Lemma is_basis_ext' :
+  forall (PE1 PE2 : E -> Prop) {n} (B1 B2 : 'E^n),
+    PE1 = PE2 -> B1 = B2 -> is_basis PE1 B1 -> is_basis PE2 B2.
+Proof. intros; subst; easy. Qed.
+
+Inductive has_dim (PE : E -> Prop) n : Prop :=
+  Dim : forall (B : 'E^n), is_basis PE B -> has_dim PE n.
+
+Lemma has_dim_EX :
+  forall (PE : E -> Prop) n, has_dim PE n -> { B : 'E^n | is_basis PE B }.
+Proof. move=>> H; apply ex_EX; induction H as [B]; exists B; easy. Qed.
+
+Lemma has_dim_ex:
+  forall (PE : E -> Prop) n,
+    (exists B : 'E^n, is_basis PE B) -> has_dim PE n.
+Proof. move=>> [B HB]. apply (Dim _ _ B HB). Qed.
+
+Lemma has_dim_ext :
+  forall (PE1 PE2 : E -> Prop) n1 n2,
+    PE1 = PE2 -> n1 = n2 -> has_dim PE1 n1 -> has_dim PE2 n2.
+Proof. intros; subst; easy. Qed.
+
+End FD_Def.
+
+
+Section Span_Subset_Facts.
+
+(** Subset properties of span. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Lemma span_inclF_diag : forall {n} (B : 'E^n), inclF B (span B).
+Proof. intros n B i; rewrite -comb_lin_kron_in_r; easy. Qed.
+
+Lemma span_ub : forall {n} (B : 'E^n) A i, A = B i -> span B A.
+Proof. intros; subst; apply span_inclF_diag. Qed.
+
+Lemma span_nil : forall (B : 'E^0), span B = zero_sub_struct.
+Proof.
+intros B; apply subset_ext_equiv; split.
+intros _ [L]; apply comb_lin_nil.
+intros x Hx; rewrite Hx -(comb_lin_zero_compat_l 0 B); easy.
+Qed.
+
+Lemma span_zero : forall {n}, span (0 : 'E^n) = zero_sub_struct.
+Proof.
+intros; apply subset_ext_equiv; split.
+intros _ [L]; apply comb_lin_zero_r.
+intros x Hx; rewrite Hx -(comb_lin_zero_compat_r (0 : 'K^n) 0); easy.
+Qed.
+
+Lemma span_zero_alt : forall {n} {B : 'E^n}, B = 0 -> span B = zero_sub_struct.
+Proof. move=>> ->; apply span_zero. Qed.
+
+Lemma span_singleF : forall (B : 'E^1), span B = in_line (B ord0).
+Proof. intros B; rewrite -> (singleF_correct B) at 1; easy. Qed.
+
+Lemma span_coupleF :
+  forall (B : 'E^2), span B = in_plane (B ord0) (B ord_max).
+Proof. intros B; rewrite -> (coupleF_correct B) at 1; easy. Qed.
+
+Lemma span_tripleF :
+  forall (B : 'E^3), span B = in_3d_space (B ord0) (B ord_1) (B ord_max).
+Proof. intros B; rewrite -> (tripleF_correct B) at 1; easy. Qed.
+
+Lemma span_inclF :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2), invalF B1 B2 -> inclF B1 (span B2).
+Proof.
+move=>> HB i; destruct (HB i) as [i2 Hi2]; rewrite Hi2; apply span_inclF_diag.
+Qed.
+
+Lemma spanF_ex :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) <-> exists M12, forall i1, B1 i1 = comb_lin (M12 i1) B2.
+Proof.
+intros n1 n2 B1 B2; split; intros HB.
+(* *)
+assert (HB' : forall i1, exists L2, B1 i1 = comb_lin L2 B2).
+  intros i1; destruct (HB i1) as [L2]; exists L2; easy.
+apply (choice _ HB').
+(* *)
+destruct HB as [H HM]; intros i1; rewrite HM; easy.
+Qed.
+
+Lemma spanF_ex_flipT :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) <-> exists M12, forall i1, B1 i1 = comb_lin (M12^~ i1) B2.
+Proof.
+intros; rewrite spanF_ex; split; intros [L HL];
+    exists (flipT L); intros j; rewrite HL; easy.
+Qed.
+
+Lemma span_incl :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) -> incl (span B1) (span B2).
+Proof.
+move=>> HB1 _ [L1]; apply spanF_ex in HB1; destruct HB1 as [M12 HM12].
+apply span_ex; eexists; apply comb_lin2_alt; easy.
+Qed.
+
+Lemma span_monot :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    invalF B1 B2 -> incl (span B1) (span B2).
+Proof. intros; apply span_incl, span_inclF; easy. Qed.
+
+Lemma span_ext :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) -> inclF B2 (span B1) -> span B1 = span B2.
+Proof. intros; apply subset_ext_equiv; split; apply span_incl; easy. Qed.
+
+Lemma span_concatF_sym_inclF :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF (concatF B1 B2) (span (concatF B2 B1)).
+Proof.
+intros n1 n2 B1 B2 i; apply span_ex.
+destruct (lt_dec i n1) as [Hi | Hi]; eexists.
+(* *)
+rewrite concatF_correct_l -comb_lin_kron_in_r.
+rewrite -(comb_lin_concatF_zero_lr _ _ B2) comb_lin_concatF.
+rewrite plus_comm; symmetry; apply comb_lin_concatF.
+(* *)
+rewrite concatF_correct_r -comb_lin_kron_in_l.
+rewrite -(comb_lin_concatF_zero_ll _ B1) comb_lin_concatF.
+rewrite plus_comm; symmetry; apply comb_lin_concatF.
+Qed.
+
+Lemma span_coupleF_sym_inclF :
+  forall (x0 x1 : E), inclF  (coupleF x1 x0) (in_plane x0 x1).
+Proof. intros; apply span_inclF, invalF_coupleF_sym. Qed.
+
+Lemma span_coupleF_sym : forall (x0 x1 : E), in_plane x0 x1 = in_plane x1 x0.
+Proof. intros; apply span_ext; apply span_coupleF_sym_inclF. Qed.
+
+Lemma span_castF_compat :
+  forall {n1 n2} (H : n1 = n2) (B1 : 'E^n1), span (castF H B1) = span B1.
+Proof.
+intros n1 n2 Hn B1; apply subset_ext_equiv; split; intros _ [L]; apply span_ex.
+subst n1; exists L; rewrite castF_id; easy.
+subst n2; exists L; rewrite castF_id; easy.
+Qed.
+
+Lemma span_concatF_sym :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    span (concatF B1 B2) = span (concatF B2 B1).
+Proof. intros; apply span_ext; apply span_concatF_sym_inclF. Qed.
+
+Lemma span_concatF_l :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B2 (span B1) -> span (concatF B1 B2) = span B1.
+Proof.
+intros n1 n2 B1 B2 HB; apply span_ext.
+apply concatF_lub_inclF; try apply span_inclF_diag; easy.
+eapply (inclF_monot_r (span _));
+    [apply span_monot, concatF_ub_l | apply span_inclF_diag].
+Qed.
+
+Lemma span_concatF_r :
+  forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) -> span (concatF B1 B2) = span B2.
+Proof.
+intros; apply span_ext.
+apply concatF_lub_inclF; try apply span_inclF_diag; easy.
+eapply (inclF_monot_r (span _));
+    [apply span_monot, concatF_ub_r | apply span_inclF_diag].
+Qed.
+
+Lemma span_span : forall {n1 n2} (B1 : 'E^n1) (B2 : 'E^n2) (L : 'K^n2),
+   (forall i, span B1 (B2 i)) -> span B1 (comb_lin L B2).
+Proof.
+intros n1 n2 B1 B2 L Hi.
+generalize (span_incl B2 B1); unfold incl; intros H.
+apply H; easy.
+Qed.
+
+End Span_Subset_Facts.
+
+
+Section Span_Linear_Facts.
+
+(** Linear properties of span. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {n : nat}.
+
+Variable B : 'E^n.
+
+Lemma span_zero_closed : span B 0.
+Proof. rewrite <- (comb_lin_zero_compat_l 0 B); easy. Qed.
+
+Lemma span_nonempty : nonempty (span B).
+Proof. exists 0; apply span_zero_closed. Qed.
+
+Lemma span_scal_closed : scal_closed (span B).
+Proof. intros l _ [Lx]; rewrite -comb_lin_scal_l; easy. Qed.
+
+Lemma span_opp_closed : opp_closed (span B).
+Proof. apply scal_opp_closed, span_scal_closed. Qed.
+
+Lemma span_plus_closed : plus_closed (span B).
+Proof. intros _ _ [Lx] [Ly]; rewrite -comb_lin_plus_l; easy. Qed.
+
+Lemma span_minus_closed : minus_closed (span B).
+Proof.
+apply plus_opp_minus_closed; [apply span_plus_closed | apply span_opp_closed].
+Qed.
+
+(* span B is a vector subspace. *)
+Lemma span_compatible_ms : compatible_ms (span B).
+Proof.
+apply plus_scal_closed_compatible_ms; [exists 0; apply span_zero_closed |
+    apply span_plus_closed | apply span_scal_closed].
+Qed.
+
+Lemma span_comb_lin_closed : comb_lin_closed (span B).
+Proof. intros m; apply compatible_ms_comb_lin, span_compatible_ms. Qed.
+
+(* Gostiaux T1, Th 6.19 pp. 168-169. *)
+Lemma span_equiv : span B = span_ms (inF^~ B).
+Proof.
+apply subset_ext_equiv; split.
+(* *)
+intros x [L]; apply span_ms_comb_lin.
+intros i; apply span_ms_incl; exists i; easy.
+(* *)
+apply span_ms_lub.
+apply span_compatible_ms.
+intros x [i Hi]; rewrite Hi; apply span_inclF_diag.
+Qed.
+
+Lemma span_ex_decomp : comb_lin_surjL (span B) B.
+Proof. intros _ [L]; exists L; easy. Qed.
+
+Lemma span_lb :
+  forall {PE}, compatible_ms PE -> inclF B PE -> incl (span B) PE.
+Proof. intros PE HPE HB x [L]; apply compatible_ms_comb_lin; easy. Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC: this could use new interp_any from Subset branch. *)
+Lemma span_glb :
+  span B = (fun x => forall PE, compatible_ms PE -> inclF B PE -> PE x).
+Proof.
+apply subset_ext_equiv; split.
+intros _ [L] PE HPE1 HPE2; apply compatible_ms_comb_lin; easy.
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC: could use span_lb with properties of interp_any on compatible_ms and inclF. *)
+intros x Hx; apply Hx; [apply span_compatible_ms | apply span_inclF_diag].
+Qed.
+
+Lemma span_filter_n0F : forall {n} {B : 'E^n}, span (filter_n0F B) = span B.
+Proof.
+clear n B; intros n B; apply subset_ext_equiv; split; intros _ [L].
+destruct (comb_lin_filter_n0F_r_ex B L) as [L' ->]; easy.
+rewrite -comb_lin_filter_n0F_r; easy.
+Qed.
+
+End Span_Linear_Facts.
+
+
+Section Comb_lin_L_Facts.
+
+(** Properties of comb_lin_surjL, comb_lin_injL and comb_lin_bijL. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {n : nat}.
+
+Variable PE : E -> Prop.
+Variable B : 'E^n.
+
+Lemma comb_lin_surjL_injL_bijL :
+  comb_lin_surjL PE B -> comb_lin_injL B -> comb_lin_bijL PE B.
+Proof.
+intros HB1 HB2 x Hx; destruct (HB1 _ Hx) as [L HL].
+exists L; split; try easy.
+intros L' HL'; apply HB2; rewrite <- HL; easy.
+Qed.
+
+Lemma comb_lin_bijL_surjL : comb_lin_bijL PE B -> comb_lin_surjL PE B.
+Proof. intros HB x Hx; destruct (HB _ Hx) as [L [HL _]]; exists L; easy. Qed.
+
+Lemma comb_lin_bijL_injL :
+  compatible_ms PE -> inclF B PE -> comb_lin_bijL PE B -> comb_lin_injL B.
+Proof.
+intros HPE HB1 HB2 L L' HL.
+destruct (HB2 (comb_lin L B)) as [Lx [_ HLx]];
+    try now apply compatible_ms_comb_lin.
+apply trans_eq with Lx; [symmetry |]; apply HLx; easy.
+Qed.
+
+End Comb_lin_L_Facts.
+
+
+Section Generator_Facts.
+
+(** Properties of is_generator (from those of span). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {n n1 n2 : nat}.
+Context {PE : E -> Prop}.
+
+Lemma is_generator_span : forall (B : 'E^n), is_generator (span B) B.
+Proof. easy. Qed.
+
+Lemma is_generator_inclF : forall {B : 'E^n}, is_generator PE B -> inclF B PE.
+Proof. intros B HB; rewrite HB; apply span_inclF_diag. Qed.
+
+Lemma is_generator_nonempty :
+  forall (B : 'E^n), is_generator PE B -> nonempty PE.
+Proof. move=>> ->; apply span_nonempty. Qed.
+
+Lemma is_generator_nil : forall (B : 'E^0), is_generator zero_sub_struct B.
+Proof. intros; unfold is_generator; rewrite span_nil; easy. Qed.
+
+Lemma is_generator_nonnil :
+  forall {B : 'E^n}, is_generator PE B -> PE <> zero_sub_struct -> B <> 0.
+Proof. move=>> HB HPE; contradict HPE; rewrite HB HPE span_zero; easy. Qed.
+
+Lemma is_generator_1 :
+  forall (x0 : E), is_generator (in_line x0) (singleF x0).
+Proof. easy. Qed.
+
+Lemma is_generator_2 :
+  forall (x0 x1 : E), is_generator (in_plane x0 x1) (coupleF x0 x1).
+Proof. easy. Qed.
+
+Lemma is_generator_3 :
+  forall (x0 x1 x2 : E),
+    is_generator (in_3d_space x0 x1 x2) (tripleF x0 x1 x2).
+Proof. easy. Qed.
+
+Lemma is_generator_coupleF_sym :
+  forall (x0 x1 : E),
+    is_generator PE (coupleF x0 x1) -> is_generator PE (coupleF x1 x0).
+Proof.
+intros x0 x1; unfold is_generator; fold (in_plane x0 x1).
+rewrite span_coupleF_sym; easy.
+Qed.
+
+Lemma is_generator_monot :
+  forall (B1 : 'E^n1) {B2 : 'E^n2},
+    invalF B1 B2 -> inclF B2 PE -> is_generator PE B1 -> is_generator PE B2.
+Proof.
+move=>> HB HB2 HB1; rewrite HB1; apply subset_ext_equiv; split.
+apply span_monot; easy. apply span_incl; rewrite -HB1; easy.
+Qed.
+
+Lemma is_generator_equiv :
+  forall (B1 : 'E^n1) (B2 : 'E^n2),
+    inclF B1 (span B2) -> inclF B2 (span B1) ->
+    is_generator PE B1 <-> is_generator PE B2.
+Proof. intros; unfold is_generator; rewrite (span_ext _ B2); easy. Qed.
+
+Lemma is_generator_castF_equiv :
+  forall (H : n1 = n2) {B1 : 'E^n1},
+    is_generator PE (castF H B1) <-> is_generator PE B1.
+Proof. intros; unfold is_generator; rewrite span_castF_compat; easy. Qed.
+
+Lemma is_generator_castF_compat :
+  forall (H : n1 = n2) {B1 : 'E^n1},
+    is_generator PE B1 -> is_generator PE (castF H B1).
+Proof. apply is_generator_castF_equiv. Qed.
+
+Lemma is_generator_castF_reg :
+  forall (H : n1 = n2) {B1 : 'E^n1},
+    is_generator PE (castF H B1) -> is_generator PE B1.
+Proof. apply is_generator_castF_equiv. Qed.
+
+Lemma is_generator_concatF_sym :
+  forall {B1 : 'E^n1} {B2 : 'E^n2},
+    is_generator PE (concatF B1 B2) -> is_generator PE (concatF B2 B1).
+Proof. intros B1 B2 HB; rewrite HB; apply span_concatF_sym. Qed.
+
+Lemma is_generator_concatF_compat_l :
+  forall {B1 : 'E^n1} {B2 : 'E^n2},
+    inclF B2 PE -> is_generator PE B1 -> is_generator PE (concatF B1 B2).
+Proof.
+move=>> HB2 HB1; rewrite HB1; apply eq_sym, span_concatF_l; rewrite -HB1; easy.
+Qed.
+
+Lemma is_generator_concatF_compat_r :
+  forall {B1 : 'E^n1} {B2 : 'E^n2},
+    inclF B1 PE -> is_generator PE B2 -> is_generator PE (concatF B1 B2).
+Proof.
+move=>> HB1 HB2; rewrite HB2; apply eq_sym, span_concatF_r; rewrite -HB2; easy.
+Qed.
+
+Lemma is_generator_zero_closed : forall (B : 'E^n), is_generator PE B -> PE 0.
+Proof. intros B HB; rewrite HB; apply span_zero_closed. Qed.
+
+Lemma is_generator_plus_closed :
+  forall (B : 'E^n), is_generator PE B -> plus_closed PE.
+Proof. intros B HB; rewrite HB; apply span_plus_closed. Qed.
+
+Lemma is_generator_opp_closed :
+  forall (B : 'E^n), is_generator PE B -> opp_closed PE.
+Proof. intros B HB; rewrite HB; apply span_opp_closed. Qed.
+
+Lemma is_generator_minus_closed :
+  forall (B : 'E^n), is_generator PE B -> minus_closed PE.
+Proof. intros B HB; rewrite HB; apply span_minus_closed. Qed.
+
+Lemma is_generator_scal_closed :
+  forall (B : 'E^n), is_generator PE B -> scal_closed PE.
+Proof. intros B HB; rewrite HB; apply span_scal_closed. Qed.
+
+Lemma is_generator_compatible_ms :
+   forall (B : 'E^n), is_generator PE B -> compatible_ms PE.
+Proof. intros B HB; rewrite HB; apply span_compatible_ms. Qed.
+
+Lemma is_generator_comb_lin_closed :
+  forall (B : 'E^n), is_generator PE B -> comb_lin_closed PE.
+Proof. intros B HB; rewrite HB; apply span_comb_lin_closed. Qed.
+
+Lemma is_generator_ex_decomp :
+  forall (B : 'E^n), is_generator PE B -> comb_lin_surjL PE B.
+Proof. intros B HB; rewrite HB; apply span_ex_decomp. Qed.
+
+Lemma is_generator_ex_decomp_rev :
+  forall {B : 'E^n},
+    compatible_ms PE -> inclF B PE -> comb_lin_surjL PE B -> is_generator PE B.
+Proof.
+intros B HB1 HB2 HB3; apply subset_ext_equiv; split.
+intros x Hx; destruct (HB3 _ Hx) as [L HL]; rewrite HL; easy.
+intros _ [L]; apply compatible_ms_comb_lin; easy.
+Qed.
+
+Lemma is_generator_uniq_decomp_rev :
+  forall {B : 'E^n},
+    compatible_ms PE -> inclF B PE -> comb_lin_bijL PE B -> is_generator PE B.
+Proof.
+intros; apply is_generator_ex_decomp_rev; try apply comb_lin_bijL_surjL; easy.
+Qed.
+
+End Generator_Facts.
+
+
+Section Generator_sub_Facts1a.
+
+(** Properties of is_generator on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_ms PE.
+Let PE_ms := sub_ms HPE.
+
+Context {n : nat}.
+Context {B : 'E^n}.
+Hypothesis HB : inclF B PE.
+Let B_ms : 'PE_ms^n := fun i => mk_sub_ms (HB i).
+
+Lemma is_generator_sub : is_generator PE B -> is_generator fullset B_ms.
+Proof.
+intros HB1; apply subset_ext_equiv; split; try easy.
+intros [x Hx] _; destruct (is_generator_ex_decomp _ HB1 x Hx) as [L HL].
+apply span_ex; exists L; apply val_inj; rewrite val_comb_lin; easy.
+Qed.
+
+Lemma is_generator_sub_rev : is_generator fullset B_ms -> is_generator PE B.
+Proof.
+intros HB1; apply subset_ext_equiv; split; intros x Hx.
+(* *)
+destruct (is_generator_ex_decomp _ HB1 (mk_sub_ms Hx)) as [L HL]; [easy |].
+apply span_ex; exists L; move: HL;
+    move=> /(f_equal val); rewrite val_comb_lin; easy.
+(* *)
+induction Hx as [L]; apply (compatible_ms_comb_lin HPE); easy.
+Qed.
+
+Lemma is_generator_sub_equiv : is_generator fullset B_ms <-> is_generator PE B.
+Proof.
+intros; split; [apply is_generator_sub_rev | apply is_generator_sub].
+Qed.
+
+End Generator_sub_Facts1a.
+
+
+Section Generator_sub_Facts1b.
+
+(** More properties of is_generator on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+
+Context {n : nat}.
+Context {B : 'E^n}.
+Hypothesis HB : is_generator PE B.
+Let HPE := is_generator_compatible_ms _ HB.
+Let PE_ms := sub_ms HPE.
+Let B_ms : 'PE_ms^n := fun i => mk_sub_ms (is_generator_inclF HB i).
+
+Lemma is_generator_sub_alt : is_generator fullset B_ms.
+Proof. apply is_generator_sub; easy. Qed.
+
+End Generator_sub_Facts1b.
+
+
+Section Generator_sub_Facts1c.
+
+(** Properties of is_generator on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_ms PE.
+Let PE_ms := sub_ms HPE.
+Context {n : nat}.
+Context {B_ms : 'PE_ms^n}.
+Let B := mapF val B_ms.
+
+Lemma is_generator_val : is_generator fullset B_ms -> is_generator PE B.
+Proof.
+rewrite -(is_generator_sub_equiv HPE (in_subF B_ms)) -(mk_subF_eq B_ms); easy.
+Qed.
+
+Lemma is_generator_val_rev : is_generator PE B -> is_generator fullset B_ms.
+Proof. rewrite (mk_subF_eq B_ms) (is_generator_sub_equiv HPE); easy. Qed.
+
+Lemma is_generator_val_equiv : is_generator PE B <-> is_generator fullset B_ms.
+Proof. rewrite (mk_subF_eq B_ms) (is_generator_sub_equiv HPE); easy. Qed.
+
+End Generator_sub_Facts1c.
+
+
+Section Generator_sub_Facts2.
+
+(** More properties of is_generator on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PEa PEb : E -> Prop}.
+Hypothesis HPE : incl PEa PEb.
+Hypothesis HPEb : compatible_ms PEb.
+Let PEb_ms := sub_ms HPEb.
+Let PEa' : PEb_ms -> Prop := preimage val PEa.
+
+Lemma image_preimage_val_ms : image val PEa' = inter PEa PEb.
+Proof. apply image_preimage_val. Qed.
+
+Lemma is_generator_preimage_val :
+  forall {n} {B : 'E^n} (HB : is_generator PEa B),
+    is_generator PEa' (fun i => mk_sub_ms (HPE _ (is_generator_inclF HB i))).
+Proof.
+intros n B HB; apply is_generator_ex_decomp_rev.
+apply preimage_val_compatible_ms, (is_generator_compatible_ms _ HB).
+apply (is_generator_inclF HB).
+intros x Hx; destruct (is_generator_ex_decomp _ HB (val x)) as [L HL]; [easy |].
+exists L; apply val_inj; rewrite val_comb_lin; easy.
+Qed.
+
+Lemma is_generator_preimage_val_rev :
+  forall {n} {B' : 'PEb_ms^n},
+    is_generator PEa' B' -> is_generator PEa (mapF val B').
+Proof.
+intros n B' HB'; apply is_generator_ex_decomp_rev.
+apply (preimage_val_compatible_ms_rev HPE HPEb),
+    (is_generator_compatible_ms _ HB').
+apply (is_generator_inclF HB').
+intros x Hx; destruct (is_generator_ex_decomp _ HB' (mk_sub_ms (HPE _ Hx)))
+    as [L HL]; [easy |].
+exists L; move: HL => /(f_equal val); rewrite val_comb_lin; easy.
+Qed.
+
+End Generator_sub_Facts2.
+
+
+Section Generator_sub_Facts3.
+
+(** More properties of is_generator on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PEb : E -> Prop}.
+Hypothesis HPEb : compatible_ms PEb.
+Let PEb_ms := sub_ms HPEb.
+Context {PEa' : PEb_ms -> Prop}.
+Let PEa := RgS PEa' val.
+
+Lemma preimage_image_val_ms : preimage val PEa = PEa'.
+Proof. apply preimage_image_val. Qed.
+
+Lemma incl_RgS_val : incl PEa PEb.
+Proof. intros _ [x Hx]; apply in_sub. Qed.
+
+Lemma is_generator_RgS_val :
+  forall {n} {B' : 'PEb_ms^n},
+    is_generator PEa' B' -> is_generator PEa (mapF val B').
+Proof.
+move=>>; rewrite -preimage_image_val_ms;
+    apply is_generator_preimage_val_rev, incl_RgS_val.
+Qed.
+
+Lemma is_generator_RgS_val_rev :
+  forall {n} {B : 'E^n} (HB : is_generator PEa B),
+    is_generator PEa'
+      (fun i => mk_sub_ms (incl_RgS_val _ (is_generator_inclF HB i))).
+Proof.
+intros; rewrite -preimage_image_val_ms; apply is_generator_preimage_val.
+Qed.
+
+End Generator_sub_Facts3.
+
+
+Section Is_finite_dim_Facts.
+
+(** Properties of is_finite_dim (from those of is_generator). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Context {PE : E -> Prop}.
+
+Lemma is_finite_dim_span : forall {n} (B : 'E^n), is_finite_dim (span B).
+Proof. intros n B; exists n, B; apply is_generator_span. Qed.
+
+Lemma is_finite_dim_nonempty : is_finite_dim PE -> nonempty PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_nonempty. Qed.
+
+Lemma is_finite_dim_0 : is_finite_dim (@zero_sub_struct E).
+Proof. exists 0, (fun _ => 0); apply is_generator_nil. Qed.
+
+Lemma is_finite_dim_1 : forall (x0 : E), is_finite_dim (in_line x0).
+Proof. intros x0; exists 1%nat, (singleF x0); apply is_generator_1. Qed.
+
+Lemma is_finite_dim_2 : forall (x0 x1 : E), is_finite_dim (in_plane x0 x1).
+Proof. intros x0 x1; exists 2%nat,  (coupleF x0 x1); apply is_generator_2. Qed.
+
+Lemma is_finite_dim_3 :
+  forall (x0 x1 x2 : E), is_finite_dim (in_3d_space x0 x1 x2).
+Proof.
+intros x0 x1 x2; exists 3%nat, (tripleF x0 x1 x2); apply is_generator_3.
+Qed.
+
+Lemma is_finite_dim_zero_closed : is_finite_dim PE -> zero_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_zero_closed. Qed.
+
+Lemma is_finite_dim_plus_closed : is_finite_dim PE -> plus_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_plus_closed. Qed.
+
+Lemma is_finite_dim_opp_closed : is_finite_dim PE -> opp_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_opp_closed. Qed.
+
+Lemma is_finite_dim_minus_closed : is_finite_dim PE -> minus_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_minus_closed. Qed.
+
+Lemma is_finite_dim_scal_closed : is_finite_dim PE -> scal_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_scal_closed. Qed.
+
+Lemma is_finite_dim_compatible_ms : is_finite_dim PE -> compatible_ms PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_compatible_ms. Qed.
+
+Lemma is_finite_dim_comb_lin_closed : is_finite_dim PE -> comb_lin_closed PE.
+Proof. intros [n [B HB]]; move: HB; apply is_generator_comb_lin_closed. Qed.
+
+Lemma is_finite_dim_ex_decomp_rev :
+  forall {n} (B : 'E^n),
+    compatible_ms PE -> inclF B PE -> comb_lin_surjL PE B -> is_finite_dim PE.
+Proof.
+intros n B; intros; exists n, B; apply is_generator_ex_decomp_rev; easy.
+Qed.
+
+Lemma is_finite_dim_uniq_decomp_rev :
+  forall {n} (B : 'E^n),
+    compatible_ms PE -> inclF B PE -> comb_lin_bijL PE B -> is_finite_dim PE.
+Proof.
+intros n B; intros; exists n, B; apply is_generator_uniq_decomp_rev; easy.
+Qed.
+
+End Is_finite_dim_Facts.
+
+
+Section Is_finite_dim_sub_Facts1a.
+
+(** Properties of is_finite_dim on sub-module spaces (from those of is_generator). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_ms PE.
+Let PE_ms := sub_ms HPE.
+
+Lemma is_finite_dim_sub : is_finite_dim PE -> is_finite_dim (@fullset PE_ms).
+Proof.
+intros [n [B HB1]]; move: (is_generator_inclF HB1) => HB2.
+exists n, (fun i => mk_sub_ms (HB2 i)).
+apply is_generator_sub; easy.
+Qed.
+
+Lemma is_finite_dim_sub_rev :
+  is_finite_dim (@fullset PE_ms) -> is_finite_dim PE.
+Proof.
+intros [n [B_ms HB1]]. move: (is_generator_inclF HB1) => HB2.
+exists n, (mapF val B_ms); apply is_generator_val; easy.
+Qed.
+
+Lemma is_finite_dim_sub_equiv :
+  is_finite_dim (@fullset PE_ms) <-> is_finite_dim PE.
+Proof.
+intros; split; [apply is_finite_dim_sub_rev | apply is_finite_dim_sub].
+Qed.
+
+End Is_finite_dim_sub_Facts1a.
+
+
+Section Is_finite_dim_sub_Facts1b.
+
+(** More properties of is_finite_dim on sub-module spaces (from those of is_generator). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : is_finite_dim PE.
+Let HPE1 := is_finite_dim_compatible_ms HPE.
+Let PE_ms := sub_ms HPE1.
+
+Lemma is_finite_dim_sub_alt : is_finite_dim (@fullset PE_ms).
+Proof. apply is_finite_dim_sub; easy. Qed.
+
+End Is_finite_dim_sub_Facts1b.
+
+
+Section Is_finite_dim_sub_Facts2.
+
+(** More properties of is_finite_dim on sub-module spaces (from those of is_generator). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PEa PEb : E -> Prop}.
+Hypothesis HPE : incl PEa PEb.
+Hypothesis HPEb : compatible_ms PEb.
+Let PEb_ms := sub_ms HPEb.
+Let PEa' : PEb_ms -> Prop := preimage val PEa.
+
+Lemma is_finite_dim_preimage_val : is_finite_dim PEa -> is_finite_dim PEa'.
+Proof.
+intros [n [B HB]];
+    exists n, (fun i => mk_sub_ms (HPE _ (is_generator_inclF HB i))).
+apply is_generator_preimage_val.
+Qed.
+
+Lemma is_finite_dim_preimage_val_rev :
+    is_finite_dim PEa' -> is_finite_dim PEa.
+Proof.
+intros [n [B' HB']]; exists n, (mapF val B');
+    apply (is_generator_preimage_val_rev HPE); easy.
+Qed.
+
+Lemma is_finite_dim_preimage_val_equiv :
+    is_finite_dim PEa' <-> is_finite_dim PEa.
+Proof.
+split; [apply is_finite_dim_preimage_val_rev |
+    apply is_finite_dim_preimage_val].
+Qed.
+
+End Is_finite_dim_sub_Facts2.
+
+
+Section Is_finite_dim_sub_Facts3.
+
+(** More properties of is_finite_dim on sub-module spaces (from those of is_generator). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PEb : E -> Prop}.
+Hypothesis HPEb : compatible_ms PEb.
+Let PEb_ms := sub_ms HPEb.
+Context {PEa' : PEb_ms -> Prop}.
+Let PEa := RgS PEa' val.
+
+Lemma is_finite_dim_RgS_val : is_finite_dim PEa' -> is_finite_dim PEa.
+Proof.
+intros [n [B' HB']]; exists n, (mapF val B'); apply is_generator_RgS_val; easy.
+Qed.
+
+Lemma is_finite_dim_RgS_val_rev : is_finite_dim PEa -> is_finite_dim PEa'.
+Proof.
+intros [n [B HB]]; exists n, (fun i => mk_sub_ms (incl_RgS_val _ _ (is_generator_inclF HB i))).
+apply is_generator_RgS_val_rev.
+Qed.
+
+Lemma is_finite_dim_RgS_val_equiv : is_finite_dim PEa <-> is_finite_dim PEa'.
+Proof.
+split; [apply is_finite_dim_RgS_val_rev | apply is_finite_dim_RgS_val].
+Qed.
+
+End Is_finite_dim_sub_Facts3.
+
+
+Section Free_Facts.
+
+(** Properties of is_lin_dep/is_free. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Hypothesis HK : nonzero_struct K.
+
+Lemma is_lin_dep_ex :
+  forall {n} {B : 'E^n},
+    is_lin_dep B <-> exists L, comb_lin L B = 0 /\ L <> 0.
+Proof.
+intros n B; split; intros HB.
+destruct (not_all_ex_not _ _ HB) as [L HL]; exists L; apply imply_to_and; easy.
+destruct HB as [L [HL1 HL2]]; apply ex_not_not_all; exists L; intros; auto.
+Qed.
+
+(** Establish first the property on is_lin_dep. *)
+
+Lemma is_lin_dep_with_zero : forall {n} {B : 'E^n}, inF 0 B -> is_lin_dep B.
+Proof.
+intros n B [i Hi]; apply is_lin_dep_ex; exists (kron _ i); split.
+rewrite comb_lin_kron_in_r; easy.
+apply nextF; exists i; rewrite kron_is_1; try easy.
+apply one_not_zero; easy.
+Qed.
+
+Lemma is_lin_dep_zero : forall {n}, is_lin_dep (0 : 'E^n.+1).
+Proof. intros; apply is_lin_dep_with_zero; exists ord0; easy.  Qed.
+
+Lemma is_lin_dep_zero_rev :
+  forall {n}, is_lin_dep (0 : 'E^n) -> (0 < n)%coq_nat.
+Proof.
+move=>> /is_lin_dep_ex [L [_ /nextF_rev [i _]]].
+apply Nat.neq_0_lt_0; contradict i; rewrite i; intros [j Hj]; easy.
+Qed.
+
+Lemma is_lin_dep_zero_equiv :
+  forall {n}, is_lin_dep (0 : 'E^n) <-> (0 < n)%coq_nat.
+Proof.
+intros; split. apply is_lin_dep_zero_rev.
+destruct n as [| n]; try easy; intros _; apply is_lin_dep_zero.
+Qed.
+
+Lemma is_lin_dep_not_injF :
+  forall {n} {B : 'E^n}, ~ injective B -> is_lin_dep B.
+Proof.
+move=> n B /not_all_ex_not [i0 /not_all_ex_not [i1 H]].
+destruct n; try now destruct i0.
+apply imply_to_and in H; destruct H as [HB Hi]; apply not_eq_sym in Hi.
+destruct n; try now rewrite 2!ord1 in Hi.
+rewrite is_lin_dep_ex; exists (kron _ i0 - kron _ i1); split.
+(* *)
+rewrite comb_lin_plus_l comb_lin_opp_l 2!comb_lin_kron_in_r -HB.
+apply: minus_eq_zero.
+(* *)
+apply nextF; exists i0.
+rewrite -> fct_minus_eq, kron_is_1, kron_is_0; try easy.
+rewrite minus_zero_r; apply one_not_zero; easy.
+apply ord_neq_compat; easy.
+Qed.
+
+Lemma is_lin_dep_castF_compat :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_lin_dep B1 -> is_lin_dep (castF H B1).
+Proof.
+intros n1 n2 H B1; rewrite 2!is_lin_dep_ex.
+intros [L [HL HL0]]; exists (castF H L); split; try now rewrite comb_lin_castF.
+contradict HL0; apply (castF_zero_reg H); easy.
+Qed.
+
+Lemma is_lin_dep_castF_reg :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_lin_dep (castF H B1) -> is_lin_dep B1.
+Proof.
+intros n1 n2 H B1 HB; rewrite -(castF_can _ (eq_sym H) B1).
+apply is_lin_dep_castF_compat; easy.
+Qed.
+
+Lemma is_lin_dep_castF_equiv :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_lin_dep (castF H B1) <-> is_lin_dep B1.
+Proof.
+intros; split. apply is_lin_dep_castF_reg. apply is_lin_dep_castF_compat.
+Qed.
+
+Lemma is_lin_dep_concatF_compat_l :
+  forall {n1 n2} {B1 : 'E^n1} (B2 : 'E^n2),
+    is_lin_dep B1 -> is_lin_dep (concatF B1 B2).
+Proof.
+move=>>; rewrite 2!is_lin_dep_ex; intros [L HL]; exists (concatF L 0); split.
+rewrite comb_lin_concatF_zero_lr; easy.
+apply concatF_zero_nextF_compat_l; easy.
+Qed.
+
+Lemma is_lin_dep_concatF_compat_r :
+  forall {n1 n2} (B1 : 'E^n1) {B2 : 'E^n2},
+    is_lin_dep B2 -> is_lin_dep (concatF B1 B2).
+Proof.
+move=>>; rewrite 2!is_lin_dep_ex; intros [L HL]; exists (concatF 0 L); split.
+rewrite comb_lin_concatF_zero_ll; easy.
+apply concatF_zero_nextF_compat_r; easy.
+Qed.
+
+Lemma is_lin_dep_concatF_not_disjF :
+  forall {n1 n2} {B1 : 'E^n1} {B2 : 'E^n2} x,
+    inF x B1 -> inF x B2 -> is_lin_dep (concatF B1 B2).
+Proof.
+intros n1 n2 B1 B2 x [i1 Hx1] [i2 Hx2]; apply is_lin_dep_not_injF.
+assert (Hf : (first_ord n2 i1 < n1)%coq_nat) by now apply /ltP; simpl.
+assert (Hl : ~ (last_ord n1 i2 < n1)%coq_nat)
+    by (apply Nat.nlt_ge; apply /leP; apply leq_addr).
+apply ex_not_not_all; exists (first_ord n2 i1).
+apply ex_not_not_all; exists (last_ord n1 i2).
+rewrite not_imp_and_equiv; split; try apply ord_neq; auto with zarith.
+rewrite concatF_correct_l concat_l_first
+    concatF_correct_r concat_r_last -Hx1; easy.
+Qed.
+
+Lemma is_lin_dep_insertF_compat :
+  forall {n} {B : 'E^n} x0 i0, is_lin_dep B -> is_lin_dep (insertF B x0 i0).
+Proof.
+intros n B x0 i0; rewrite 2!is_lin_dep_ex; move=> [L [HL /nextF_rev [i Hi]]].
+exists (insertF L 0 i0); split.
+rewrite comb_lin_insertF scal_zero_l HL plus_zero_l; easy.
+apply nextF; exists (skip_ord i0 i); rewrite insertF_correct; easy.
+Qed.
+
+Lemma is_lin_dep_skipF_reg :
+  forall {n} {B : 'E^n.+1} i0, is_lin_dep (skipF B i0) -> is_lin_dep B.
+Proof.
+intros n B i0; rewrite -{2}(insertF_skipF B i0).
+apply is_lin_dep_insertF_compat.
+Qed.
+
+Lemma is_lin_dep_in_line :
+  forall {n} {B : 'E^n} {i j}, i <> j -> in_line (B i) (B j) -> is_lin_dep B.
+Proof.
+move=> n B i j Hi /span_EX [Lj HLj]; apply is_lin_dep_ex;
+    exists (kron _ j - scal (Lj ord0) (kron _ i)); split.
+(* *)
+rewrite comb_lin_minus_l comb_lin_scal_l 2!comb_lin_kron_in_r HLj.
+rewrite comb_lin_1; apply: minus_eq_zero.
+(* *)
+apply nextF; exists j.
+rewrite -> fct_minus_eq, fct_scal_eq, kron_is_1, kron_is_0; try easy.
+rewrite scal_zero_r minus_zero_r; apply one_not_zero; easy.
+intros; contradict Hi; apply ord_inj; easy.
+Qed.
+
+Lemma is_lin_dep_in_plane :
+  forall {n} {B : 'E^n} {i j k},
+    i <> k -> j <> k -> in_plane (B i) (B j) (B k) -> is_lin_dep B.
+Proof.
+move=> n B i j k Hi Hj /span_EX [Lk HLk]; apply is_lin_dep_ex;
+  exists (kron _ k
+      - scal (Lk ord0) (kron _ i) - scal (Lk ord_max) (kron _ j)); split.
+(* *)
+rewrite 2!comb_lin_minus_l 2!comb_lin_scal_l 3!comb_lin_kron_in_r HLk.
+rewrite comb_lin_2 coupleF_0 coupleF_1; apply: minus2_eq_zero.
+(* *)
+apply nextF; exists k.
+rewrite -> 2!fct_minus_eq, 2!fct_scal_eq,
+    kron_is_1, 2!kron_is_0; try apply ord_neq_compat; try easy.
+rewrite 2!scal_zero_r 2!minus_zero_r; apply one_not_zero; easy.
+Qed.
+
+Lemma is_free_without_zero : forall {n} {B : 'E^n}, is_free B -> ~ inF 0 B.
+Proof. move=>>; rewrite contra_not_r_equiv; apply is_lin_dep_with_zero. Qed.
+
+Lemma is_free_zero_rev : forall {n}, is_free (0 : 'E^n) -> n = O.
+Proof.
+intros n; rewrite -Nat.le_0_r -Nat.nlt_ge contra_not_r_equiv.
+destruct n as [| n]; try easy; intros _; apply is_lin_dep_zero.
+Qed.
+
+Lemma is_free_zero_equiv : forall {n}, is_free (0 : 'E^n) <-> n = O.
+Proof.
+move=>>; rewrite -Nat.le_0_r -Nat.nlt_ge iff_not_r_equiv.
+apply is_lin_dep_zero_equiv.
+Qed.
+
+Lemma is_free_nil_or_nonzero :
+  forall {n} {B : 'E^n}, is_free B -> n = O \/ B <> 0.
+Proof.
+intros n; destruct n as [| n]; [intros; left; easy |].
+move=>> HB; right; contradict HB; rewrite HB; apply is_lin_dep_zero.
+Qed.
+
+Lemma is_free_nonzero : forall {n} {B : 'E^n.+1}, is_free B -> B <> 0.
+Proof. move=>> /is_free_nil_or_nonzero [H | H]; easy. Qed.
+
+Lemma is_free_nonnil_or_zero :
+  forall {n} {B : 'E^n}, is_free B -> (0 < n)%coq_nat \/ B = 0.
+Proof.
+intros n; destruct n as [| n]; [intros; right; apply nil_is_zero |].
+move=>>; left; auto with arith.
+Qed.
+
+Lemma is_free_nonnil :
+  forall {n} (B : 'E^n), is_free B -> B <> 0 -> (0 < n)%coq_nat.
+Proof. move=>> /is_free_nonnil_or_zero [H | H]; easy. Qed.
+
+Lemma is_free_nonnil_nonzero_equiv :
+  forall {n} {B : 'E^n}, is_free B -> (0 < n)%coq_nat <-> B <> 0.
+Proof.
+intros; apply iff_sym; split; [apply is_free_nonnil; easy |].
+destruct n as [| n]; try easy; intros _; apply is_free_nonzero; easy.
+Qed.
+
+Lemma is_free_nil_zero_equiv :
+  forall {n} {B : 'E^n}, is_free B -> n = O <-> B = 0.
+Proof.
+intros; rewrite -Nat.le_0_r iff_not_equiv Nat.nle_gt.
+apply is_free_nonnil_nonzero_equiv; easy.
+Qed.
+
+Lemma is_free_injF :
+  forall {n} {B : 'E^n}, is_free B -> injective B.
+Proof. move=>>; rewrite contra_equiv; apply is_lin_dep_not_injF. Qed.
+
+Lemma is_free_castF_compat :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_free B1 -> is_free (castF H B1).
+Proof. move=>>; apply contra_equiv, is_lin_dep_castF_reg. Qed.
+
+Lemma is_free_castF_reg :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_free (castF H B1) -> is_free B1.
+Proof. move=>>; apply contra_equiv, is_lin_dep_castF_compat. Qed.
+
+Lemma is_free_castF_equiv :
+  forall {n1 n2} (H : n1 = n2) {B1 : 'E^n1},
+    is_free (castF H B1) <-> is_free B1.
+Proof. move=>>; apply iff_not_equiv, is_lin_dep_castF_equiv. Qed.
+
+Lemma is_free_concatF_reg_l :
+  forall {n1 n2} {B1 : 'E^n1} (B2 : 'E^n2),
+    is_free (concatF B1 B2) -> is_free B1.
+Proof. move=>>; apply contra_equiv; apply is_lin_dep_concatF_compat_l. Qed.
+
+Lemma is_free_concatF_reg_r :
+  forall {n1 n2} (B1 : 'E^n1) {B2 : 'E^n2},
+    is_free (concatF B1 B2) -> is_free B2.
+Proof. move=>>; apply contra_equiv; apply is_lin_dep_concatF_compat_r. Qed.
+
+Lemma is_free_concatF_disjF :
+  forall {n1 n2} {B1 : 'E^n1} {B2 : 'E^n2} x,
+    is_free (concatF B1 B2) -> ~ inF x B1 \/ ~ inF x B2.
+Proof.
+move=>>; rewrite -not_and_equiv contra_not_r_equiv.
+intros [H1 H2]; eapply is_lin_dep_concatF_not_disjF; [apply H1 | easy].
+Qed.
+
+Lemma is_free_insertF_reg :
+  forall {n} {B : 'E^n} x0 i0, is_free (insertF B x0 i0) -> is_free B.
+Proof. move=>>; rewrite contra_equiv; apply is_lin_dep_insertF_compat. Qed.
+
+Lemma is_free_skipF_compat :
+  forall {n} {B : 'E^n.+1} i0, is_free B -> is_free (skipF B i0).
+Proof. move=>>; rewrite contra_equiv; apply is_lin_dep_skipF_reg. Qed.
+
+Lemma is_free_not_in_line :
+  forall {n} {B : 'E^n},
+    is_free B -> forall i j, i <> j -> ~ in_line (B i) (B j).
+Proof.
+move=>> HB; move=>> Hi H; contradict HB;
+    eapply is_lin_dep_in_line; [apply Hi | apply H].
+Qed.
+
+Lemma is_free_not_in_plane :
+  forall {n} {B : 'E^n},
+    is_free B ->
+    forall i j k, i <> k -> j <> k -> ~ in_plane (B i) (B j) (B k).
+Proof.
+move=>> HB; move=>> Hi Hj H; contradict HB;
+    eapply is_lin_dep_in_plane; [apply Hi | apply Hj | apply H].
+Qed.
+
+(** Establish first the property on is_free. *)
+
+Lemma is_free_uniq_decomp :
+  forall {n} {B : 'E^n}, is_free B -> comb_lin_injL B.
+Proof.
+intros n B HB L L' HL%minus_zero_compat.
+rewrite -comb_lin_minus_l in HL; apply minus_zero_reg; auto.
+Qed.
+
+Lemma is_free_uniq_decomp_rev :
+  forall {n} {B : 'E^n}, comb_lin_injL B -> is_free B.
+Proof.
+intros n B HB L HL; apply HB; rewrite (comb_lin_zero_compat_l 0); easy.
+Qed.
+
+Lemma is_free_uniq_decomp_equiv :
+  forall {n} {B : 'E^n}, is_free B <-> comb_lin_injL B.
+Proof.
+intros; split. apply is_free_uniq_decomp. apply is_free_uniq_decomp_rev.
+Qed.
+
+Lemma is_free_nil : forall (B : 'E^0), is_free B.
+Proof. intros B L HL; apply extF; intros [i Hi]; easy. Qed.
+
+Lemma is_free_permutF :
+  forall {n} {p} {B : 'E^n}, bijective p -> is_free B -> is_free (permutF p B).
+Proof.
+intros n p B Hp HB L; rewrite {1}(permutF_f_inv_r Hp L) comb_lin_permutF;
+    try now apply bij_inj.
+move=> /HB /extF_rev HL; apply extF; intros i.
+specialize (HL (p i)); rewrite zeroF in HL; rewrite zeroF -HL.
+unfold permutF; rewrite f_inv_can_l; easy.
+Qed.
+
+Lemma is_free_permutF_equiv :
+  forall {n} {p} {B : 'E^n},
+    bijective p -> is_free B <-> is_free (permutF p B).
+Proof.
+intros n p B Hp; split. apply is_free_permutF; easy.
+rewrite <- (permutF_can p (f_inv Hp) B), permutF_can.
+apply is_free_permutF, f_inv_bij.
+apply f_inv_can_l.
+apply f_inv_can_r.
+Qed.
+
+Lemma is_lin_dep_not_uniq_decomp :
+  forall {n} {B : 'E^n}, is_lin_dep B -> ~ comb_lin_injL B.
+Proof. move=>>; rewrite -contra_equiv; apply is_free_uniq_decomp_rev. Qed.
+
+Lemma is_lin_dep_not_uniq_decomp_rev :
+  forall {n} {B : 'E^n}, ~ comb_lin_injL B -> is_lin_dep B.
+Proof. move=>>; rewrite -contra_equiv; apply is_free_uniq_decomp. Qed.
+
+Lemma is_lin_dep_not_uniq_decomp_equiv :
+  forall {n} {B : 'E^n}, is_lin_dep B <-> ~ comb_lin_injL B.
+Proof. move=>>; apply not_iff_compat, is_free_uniq_decomp_equiv. Qed.
+
+Lemma is_lin_dep_nonnil :
+  forall {n} (B : 'E^n), is_lin_dep B -> (0 < n)%coq_nat.
+Proof.
+move=>>; rewrite contra_not_l_equiv Nat.nlt_ge Nat.le_0_r.
+intros; subst; apply is_free_nil.
+Qed.
+
+(** Mixed properties of is_lin_dep/is_free. *)
+
+(* Gostiaux T1, Th 6.61 pp. 189-190. (=>) *)
+Lemma is_lin_dep_insertF :
+  forall {n} {B : 'E^n} {x0} i0, span B x0 -> is_lin_dep (insertF B x0 i0).
+Proof.
+intros n B x0 i0 [L]; apply is_lin_dep_ex; exists (insertF L (- 1%K) i0); split.
+rewrite comb_lin_insertF scal_opp_l scal_one; apply: plus_opp_l.
+rewrite -(insertF_zero i0); apply insertF_nextF_compat_r.
+apply one_not_zero in HK; contradict HK; apply opp_inj; rewrite opp_zero; easy.
+Qed.
+
+Lemma is_free_insertF_rev :
+  forall {n} {B : 'E^n} {x0} i0, is_free (insertF B x0 i0) -> ~ span B x0.
+Proof. move=>>; rewrite contra_not_r_equiv; apply is_lin_dep_insertF. Qed.
+
+End Free_Facts.
+
+
+Section Free_sub_Facts1a.
+
+(** Properties of is_free on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_ms PE.
+Let PE_ms := sub_ms HPE.
+
+Context {n : nat}.
+Context {B : 'E^n}.
+Hypothesis HB : inclF B PE.
+Let B_ms : 'PE_ms^n := fun i => mk_sub_ms (HB i).
+
+Lemma is_free_sub : is_free B -> is_free B_ms.
+Proof.
+intros HB1 L HL; apply HB1;
+    rewrite mk_sub_comb_lin in HL; apply mk_sub_inj in HL; easy.
+Qed.
+
+Lemma is_free_sub_rev : is_free B_ms -> is_free B.
+Proof.
+intros HB1 L HL; apply HB1;
+    apply val_inj; rewrite val_comb_lin; easy.
+Qed.
+
+Lemma is_free_sub_equiv : is_free B_ms <-> is_free B.
+Proof. split; [apply is_free_sub_rev | apply is_free_sub]. Qed.
+
+End Free_sub_Facts1a.
+
+
+Section Free_sub_Facts1b.
+
+(** Properties of is_free on sub-module spaces. *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+Hypothesis HPE : compatible_ms PE.
+Let PE_ms := sub_ms HPE.
+Context {n : nat}.
+Context {B_ms : 'PE_ms^n}.
+Let B := mapF val B_ms.
+
+Lemma is_free_val : is_free B_ms -> is_free B.
+Proof.
+rewrite -(is_free_sub_equiv HPE (in_subF B_ms)) -(mk_subF_eq B_ms); easy.
+Qed.
+
+Lemma is_free_val_rev : is_free B -> is_free B_ms.
+Proof. rewrite (mk_subF_eq B_ms) (is_free_sub_equiv HPE); easy. Qed.
+
+Lemma is_free_val_equiv : is_free B <-> is_free B_ms.
+Proof. rewrite (mk_subF_eq B_ms) (is_free_sub_equiv HPE); easy. Qed.
+
+End Free_sub_Facts1b.
+
+
+Section Basis_Facts1.
+
+(** Properties of is_basis (from those of is_generator and is_free). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Hypothesis HK : nonzero_struct K.
+
+Context {PE : E -> Prop}.
+
+Context {n n1 n2 : nat}.
+
+Context {B : 'E^n}.
+Context {B1 : 'E^n1}.
+Context {B2 : 'E^n2}.
+
+Lemma is_basis_span_equiv : is_basis (span B) B <-> is_free B.
+Proof.
+assert (H : forall P Q R : Prop, P -> Q <-> R -> P /\ Q <-> R) by tauto.
+intros; apply H; easy.
+Qed.
+
+Lemma is_basis_injF : is_basis PE B -> injective B.
+Proof. intros [_ HB]; apply is_free_injF; easy. Qed.
+
+Lemma is_basis_castF_equiv :
+  forall (H : n1 = n2), is_basis PE (castF H B1) <-> is_basis PE B1.
+Proof.
+intros; unfold is_basis;
+    rewrite is_generator_castF_equiv is_free_castF_equiv; easy.
+Qed.
+
+Lemma is_basis_castF_compat :
+  forall (H : n1 = n2), is_basis PE B1 -> is_basis PE (castF H B1).
+Proof. apply is_basis_castF_equiv. Qed.
+
+Lemma is_basis_castF_reg :
+  forall (H : n1 = n2), is_basis PE (castF H B1) -> is_basis PE B1.
+Proof. apply is_basis_castF_equiv. Qed.
+
+Lemma is_basis_is_generator : is_basis PE B -> is_generator PE B.
+Proof. intros [HB _]; easy. Qed.
+
+Lemma is_basis_is_free : forall PE {n} {B : 'E^n}, is_basis PE B -> is_free B.
+Proof. move=>> [_ HB]; easy. Qed.
+
+Lemma is_basis_inclF : is_basis PE B -> inclF B PE.
+Proof. intros [HB _]; apply is_generator_inclF; easy. Qed.
+
+Lemma is_basis_zero_closed :
+  forall (B : 'E^n), is_basis PE B -> PE 0.
+Proof. move=>> [HB _]; apply (is_generator_zero_closed _ HB). Qed.
+
+Lemma is_basis_plus_closed :
+  forall (B : 'E^n), is_basis PE B -> plus_closed PE.
+Proof. move=>> [HB _]; apply (is_generator_plus_closed _ HB). Qed.
+
+Lemma is_basis_opp_closed :
+  forall (B : 'E^n), is_basis PE B -> opp_closed PE.
+Proof. move=>> [HB _]; apply (is_generator_opp_closed _ HB). Qed.
+
+Lemma is_basis_minus_closed :
+  forall (B : 'E^n), is_basis PE B -> minus_closed PE.
+Proof. move=>> [HB _]; apply (is_generator_minus_closed _ HB). Qed.
+
+Lemma is_basis_scal_closed :
+  forall (B : 'E^n), is_basis PE B -> scal_closed PE.
+Proof. move=>> [HB _]; apply (is_generator_scal_closed _ HB). Qed.
+
+Lemma is_basis_compatible_ms :
+  forall (B : 'E^n), is_basis PE B -> compatible_ms PE.
+Proof. move=>> [HB _]; apply (is_generator_compatible_ms _ HB). Qed.
+
+Lemma is_basis_comb_lin_closed :
+  forall (B : 'E^n), is_basis PE B -> comb_lin_closed PE.
+Proof. move=>> [HB _]; apply (is_generator_comb_lin_closed _ HB). Qed.
+
+Lemma is_basis_ex_decomp :
+  forall {PE} {B : 'E^n}, is_basis PE B -> comb_lin_surjL PE B.
+Proof. move=>> [HB _]; apply is_generator_ex_decomp; easy. Qed.
+
+Lemma is_basis_uniq_decomp : forall PE, is_basis PE B -> comb_lin_injL B.
+Proof. move=>> [_ HB]; apply is_free_uniq_decomp; easy. Qed.
+
+Theorem is_basis_ex_uniq_decomp :
+  forall PE, is_basis PE B -> comb_lin_bijL PE B.
+Proof.
+intros PE' HB; apply comb_lin_surjL_injL_bijL.
+apply is_basis_ex_decomp; easy.
+apply (is_basis_uniq_decomp PE'); easy.
+Qed.
+
+Theorem is_basis_ex_uniq_decomp_rev :
+  compatible_ms PE -> inclF B PE -> comb_lin_bijL PE B -> is_basis PE B.
+Proof.
+intros HPE HB1 HB2; split.
+apply is_generator_uniq_decomp_rev; easy.
+apply is_free_uniq_decomp_rev, (comb_lin_bijL_injL PE); easy.
+Qed.
+
+Lemma is_basis_ex_uniq_decomp_equiv :
+  compatible_ms PE -> inclF B PE -> is_basis PE B <-> comb_lin_bijL PE B.
+Proof.
+intros; split.
+apply is_basis_ex_uniq_decomp. apply is_basis_ex_uniq_decomp_rev; easy.
+Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (27/02/2024): reproduce sections Generator_sub_Facts* for is_basis. *)
+
+Lemma is_basis_fullset_compat_gen :
+  forall (HPE : compatible_ms PE) (HB : inclF B PE),
+    let PE_sub := @fullset (sub_ms HPE) in
+    let B_sub i := mk_sub_ms (HB i) in
+    is_basis PE B -> is_basis PE_sub B_sub.
+Proof.
+move=>> [HB1 HB2]; split.
+apply is_generator_sub; easy.
+apply is_free_sub; easy.
+Qed.
+
+Lemma is_basis_fullset_compat_gen_equiv :
+  forall (HPE : compatible_ms PE) (HB : inclF B PE),
+    let PE_sub := @fullset (sub_ms HPE) in
+    let B_sub i := mk_sub_ms (HB i) in
+    is_basis PE B <-> is_basis PE_sub B_sub.
+Proof.
+intros HPE HB PE_sub B_sub; split; intros (HB1,HB2); split.
+apply is_generator_sub; easy.
+apply is_free_sub; easy.
+(* *)
+2: generalize HB2; apply is_free_sub_equiv.
+apply subset_ext_equiv; split; try easy.
+intros t Ht.
+pose (tt:= (mk_sub_ms_ HPE Ht)).
+assert (H1: span B_sub tt).
+rewrite -HB1; easy.
+inversion H1 as (L,HL).
+apply span_ex; exists L.
+apply trans_eq with (val tt); try easy.
+rewrite -HL.
+rewrite val_comb_lin; easy.
+apply span_lb; easy.
+Qed.
+
+Lemma is_basis_fullset_compat :
+  forall (HB : is_basis PE B),
+    let HB1 := is_basis_inclF HB in
+    let HPE := is_basis_compatible_ms _ HB in
+    let PE_sub := @fullset (sub_ms HPE) in
+    let B_sub i := mk_sub_ms (HB1 i) in
+    is_basis PE_sub B_sub.
+Proof. intros; apply is_basis_fullset_compat_gen; easy. Qed.
+
+End Basis_Facts1.
+
+
+Section Basis_Facts2.
+
+(** More properties of is_basis (from those of is_generator and is_free). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Hypothesis HK : nonzero_struct K.
+
+Variable PE : E -> Prop.
+
+Lemma is_basis_nil : forall (B : 'E^0), is_basis zero_sub_struct B.
+Proof. intros; split; [apply is_generator_nil | apply is_free_nil]. Qed.
+
+Lemma is_basis_nonzero : forall {n} {B : 'E^n.+1}, is_basis PE B -> B <> 0.
+Proof. move=>> [_ /(is_free_nonzero HK) HB]; easy. Qed.
+
+Lemma is_basis_nonzero_rev :
+  forall {n} (B : 'E^n), is_basis PE B -> B <> 0 -> (0 < n)%coq_nat.
+Proof. move=>> [_ HB1] HB2; apply (is_free_nonnil _ HB1 HB2). Qed.
+
+Lemma is_basis_nonzero_equiv :
+  forall {n} (B : 'E^n), is_basis PE B -> B <> 0 <-> (0 < n)%coq_nat.
+Proof.
+intros n B HB; split. apply is_basis_nonzero_rev; easy.
+intros; destruct n as [| n]; try apply is_basis_nonzero; easy.
+Qed.
+
+Lemma is_basis_nonempty :
+  forall {PE} {n} (B : 'E^n), is_basis PE B -> nonempty PE.
+Proof. move=>> /is_basis_compatible_ms; apply compatible_ms_nonempty. Qed.
+
+Lemma is_basis_nonzero_sub_struct :
+  forall {PE} {n} (B : 'E^n.+1), is_basis PE B -> PE <> zero_sub_struct.
+Proof.
+move=> PE' n B [/is_generator_inclF HB /(is_free_nonzero HK) /nextF_rev [i Hi]].
+apply nonzero_sub_struct_ex; exists (B i); easy.
+Qed.
+
+Lemma is_basis_nonzero_sub_struct_rev :
+  forall {n} (B : 'E^n), is_basis PE B -> PE <> zero_sub_struct -> (0 < n)%coq_nat.
+Proof.
+move=>> HB /nonzero_sub_struct_ex HPE; specialize (HPE (is_basis_nonempty _ HB)).
+apply (is_basis_nonzero_rev _ HB); contradict HPE; rewrite HPE in HB.
+intros [x [Hx Hx2]]; contradict Hx2.
+destruct (is_basis_ex_decomp HB x Hx) as [L HL].
+rewrite HL comb_lin_zero_r; easy.
+Qed.
+
+Lemma is_basis_nonzero_sub_struct_equiv :
+  forall {n} (B : 'E^n), is_basis PE B -> PE <> zero_sub_struct <-> (0 < n)%coq_nat.
+Proof.
+intros n B HB; split. apply (is_basis_nonzero_sub_struct_rev _ HB).
+intros; destruct n as [| n]; try apply (is_basis_nonzero_sub_struct B); easy.
+Qed.
+
+End Basis_Facts2.
+
+
+Section Has_dim_Facts.
+
+(** Properties of has_dim (from those of is_basis). *)
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Hypothesis HK : nonzero_struct K.
+
+Context {PE : E -> Prop}.
+
+Context {n : nat}.
+
+Lemma has_dim_is_finite_dim : has_dim PE n -> is_finite_dim PE.
+Proof. intros [B [HB _]]; exists n, B; easy. Qed.
+
+Lemma has_dim_nonempty : has_dim PE n -> nonempty PE.
+Proof. intros [B HB]; apply (is_basis_nonempty _ HB). Qed.
+
+Lemma has_dim_nonzero_sub_struct : has_dim PE n.+1 -> PE <> zero_sub_struct.
+Proof. intros [B HB]; apply (is_basis_nonzero_sub_struct HK _ HB). Qed.
+
+Lemma has_dim_nonzero_sub_struct_rev :
+  has_dim PE n -> PE <> zero_sub_struct -> (0 < n)%coq_nat.
+Proof. intros [B HB]; apply (is_basis_nonzero_sub_struct_rev _ _ HB). Qed.
+
+Lemma has_dim_nonzero_sub_struct_equiv :
+  has_dim PE n -> PE <> zero_sub_struct <-> (0 < n)%coq_nat.
+Proof. intros [B HB]; apply (is_basis_nonzero_sub_struct_equiv HK _ _ HB). Qed.
+
+Lemma has_dim_zero_closed : has_dim PE n -> PE 0.
+Proof. intros [B HB]; apply (is_basis_zero_closed _ HB). Qed.
+
+Lemma has_dim_plus_closed : has_dim PE n -> plus_closed PE.
+Proof. intros [B HB]; apply (is_basis_plus_closed _ HB). Qed.
+
+Lemma has_dim_opp_closed : has_dim PE n -> opp_closed PE.
+Proof. intros [B HB]; apply (is_basis_opp_closed _ HB). Qed.
+
+Lemma has_dim_minus_closed : has_dim PE n -> minus_closed PE.
+Proof. intros [B HB]; apply (is_basis_minus_closed _ HB). Qed.
+
+Lemma has_dim_scal_closed : has_dim PE n -> scal_closed PE.
+Proof. intros [B HB]; apply (is_basis_scal_closed _ HB). Qed.
+
+Lemma has_dim_compatible_ms : has_dim PE n -> compatible_ms PE.
+Proof. intros [B HB]; apply (is_basis_compatible_ms _ HB). Qed.
+
+Lemma has_dim_comb_lin_closed : has_dim PE n -> comb_lin_closed PE.
+Proof. intros [B HB]; apply (is_basis_comb_lin_closed _ HB). Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (27/02/2024): reproduce sections Generator_sub_Facts* for has_dim. *)
+
+Lemma has_dim_fullset_compat_gen :
+  forall (HPE : compatible_ms PE),
+    let PE_sub := @fullset (sub_ms HPE) in
+    has_dim PE n -> has_dim PE_sub n.
+Proof.
+move=>> [B HB]; eexists.
+apply (is_basis_fullset_compat_gen _ (is_basis_inclF HB)); easy.
+Qed.
+
+Lemma has_dim_fullset_compat :
+  forall (HPE : has_dim PE n),
+    let PE_sub := @fullset (sub_ms (has_dim_compatible_ms HPE)) in
+    has_dim PE_sub n.
+Proof. intros; apply has_dim_fullset_compat_gen; easy. Qed.
+
+Lemma has_dim_0 : has_dim (@zero_sub_struct E) 0.
+Proof. apply (Dim _ _ (fun _ : 'I_0 => 0)), is_basis_nil. Qed.
+
+Lemma has_dim_0_rev : has_dim PE 0 -> PE = zero_sub_struct.
+Proof. intros [B [HB1 HB2]]; rewrite HB1; apply span_nil. Qed.
+
+Lemma has_dim_0_equiv : has_dim PE 0 <-> PE = zero_sub_struct.
+Proof.
+split. apply has_dim_0_rev. intros HPE; rewrite HPE; apply has_dim_0.
+Qed.
+
+Lemma is_free_has_dim_span :
+  forall {B : 'E^n}, is_free B -> has_dim (span B) n.
+Proof. intros B HB; apply (Dim _ _ B), is_basis_span_equiv; easy. Qed.
+
+Lemma is_free_has_dim_span_alt :
+  forall {p} {B : 'E^n}, p = n -> is_free B -> has_dim (span B) p.
+Proof. intros; subst; apply is_free_has_dim_span; easy. Qed.
+
+End Has_dim_Facts.
+
+
+Section Linear_mapping_Facts1.
+
+Context {K : Ring}.
+Context {E F : ModuleSpace K}.
+
+Variable f : E -> F.
+Hypothesis Hf : is_linear_mapping f.
+
+Variable PE : E -> Prop.
+Hypothesis HPE : compatible_ms PE.
+
+Lemma is_linear_mapping_is_free_compat :
+  forall {n} (B : 'E^n),
+    injS PE f -> (forall i, PE (B i)) -> is_free B -> is_free (mapF f B).
+Proof.
+intros n B H1 H2 H3 L HL.
+rewrite <- linear_mapping_comb_lin_compat in HL; try easy.
+apply H3.
+apply H1.
+apply compatible_ms_comb_lin; try easy.
+apply compatible_ms_zero; easy.
+rewrite is_linear_mapping_zero; easy.
+Qed.
+
+Lemma is_linear_mapping_is_generator_compat :
+  forall {n} (B : 'E^n),
+    (forall i, PE (B i)) ->
+    is_generator PE B -> is_generator (image f PE) (mapF f B).
+Proof.
+intros n B H1 H2.
+apply subset_ext_equiv; split; intros x Hx.
+(* *)
+induction Hx as [y Hy].
+rewrite H2 in Hy.
+destruct Hy as [L].
+rewrite linear_mapping_comb_lin_compat; easy.
+(* *)
+destruct Hx as [L].
+rewrite <- linear_mapping_comb_lin_compat; try easy.
+apply Im.
+apply compatible_ms_comb_lin; try easy.
+Qed.
+
+Lemma is_linear_mapping_is_basis_compat :
+  forall n (B : 'E^n),
+    injS PE f -> (forall i, PE (B i)) ->
+    is_basis PE B -> is_basis (image f PE) (mapF f B).
+Proof.
+move=>> Hf1 HB1 [HB2 HB3]; split.
+apply is_linear_mapping_is_generator_compat; easy.
+apply is_linear_mapping_is_free_compat; easy.
+Qed.
+
+Lemma is_linear_mapping_preimageF :
+  forall {n} (B : 'F^n), inclF B (image f PE) ->
+    exists (A : 'E^n), inclF A PE /\ B = mapF f A.
+Proof.
+intros n B HB; rewrite image_eq in HB.
+destruct (choice _ HB) as [A HA]; exists A; split;
+    [| apply extF]; intro; apply HA.
+Qed.
+
+End Linear_mapping_Facts1.
+
+
+Section Linear_mapping_Facts2.
+
+Context {K : Ring}.
+Context {E F : ModuleSpace K}.
+
+Variable f : E -> F.
+Hypothesis Hf : is_linear_mapping f.
+
+Variable PE : E -> Prop.
+Hypothesis HPE : compatible_ms PE.
+
+Lemma is_linear_mapping_is_basis_compat_equiv :
+  forall n (B : 'E^n),
+    bijS PE (image f PE) f (* injS ? *) -> (forall i, PE (B i)) ->
+    is_basis PE B <-> is_basis (image f PE) (fun i => f (B i)).
+Proof.
+intros n B Hf1 HB1; split.
+(* *)
+apply bijS_equiv in Hf1; try apply inhabited_ms.
+apply is_linear_mapping_is_basis_compat; easy.
+(* *)
+intros HB2.
+apply (is_basis_fullset_compat_gen_equiv HPE HB1).
+pose (PEs := sub_ms HPE).
+assert (HPI: compatible_ms (image f PE)).
+apply compatible_ms_image; easy.
+pose (PIs := sub_ms HPI).
+pose (Bs := fun i : 'I_n => mk_sub_ms (HB1 i): PEs); fold Bs.
+assert (V: forall t:PEs, (image f PE) (f (val t))).
+intros t; apply Im; try easy.
+destruct t; easy.
+pose (fs := fun t:PEs => mk_sub_ms (V t): PIs).
+pose (fBs := fun i => fs (Bs i)).
+(* *)
+assert (Hfs: bijective fs).
+apply sub_ms_bij with f; try easy.
+pose (gs:= f_inv Hfs).
+apply is_basis_ext with (mapF gs fBs).
+apply extF; intros i.
+rewrite mapF_correct; unfold fBs, gs.
+now rewrite f_inv_can_l.
+assert (W: (@fullset (@sub_ms K E PE HPE) = image gs (image fs fullset))).
+apply subset_ext_equiv; split; intros x Hx; try easy.
+rewrite -(f_inv_can_l Hfs x); easy.
+assert (Hfs2 : is_linear_mapping fs).
+apply sub_ms_linear_mapping with f; try easy.
+(* *)
+rewrite W; apply is_linear_mapping_is_basis_compat; try easy.
+apply is_linear_mapping_bij_compat; try easy.
+apply compatible_ms_image; try easy.
+apply compatible_ms_fullset.
+apply (bijS_injS fullset).
+replace (image fs fullset) with (@fullset PIs).
+apply bij_S_equiv; try apply inhabited_ms.
+apply f_inv_bij.
+apply subset_ext_equiv; split; intros (x,Hx1) Hx2; try easy.
+rewrite image_eq; inversion Hx1 as (y,Hy).
+exists (mk_sub_ms Hy); split; try easy.
+apply val_inj; simpl; easy.
+assert (T: inclF (fun i : 'I_n => f (B i)) (image f PE)).
+intros i; easy.
+(* *)
+generalize (is_basis_fullset_compat_gen_equiv 
+    HPI T); intros (TT1,_).
+specialize (TT1 HB2).
+apply is_basis_ext' with (3:=TT1).
+apply subset_ext_equiv; split; intros (x,Hx1) Hx2; try easy.
+rewrite image_eq; inversion Hx1 as (y,Hy).
+exists (mk_sub_ms Hy); split; try easy.
+apply val_inj; simpl; easy.
+apply extF; intros i.
+apply val_inj; now simpl.
+Qed.
+
+End Linear_mapping_Facts2.
+
+
+Section AffineSpace_FD_Def.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (07/06/2023): try to use new span_struct. *)
+
+Inductive aff_span {n} (A : 'E^n) : E -> Prop :=
+  Aff_span : forall L, invertible (sum L) -> aff_span A (barycenter L A).
+
+Lemma aff_span_EX :
+  forall {n} (A : 'E^n) G,
+    aff_span A G -> { L | sum L = 1 /\ G = barycenter L A }.
+Proof.
+move=>> Hx; apply ex_EX; induction Hx as [L HL];
+    eexists; apply (barycenter_normalized _ _ _ HL); easy.
+Qed.
+
+Lemma aff_span_ex :
+  forall {n} (A : 'E^n) G,
+    (exists L, invertible (sum L) /\ G = barycenter L A) -> aff_span A G.
+Proof. move=>> [L [HL1 HL2]]; rewrite HL2; easy. Qed.
+
+Definition in_aff_point (A0 : E) := aff_span (singleF A0).
+Definition in_aff_line (A0 A1 : E) := aff_span (coupleF A0 A1).
+Definition in_aff_plane (A0 A1 A2 : E) := aff_span (tripleF A0 A1 A2).
+(*Definition in_aff_3d_space (A0 A1 A2 A3 : E) := aff_span (quadrupleF A0 A1 A2 A3).*)
+
+Definition barycenter_surjL (PE : E -> Prop) {n} (A : 'E^n) :=
+  forall G, PE G -> exists L, invertible (sum L) /\ G = barycenter L A.
+Definition barycenter_injL {n} (A : 'E^n) :=
+  forall L L', invertible (sum L) -> sum L = sum L' ->
+    barycenter L A = barycenter L' A -> L = L'.
+Definition barycenter_bijL (PE : E -> Prop) {n} (A : 'E^n) :=
+  forall G, PE G -> exists! L, sum L = 1 /\ G = barycenter L A.
+
+Definition is_aff_gen (PE : E -> Prop) {n} (A : 'E^n) := PE = aff_span A.
+
+Lemma is_aff_gen_ext :
+  forall (PE : E -> Prop) {n} (A1 A2 : 'E^n),
+    A1 = A2 -> is_aff_gen PE A1 -> is_aff_gen PE A2.
+Proof. intros; subst; easy. Qed.
+
+Definition is_aff_indep {n} (A : 'E^n.+1) :=
+  forall i0, is_free (frameF A i0).
+
+Lemma is_aff_indep_ext :
+  forall {n} (A1 A2 : 'E^n.+1), A1 = A2 -> is_aff_indep A1 -> is_aff_indep A2.
+Proof. intros; subst; easy. Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC: this could be a notation... *)
+Definition is_aff_dep {n} (A : 'E^n.+1) := ~ is_aff_indep A.
+
+Lemma is_aff_dep_ext :
+  forall {n} (A1 A2 : 'E^n.+1), A1 = A2 -> is_aff_dep A1 -> is_aff_dep A2.
+Proof. intros; subst; easy. Qed.
+
+Definition is_aff_basis (PE : E -> Prop) {n} (A : 'E^n.+1) :=
+  is_aff_gen PE A /\ is_aff_indep A.
+
+Lemma is_aff_basis_ext :
+  forall (PE : E -> Prop) {n} (A1 A2 : 'E^n.+1),
+    A1 = A2 -> is_aff_basis PE A1 -> is_aff_basis PE A2.
+Proof. intros; subst; easy. Qed.
+
+Inductive has_aff_dim (PE : E -> Prop) n : Prop :=
+  Aff_dim : forall (A : 'E^n.+1), is_aff_basis PE A -> has_aff_dim PE n.
+
+Lemma has_aff_dim_EX :
+  forall (PE : E -> Prop) n, has_aff_dim PE n -> { A : 'E^n.+1 | is_aff_basis PE A }.
+Proof. move=>> H; apply ex_EX; induction H as [A]; exists A; easy. Qed.
+
+Lemma has_aff_dim_ex:
+  forall (PE : E -> Prop) n,
+    (exists A : 'E^n.+1, is_aff_basis PE A) -> has_aff_dim PE n.
+Proof. move=>> [A HA]. apply (Aff_dim _ _ A HA). Qed.
+
+Lemma has_aff_dim_ext :
+  forall (PE1 PE2 : E -> Prop) n1 n2,
+    PE1 = PE2 -> n1 = n2 -> has_aff_dim PE1 n1 -> has_aff_dim PE2 n2.
+Proof. intros; subst; easy. Qed.
+
+End AffineSpace_FD_Def.
+
+
+Section Aff_span_Subset_Facts.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma aff_span_eq :
+  forall {n} (A : 'E^n), aff_span A = barycenter_closure (inF^~ A).
+Proof.
+intros n A; apply subset_ext_equiv; split; intros C HC.
+induction HC; apply Barycenter_closure; [easy | apply invalF_refl].
+induction HC as [m L C HL HC].
+destruct (barycenter_injF_ex L C HL) as [p [N [D [HN [HD1 [HD2 ->]]]]]].
+clear L HL; apply ((invalF_trans _ D _)^~ HC) in HD2.
+destruct (barycenter_unfun0F N D A) as [L [HL1 [HL2 ->]]]; easy.
+Qed.
+
+Lemma aff_span_inclF_diag : forall {n} (A : 'E^n), inclF A (aff_span A).
+Proof.
+intros n A i; rewrite -barycenter_kron_r; apply Aff_span.
+rewrite sum_kron_r; apply invertible_one.
+Qed.
+
+Lemma aff_span_ub : forall {n} (A : 'E^n) C i, C = A i -> aff_span A C.
+Proof. intros; subst; apply aff_span_inclF_diag. Qed.
+
+Lemma aff_span_inclF :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2), invalF A1 A2 -> inclF A1 (aff_span A2).
+Proof.
+move=>> HA i; destruct (HA i) as [i2 Hi2]; rewrite Hi2; apply aff_span_inclF_diag.
+Qed.
+
+Lemma aff_spanF_ex :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2),
+    inclF A1 (aff_span A2) <->
+    exists M12, forall i1,
+      invertible (sum (M12 i1)) /\ A1 i1 = barycenter (M12 i1) A2.
+Proof.
+intros n1 n2 A1 A2; split; intros HA.
+(* *)
+assert (HA' : forall i1, exists L2,
+    invertible (sum L2) /\ A1 i1 = barycenter L2 A2).
+  intros i1; destruct (HA i1) as [L2]; exists L2; easy.
+apply (choice _ HA').
+(* *)
+intros i1; destruct HA as [M HM], (HM i1) as [HM1a HM1b]; rewrite HM1b; easy.
+Qed.
+
+Lemma aff_spanF_ex1 :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2),
+    inclF A1 (aff_span A2) <->
+    exists M12, forall i1, sum (M12 i1) = 1 /\ A1 i1 = barycenter (M12 i1) A2.
+Proof.
+move=>>; rewrite aff_spanF_ex; split; move=> [M12 /all_and_equiv [HM12 HA1]].
+(* *)
+exists (fun i1 => scal (/ sum (M12 i1)) (M12 i1)); intros i1.
+apply barycenter_normalized; easy.
+(* *)
+exists M12; intros; rewrite HM12; split; [apply invertible_one | easy].
+Qed.
+
+Lemma aff_span_incl :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2),
+    inclF A1 (aff_span A2) -> incl (aff_span A1) (aff_span A2).
+Proof.
+move=>> HA1 _ [L1 HL1]; apply aff_spanF_ex1 in HA1; destruct HA1 as [M HM].
+apply all_and_equiv in HM; destruct HM as [HM HA1].
+apply aff_span_ex; exists (fun i2 => comb_lin L1 (M^~ i2)); split.
+2: apply barycenter2_r; easy.
+replace (sum _) with (sum L1); try easy.
+rewrite -comb_lin_sum_r -(comb_lin_ext_r _ ones).
+rewrite comb_lin_ones_r; easy.
+intros i1; rewrite (HM i1); easy.
+Qed.
+
+Lemma aff_span_monot :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2),
+    invalF A1 A2 -> incl (aff_span A1) (aff_span A2).
+Proof. intros; apply aff_span_incl, aff_span_inclF; easy. Qed.
+
+Lemma aff_span_ext :
+  forall {n1 n2} (A1 : 'E^n1) (A2 : 'E^n2),
+    inclF A1 (aff_span A2) -> inclF A2 (aff_span A1) ->
+    aff_span A1 = aff_span A2.
+Proof. intros; apply subset_ext_equiv; split; apply aff_span_incl; easy. Qed.
+
+Lemma aff_span_inv_frameF_orig :
+  forall {n} (B : 'V^n) (O : E) i0, aff_span (inv_frameF O B i0) O.
+Proof.
+intros n B O i0; rewrite -{2}(transl_zero O); apply (aff_span_ub _ _ i0).
+unfold inv_frameF; rewrite translF_insertF insertF_correct_l; easy.
+Qed.
+
+Lemma is_aff_basis_is_aff_gen :
+  forall (PE : E -> Prop) {n} (A : 'E^n.+1),
+    is_aff_basis PE A -> is_aff_gen PE A.
+Proof. move=>> [HA _]; easy. Qed.
+
+Lemma is_aff_basis_is_aff_indep :
+  forall (PE : E -> Prop) {n} (A : 'E^n.+1),
+    is_aff_basis PE A -> is_aff_indep A.
+Proof. move=>> [_ HA]; easy. Qed.
+
+Lemma is_aff_gen_nonempty :
+  forall {PE : E -> Prop} {n} {A : 'E^n.+1}, is_aff_gen PE A -> exists O, PE O.
+Proof. move=> PE n A ->; exists (A ord0); apply aff_span_inclF_diag. Qed.
+
+Lemma is_aff_basis_nonempty :
+  forall {PE : E -> Prop} {n} {A : 'E^n.+1},
+    is_aff_basis PE A -> exists O, PE O.
+Proof. move=>> [HA _]; apply (is_aff_gen_nonempty HA). Qed.
+
+Lemma has_aff_dim_is_nonempty :
+  forall {PE : E -> Prop} {n}, has_aff_dim PE n -> exists O, PE O.
+Proof. move=>> [A HA]; apply (is_aff_basis_nonempty HA). Qed.
+
+End Aff_span_Subset_Facts.
+
+
+Section Aff_span_Affine_Facts.
+
+Context {K : Ring}.
+Context {V : ModuleSpace K}.
+Context {E : AffineSpace V}.
+
+Lemma aff_span_exs :
+  forall {n1 n2} {A2 : 'E^n2} {G1 : 'E^n1},
+    inclF G1 (aff_span A2) ->
+    exists L, (forall i, sum (L i) = 1) /\ G1 = fun i => barycenter (L i) A2.
+Proof.
+move=> n1 n2 A2 G1 HG1.
+destruct (choice (fun i L => sum L = 1 /\ G1 i = barycenter L A2)) as [L HL].
+  intros i; destruct (aff_span_EX _ _ (HG1 i)) as [L HL]; exists L; easy.
+apply all_and_equiv in HL; exists L; split; try apply extF; easy.
+Qed.
+
+Lemma is_aff_indep_alt :
+  forall {n} {A : 'E^n.+1} i0, is_free (frameF A i0) -> is_aff_indep A.
+Proof.
+intros n A i0 HA i; destruct (ord_eq_dec i i0) as [Hi | Hi]; try now rewrite Hi.
+destruct n as [|n]. contradict Hi; rewrite !ord1; easy.
+intros L; unfold frameF; rewrite (vectF_change_orig (A i0))
+    comb_lin_plus_r -scal_sum_l scal_opp -vect_sym vectF_skipF.
+rewrite -comb_lin_insertF_l (comb_lin_skip_zero_r _ _ i0 (vectF_zero_alt _ _)).
+move=> /HA /extF_rev HL1.
+assert (HL1a : skipF L (insert_ord (not_eq_sym Hi)) = 0).
+  apply skipF_zero_compat; move=> j /skip_insert_ord_neq Hj.
+  specialize (HL1 (insert_ord Hj)).
+  rewrite skipF_correct insertF_correct in HL1; easy.
+apply (extF_zero_skipF _ (insert_ord (not_eq_sym Hi))); try easy.
+rewrite -sum_one // opp_zero_equiv; move: (HL1 (insert_ord Hi)).
+rewrite zeroF skipF_correct insertF_correct_l; easy.
+Qed.
+
+Lemma is_aff_dep_alt :
+  forall {n} {A : 'E^n.+1} i0, is_lin_dep (frameF A i0) -> is_aff_dep A.
+Proof. move=>>; rewrite -contra_equiv; easy. Qed.
+
+End Aff_span_Affine_Facts.
diff --git a/FEM/Algebra/Ring_compl.v b/FEM/Algebra/Ring_compl.v
index 83113a0534a9d3637d544a4a979191d0ec687282..5199b9df9dae4851719638635a7acb8b4960d0e8 100644
--- a/FEM/Algebra/Ring_compl.v
+++ b/FEM/Algebra/Ring_compl.v
@@ -96,7 +96,7 @@ Definition inv (x : K) : K :=
   | right _ => 0
   end.
 
-Definition div x y : K := x * inv y.
+Definition div (x y : K) : K := x * inv y.
 
 Lemma inv_correct : forall {x y}, is_inverse x y -> inv x = y.
 Proof.
@@ -109,7 +109,7 @@ Qed.
 Lemma inv_correct_rev :
   forall {x y}, invertible x -> y = inv x -> is_inverse x y.
 Proof.
-move=>> [y' Hy'] Hy; rewrite Hy; apply (is_inverse_compat_r  y'); try easy.
+move=>> [y' Hy'] Hy; rewrite Hy; apply (is_inverse_compat_r y'); try easy.
 apply eq_sym, inv_correct; easy.
 Qed.
 
diff --git a/FEM/Algebra/Ring_compl_new.v b/FEM/Algebra/Ring_compl_new.v
new file mode 100644
index 0000000000000000000000000000000000000000..1b624ea24d78208456eca91ae1aa83c265ef7e55
--- /dev/null
+++ b/FEM/Algebra/Ring_compl_new.v
@@ -0,0 +1,1512 @@
+(**
+This file is part of the Elfic library
+
+Copyright (C) Boldo, Clément, Martin, Mayero, Mouhcine
+
+This library is free software; you can redistribute it and/or
+modify it under the terms of the GNU Lesser General Public
+License as published by the Free Software Foundation; either
+version 3 of the License, or (at your option) any later version.
+
+This library is distributed in the hope that it will be useful,
+but WITHOUT ANY WARRANTY; without even the implied warranty of
+MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the
+COPYING file for more details.
+*)
+
+From FEM.Compl Require Import stdlib.
+
+Set Warnings "-notation-overridden".
+From FEM.Compl Require Import ssr.
+
+(* Next import provides Reals, Coquelicot.Hierarchy, and functions to rings. *)
+From LM Require Import linear_map.
+Close Scope R_scope.
+Set Warnings "notation-overridden".
+
+From Lebesgue Require Import nat_compl Subset Subset_any Function.
+
+From FEM Require Import Compl Sub_struct Monoid_compl Group_compl.
+
+
+(** Results needing a commutative Ring are only stated for
+ the ring of real numbers R_Ring. *)
+
+
+Section Ring_Def0.
+
+Context {K : Ring}.
+
+Definition two : K := plus one one.
+
+End Ring_Def0.
+
+
+Declare Scope Ring_scope.
+Delimit Scope Ring_scope with K.
+Notation "x * y" := (mult x y) : Ring_scope.
+Notation "1" := (one) : Ring_scope.
+Notation "2" := (two) : Ring_scope.
+
+Local Open Scope Monoid_scope.
+Local Open Scope Group_scope.
+Local Open Scope Ring_scope.
+
+
+Section Ring_Def1.
+
+Context {K : Ring}.
+
+Definition is_inverse (x y : K) : Prop := y * x = 1 /\ x * y = 1.
+
+Lemma is_inverse_compat_l :
+  forall x1 x2 y, x1 = x2 -> is_inverse x1 y -> is_inverse x2 y.
+Proof. intros; subst; easy. Qed.
+
+Lemma is_inverse_compat_r :
+  forall y1 y2 x, y1 = y2 -> is_inverse x y1 -> is_inverse x y2.
+Proof. intros; subst; easy. Qed.
+
+Lemma is_inverse_uniq :
+  forall x y1 y2, is_inverse x y1 -> is_inverse x y2 -> y1 = y2.
+Proof.
+intros x y1 y2 [Hy1 _] [_ Hy2].
+rewrite -(mult_one_r y1) -Hy2 mult_assoc Hy1 mult_one_l; easy.
+Qed.
+
+(* FIXME: essai de définition avec un existentiel fort. *)
+Definition invertible (x : K) : Type := { y | is_inverse x y }.
+
+(*
+Inductive invertible (x : K) : Prop :=
+  | Invertible : forall y, is_inverse x y -> invertible x.
+
+Lemma invertible_ext :
+  forall {x} (y : K), x = y -> invertible x <-> invertible y.
+Proof. intros; subst; easy. Qed.
+
+Lemma invertible_ex : forall {x}, (exists y, is_inverse x y) -> invertible x.
+Proof. intros x [y Hy]; apply (Invertible _ _ Hy). Qed.
+
+Lemma invertible_EX : forall {x}, invertible x -> { y | is_inverse x y }.
+Proof. move=>> Hx; apply ex_EX; destruct Hx as [y Hy]; exists y; easy. Defined.
+ *)
+
+(* FIXME: pour avoir des fonctions totales inv et div, il faut invertible_dec.
+ Pas possible sans ex_EX ? *)
+Lemma invertible_dec : forall x, invertible x + { forall y, ~ is_inverse x y }.
+Proof.
+intros x; destruct (classic_dec (exists y, is_inverse x y)) as [Hx | Hx].
+left; apply ex_EX; easy.
+right; apply not_ex_all_not; easy.
+Qed.
+
+Definition inv (x : K) : K :=
+  match invertible_dec x with
+  | inleft Hx => proj1_sig Hx
+  | inright _ => 0
+  end.
+
+Definition div (x y : K) : K := x * inv y.
+
+Lemma inv_correct : forall {x y}, is_inverse x y -> inv x = y.
+Proof.
+intros x y H; apply (is_inverse_uniq x); try easy.
+unfold inv; destruct (invertible_dec _) as [H' | H'].
+apply (proj2_sig H').
+contradict H'; apply not_all_not_ex_equiv; exists y; easy. 
+Qed.
+
+Lemma inv_correct_rev :
+  forall {x y}, invertible x -> y = inv x -> is_inverse x y.
+Proof.
+move=>> [y' Hy'] Hy; rewrite Hy; apply (is_inverse_compat_r y'); try easy.
+apply eq_sym, inv_correct; easy.
+Qed.
+
+End Ring_Def1.
+
+
+Notation "/ x" := (inv x) : Ring_scope.
+Notation "x / y" := (div x y) : Ring_scope.
+
+
+Section Ring_Def2.
+
+Variable K : Ring.
+
+Definition is_comm : Prop := forall x y : K, x * y = y * x.
+
+(* FIXME: invertible est un Type et non un Prop ! *)
+Definition has_inv : Prop := forall x : K, x <> 0 -> invertible x.
+
+Definition is_field : Prop := nonzero_struct K /\ has_inv.
+
+Definition is_comm_field : Prop := is_comm /\ is_field.
+
+End Ring_Def2.
+
+
+Section Ring_Facts.
+
+Context {K : Ring}.
+
+Lemma inhabited_r : inhabited K.
+Proof. apply inhabited_g. Qed.
+
+Lemma inhabited_fct_r : forall {T : Type}, inhabited (T -> K).
+Proof. intros; apply inhabited_fct_g. Qed.
+
+Lemma one_is_zero : (1 : K) = 0 -> zero_struct K.
+Proof. intros H x; rewrite -(mult_one_l x) H mult_zero_l; easy. Qed.
+
+Lemma one_not_zero : nonzero_struct K -> (1 : K) <> 0.
+Proof.
+intros [x Hx]; generalize Hx; rewrite -contra_equiv.
+intros; apply one_is_zero; easy.
+Qed.
+
+Lemma one_is_zero_equiv : (1 : K) = 0 <-> zero_struct K.
+Proof. split; [apply one_is_zero | easy]. Qed.
+
+Lemma one_not_zero_equiv : (1 : K) <> 0 <-> nonzero_struct K.
+Proof. split; [intros; exists 1; easy | apply one_not_zero]. Qed.
+
+Lemma minus_one_not_zero : nonzero_struct K -> - (1 : K) <> 0.
+Proof.
+intros; rewrite -opp_zero; apply opp_neq_compat, one_not_zero; easy.
+Qed.
+
+Lemma mult_compat_l : forall x x1 x2 : K, x1 = x2 -> x * x1 = x * x2.
+Proof. move=>> H; rewrite H; easy. Qed.
+
+Lemma mult_compat_r : forall x x1 x2 : K, x1 = x2 -> x1 * x = x2 * x.
+Proof. move=>> H; rewrite H; easy. Qed.
+
+Lemma mult_zero_rev_l :
+  forall x1 x2 : K, x1 * x2 = 0 -> ~ invertible x1 \/ x2 = 0.
+Proof.
+intros x1 x2 H.
+destruct (classic (invertible x1)) as [[y1 [H1 _]] | H1]; [right | now left].
+rewrite -(mult_one_l x2) -H1 -mult_assoc H mult_zero_r; easy.
+Qed.
+
+Lemma mult_zero_rev_r :
+  forall x1 x2 : K, x1 * x2 = 0 -> x1 = 0 \/ ~ invertible x2.
+Proof.
+intros x1 x2 H.
+destruct (classic (invertible x2)) as [[y2 [_ H2]] | H2]; [left | now right].
+rewrite -(mult_one_r x1) -H2 mult_assoc H mult_zero_l; easy.
+Qed.
+
+Lemma mult_not_zero_l :
+  forall x1 x2 : K, invertible x1 -> x2 <> 0 -> x1 * x2 <> 0.
+Proof. move=>> H1 H2 H; destruct (mult_zero_rev_l _ _ H); easy. Qed.
+
+Lemma mult_not_zero_r :
+  forall x1 x2 : K, x1 <> 0 -> invertible x2 -> x1 * x2 <> 0.
+Proof. move=>> H1 H2 H; destruct (mult_zero_rev_r _ _ H); easy. Qed.
+
+Lemma mult_reg_l :
+  forall x x1 x2 : K, invertible x -> x * x1 = x * x2 -> x1 = x2.
+Proof.
+move=>> [y [Hy _]] /(mult_compat_l y) Hx.
+rewrite 2!mult_assoc Hy 2!mult_one_l in Hx; easy.
+Qed.
+
+Lemma mult_reg_r :
+  forall x x1 x2 : K, invertible x -> x1 * x = x2 * x -> x1 = x2.
+Proof.
+move=>> [y [_ Hy]] /(mult_compat_r y) Hx.
+rewrite -2!mult_assoc Hy 2!mult_one_r in Hx; easy.
+Qed.
+
+Lemma mult_minus_distr_l : forall x y1 y2 : K, x * (y1 - y2) = x * y1 - x * y2.
+Proof. intros; rewrite !minus_eq opp_mult_r; apply mult_distr_l. Qed.
+
+Lemma mult_minus_distr_r : forall x1 x2 y : K, (x1 - x2) * y = x1 * y - x2 * y.
+Proof. intros; rewrite !minus_eq opp_mult_l; apply mult_distr_r. Qed.
+
+Lemma is_inverse_sym : forall x y : K, is_inverse x y -> is_inverse y x.
+Proof. move=>> [H1 H2]; easy. Qed.
+
+Lemma is_inverse_sym_equiv : forall x y : K, is_inverse x y <-> is_inverse y x.
+Proof. intros; split; apply is_inverse_sym. Qed.
+
+Lemma is_inverse_zero_struct : forall x y : K, zero_struct K -> is_inverse x y.
+Proof.
+intros x y HK; split; rewrite (HK x) (HK y) (HK 1); apply mult_zero_l.
+Qed.
+
+Lemma invertible_zero_struct : forall x : K, zero_struct K -> invertible x.
+Proof. move=> x /is_inverse_zero_struct Hx; apply (Invertible x 0); easy. Qed.
+
+Lemma invertible_zero : invertible (0 : K) -> zero_struct K.
+Proof.
+intros [z [_ H]]; rewrite mult_zero_l in H; apply one_is_zero; easy.
+Qed.
+
+Lemma is_inverse_zero : forall x : K, is_inverse 0 x -> zero_struct K.
+Proof. move=>> Hx; apply invertible_zero; apply (Invertible _ _ Hx). Qed.
+
+Lemma invertible_one : invertible (1 : K).
+Proof. apply (Invertible _ 1); split; apply mult_one_l. Qed.
+
+Lemma invertible_eq_one : forall {a : K}, a = 1 -> invertible a.
+Proof. intros; subst; apply invertible_one. Qed.
+
+Lemma noninvertible_zero : nonzero_struct K -> ~ invertible (0 : K).
+Proof. intros [x0 Hx0]; contradict Hx0; apply invertible_zero; easy. Qed.
+
+Lemma is_inverse_nonzero_l :
+  forall x y : K, nonzero_struct K -> is_inverse x y -> x <> 0.
+Proof.
+move=>> HK; apply contra_not_r_equiv.
+intros; subst; move=> [_ H]; contradict H.
+rewrite mult_zero_l; apply not_eq_sym, one_not_zero; easy.
+Qed.
+
+Lemma is_inverse_nonzero_r :
+  forall x y : K, nonzero_struct K -> is_inverse x y -> y <> 0.
+Proof. move=>> HK /is_inverse_sym; apply is_inverse_nonzero_l; easy. Qed.
+
+Lemma is_inverse_invertible_l : forall x y : K, is_inverse x y -> invertible x.
+Proof. intros x y H; apply (Invertible x y H). Qed.
+
+Lemma is_inverse_invertible_r : forall x y : K, is_inverse x y -> invertible y.
+Proof. move=>> /is_inverse_sym; apply is_inverse_invertible_l. Qed.
+
+Lemma inv_0 : / (0 : K) = 0.
+Proof.
+unfold inv; destruct (invertible_dec _); try easy.
+apply invertible_zero; easy.
+Qed.
+
+Lemma mult_invertible_compat :
+  forall x y : K, invertible x -> invertible y -> invertible (x * y).
+Proof.
+move=>> [x1 [Hxa Hxb]] [y1 [Hya Hyb]]; apply (Invertible _ (y1 * x1)); split.
+rewrite -mult_assoc (mult_assoc x1) Hxa mult_one_l; easy.
+rewrite -mult_assoc (mult_assoc _ y1) Hyb mult_one_l; easy.
+Qed.
+
+Lemma mult_invertible_reg_l :
+  forall x y : K, invertible y -> invertible (x * y) ->  invertible x.
+Proof.
+intros x y Hy [z [Hza Hzb]]; apply (Invertible _ (y * z)); split.
+(* *)
+generalize Hy; intros [y1 Hy1].
+apply (mult_reg_l y1); try now apply (is_inverse_invertible_r y).
+apply (mult_reg_r y); try easy.
+destruct Hy1 as [Hy1 _].
+rewrite 2!mult_assoc Hy1 mult_one_l mult_one_r -mult_assoc Hy1; easy.
+(* *)
+rewrite mult_assoc; easy.
+Qed.
+
+Lemma mult_invertible_reg_r :
+  forall x y : K, invertible x -> invertible (x * y) ->  invertible y.
+Proof.
+intros x y Hx [z [Hza Hzb]]; apply (Invertible _ (z * x)); split.
+(* *)
+rewrite -mult_assoc; easy.
+(* *)
+generalize Hx; intros [x1 Hx1].
+apply (mult_reg_l x); try easy.
+apply (mult_reg_r x1); try now apply (is_inverse_invertible_r x).
+destruct Hx1 as [_ Hx1].
+rewrite mult_assoc -mult_assoc -(mult_assoc z) Hx1 !mult_one_r Hx1; easy.
+Qed.
+
+Lemma mult_invertible_equiv_l :
+  forall x y : K, invertible y -> invertible (x * y) <-> invertible x.
+Proof.
+intros; split.
+apply mult_invertible_reg_l; easy.
+intros; apply mult_invertible_compat; easy.
+Qed.
+
+Lemma mult_invertible_equiv_r :
+  forall x y : K, invertible x -> invertible (x * y) <-> invertible y.
+Proof.
+intros; split.
+apply mult_invertible_reg_r; easy.
+intros; apply mult_invertible_compat; easy.
+Qed.
+
+Lemma mult_inv_l : forall {x : K}, invertible x -> / x * x = 1.
+Proof. intros x [y Hy]; rewrite (inv_correct Hy); apply Hy. Qed.
+
+Lemma mult_inv_r : forall {x : K}, invertible x -> x * / x = 1.
+Proof. intros x [y Hy]; rewrite (inv_correct Hy); apply Hy. Qed.
+
+Lemma is_inverse_inv_r : forall {x : K}, invertible x -> is_inverse x (/ x).
+Proof. intros x Hx; apply (inv_correct_rev Hx); easy. Qed.
+
+Lemma is_inverse_inv_l : forall {x : K}, invertible x -> is_inverse (/ x) x.
+Proof. intros; apply is_inverse_sym, is_inverse_inv_r; easy. Qed.
+
+Lemma inv_invertible : forall {x : K}, invertible x -> invertible (/ x).
+Proof. intros x Hx; apply (Invertible _ x), (is_inverse_inv_l Hx). Qed.
+
+Lemma inv_invol : forall {x : K}, invertible x -> / (/ x) = x.
+Proof. intros; apply inv_correct, is_inverse_inv_l; easy. Qed.
+
+Lemma div_eq_one : forall {x : K}, invertible x -> x / x = 1.
+Proof. move=>>; apply mult_inv_r. Qed.
+
+Lemma div_one_compat : forall {x y : K}, invertible y -> x = y -> x / y = 1.
+Proof. move=>> Hy ->; apply div_eq_one; easy. Qed.
+
+Lemma div_one_reg : forall {x y : K}, invertible y -> x / y = 1 -> x = y.
+Proof.
+move=> x y Hy /(mult_compat_r y); rewrite -mult_assoc (mult_inv_l Hy).
+rewrite mult_one_l mult_one_r; easy.
+Qed.
+
+Lemma div_one_equiv : forall {x y : K}, invertible y -> x / y = 1 <-> x = y.
+Proof. intros; split; [apply div_one_reg | apply div_one_compat]; easy. Qed.
+
+Lemma mult_sum_distr_l :
+  forall {n} x (u : 'K^n), x * sum u = sum (fun i => x * u i).
+Proof.
+intros n; induction n as [|n Hn]; intros.
+rewrite !sum_nil; apply mult_zero_r.
+rewrite !sum_ind_l mult_distr_l Hn; easy.
+Qed.
+
+Lemma mult_sum_distr_r :
+  forall {n} x (u : 'K^n), sum u * x = sum (fun i => u i * x).
+Proof.
+intros n; induction n as [|n Hn]; intros.
+rewrite !sum_nil; apply mult_zero_l.
+rewrite !sum_ind_l mult_distr_r Hn; easy.
+Qed.
+
+Lemma sum_nil_noninvertible :
+  forall (u : 'K^0), nonzero_struct K -> ~ invertible (sum u).
+Proof. move=>>; rewrite sum_nil; apply noninvertible_zero. Qed.
+
+Lemma invertible_sum_nonnil :
+  forall {n} (u : 'K^n),
+    nonzero_struct K ->  invertible (sum u) -> n <> 0%nat.
+Proof.
+move=>> HK; apply contra_not_r_equiv; intros; subst.
+apply sum_nil_noninvertible; easy.
+Qed.
+
+Lemma sum_singleF_invertible :
+  forall {a : K}, invertible a -> invertible (sum (singleF a)).
+Proof. move=>>; rewrite sum_singleF; easy. Qed.
+
+Lemma invertible_sum_equiv :
+  forall {n} (u : 'K^n),
+    invertible (sum u) <-> invertible (sum (filter_n0F u)).
+Proof. intros; rewrite sum_filter_n0F; easy. Qed.
+
+Lemma sum_insertF_baryc :
+  forall {n} (u : 'K^n) i0, sum (insertF u (1 - sum u) i0) = 1.
+Proof. intros; rewrite sum_insertF plus_minus_l; easy. Qed.
+
+Lemma invertible_sum_squeezeF :
+  forall {n} (u : 'K^n.+1) {i0 i1},
+    i1 <> i0 -> invertible (sum u) -> invertible (sum (squeezeF u i0 i1)).
+Proof. intros; rewrite sum_squeezeF; easy. Qed.
+
+(** Functions to ring. *)
+
+Context {U : Type}.
+
+Lemma fct_one_eq : forall x, (1 : U -> K) x = 1.
+Proof. easy. Qed.
+
+Lemma fct_mult_eq : forall (f g : U -> K) x, (f * g) x = f x * g x.
+Proof. easy. Qed.
+
+End Ring_Facts.
+
+
+Section Ring_Morphism_Def.
+
+Context {K1 K2 : Ring}.
+
+Definition f_mult_compat (f : K1 -> K2) : Prop :=
+  forall x1 y1, f (x1 * y1) = f x1 * f y1.
+
+Definition f_one_compat (f : K1 -> K2) : Prop := f 1 = 1.
+
+Definition morphism_r f := morphism_g f /\ f_mult_compat f /\ f_one_compat f.
+
+End Ring_Morphism_Def.
+
+
+Section Ring_Morphism_Facts1.
+
+Context {K1 K2 : Ring}.
+
+Lemma morphism_r_g : forall f : K1 -> K2, morphism_r f -> morphism_g f.
+Proof. move=>> H; apply H. Qed.
+
+Lemma morphism_r_mult : forall f : K1 -> K2, morphism_r f -> f_mult_compat f.
+Proof. move=>> H; apply H. Qed.
+
+Lemma morphism_r_one : forall f : K1 -> K2, morphism_r f -> f_one_compat f.
+Proof. move=>> H; apply H. Qed.
+
+Lemma morphism_r_m : forall f : K1 -> K2, morphism_r f -> morphism_m f.
+Proof. move=>> /morphism_r_g; apply: morphism_g_m. Qed.
+
+Lemma morphism_r_plus : forall f : K1 -> K2, morphism_r f -> f_plus_compat f.
+Proof. move=>> /morphism_r_g; apply: morphism_g_plus. Qed.
+
+Lemma morphism_r_zero : forall f : K1 -> K2, morphism_r f -> f_zero_compat f.
+Proof. move=>> /morphism_r_g; apply: morphism_g_zero. Qed.
+
+Lemma morphism_r_sum : forall f : K1 -> K2, morphism_r f -> f_sum_compat f.
+Proof. move=>> /morphism_r_g; apply: morphism_g_sum. Qed.
+
+Lemma morphism_r_opp : forall f : K1 -> K2, morphism_r f -> f_opp_compat f.
+Proof. move=>> /morphism_r_g; apply morphism_g_opp. Qed.
+
+Lemma morphism_r_minus : forall f : K1 -> K2, morphism_r f -> f_minus_compat f.
+Proof. move=>> /morphism_r_g; apply morphism_g_minus. Qed.
+
+(* Apparently, there is no structure on morphism_r! *)
+
+Lemma morphism_r_ext :
+  forall f g : K1 -> K2, same_fun f g -> morphism_r f -> morphism_r g.
+Proof. intros f g H; replace g with f; try apply fun_ext; easy. Qed.
+
+End Ring_Morphism_Facts1.
+
+
+Section FF_Ring_Facts1.
+
+(** Properties with the second identity element of rings (one). *)
+
+Context {K : Ring}.
+
+Definition ones {n} : 'K^n := constF n 1.
+
+Definition alt_ones {n} : 'K^n :=
+  fun i => if Nat.Even_Odd_dec i then 1 else - (1 : K).
+
+Definition alt_ones_shift {n} : 'K^n :=
+  fun i => if Nat.Even_Odd_dec i then - (1 : K) else 1.
+
+Lemma ones_correct : forall {n} (i : 'I_n), ones i = 1.
+Proof. easy. Qed.
+
+Lemma alt_ones_correct_even :
+  forall {n} (i : 'I_n), Nat.Even i -> alt_ones i = 1.
+Proof.
+intros n i Hi; unfold alt_ones; destruct (Nat.Even_Odd_dec i); simpl; try easy.
+exfalso; apply (Nat.Even_Odd_False i); easy.
+Qed.
+
+Lemma alt_ones_correct_odd :
+  forall {n} (i : 'I_n), Nat.Odd i -> alt_ones i = - (1 : K).
+Proof.
+intros n i Hi; unfold alt_ones; destruct (Nat.Even_Odd_dec i); simpl; try easy.
+exfalso; apply (Nat.Even_Odd_False i); easy.
+Qed.
+
+Lemma alt_ones_shift_correct_even :
+  forall {n} (i : 'I_n), Nat.Even i -> alt_ones_shift i = - (1 : K).
+Proof.
+intros n i Hi; unfold alt_ones_shift;
+    destruct (Nat.Even_Odd_dec i); simpl; try easy.
+exfalso; apply (Nat.Even_Odd_False i); easy.
+Qed.
+
+Lemma alt_ones_shift_correct_odd :
+  forall {n} (i : 'I_n), Nat.Odd i -> alt_ones_shift i = 1.
+Proof.
+intros n i Hi; unfold alt_ones_shift;
+    destruct (Nat.Even_Odd_dec i); simpl; try easy.
+exfalso; apply (Nat.Even_Odd_False i); easy.
+Qed.
+
+Lemma alt_ones_eq : forall {n}, @alt_ones n = liftF_S (@alt_ones_shift n.+1).
+Proof.
+intros n; apply extF; intros i; unfold liftF_S; destruct (Nat.Even_Odd_dec i).
+(* *)
+rewrite -> alt_ones_shift_correct_odd, alt_ones_correct_even; try easy.
+rewrite lift_S_correct; apply Nat.Odd_succ; easy.
+(* *)
+rewrite -> alt_ones_shift_correct_even, alt_ones_correct_odd; try easy.
+rewrite lift_S_correct; apply Nat.Even_succ; easy.
+Qed.
+
+Lemma alt_ones_shift_eq :
+  forall {n}, @alt_ones_shift n = liftF_S (@alt_ones n.+1).
+Proof.
+intros n; apply extF; intros i; unfold liftF_S; destruct (Nat.Even_Odd_dec i).
+(* *)
+rewrite -> alt_ones_shift_correct_even, alt_ones_correct_odd; try easy.
+rewrite lift_S_correct; apply Nat.Odd_succ; easy.
+(* *)
+rewrite -> alt_ones_shift_correct_odd, alt_ones_correct_even; try easy.
+rewrite lift_S_correct; apply Nat.Even_succ; easy.
+Qed.
+
+Lemma alt_ones_shift2 :
+  forall {n}, @alt_ones n = liftF_S (liftF_S (@alt_ones n.+2)).
+Proof. intros; rewrite alt_ones_eq alt_ones_shift_eq; easy. Qed.
+
+Lemma ones_not_zero : forall {n}, nonzero_struct K -> (ones : 'K^n.+1) <> 0.
+Proof. intros; apply nextF; exists ord0; apply one_not_zero; easy. Qed.
+
+Lemma alt_ones_not_zero :
+  forall {n}, nonzero_struct K -> (alt_ones : 'K^n.+1) <> 0.
+Proof.
+intros; apply nextF; exists ord0; rewrite alt_ones_correct_even.
+apply one_not_zero; easy.
+apply Even_0.
+Qed.
+
+Lemma sum_alt_ones_even : forall {n}, Nat.Even n -> sum (@alt_ones n) = 0.
+Proof.
+intros n [m Hm]; subst; induction m as [| m Hm]. apply sum_nil.
+rewrite -(sum_castF (nat_double_S _)).
+rewrite 2!sum_ind_l; unfold castF at 1, liftF_S at 1, castF at 1.
+rewrite {1}alt_ones_correct_even; try apply Even_0;
+    rewrite alt_ones_correct_odd; try apply Odd_1.
+rewrite plus_assoc plus_opp_r plus_zero_l.
+rewrite -(sum_eq (@alt_ones (2 * m)%coq_nat)); try easy.
+apply alt_ones_shift2.
+Qed.
+
+Lemma sum_alt_ones_odd : forall {n}, Nat.Odd n -> sum (@alt_ones n) = 1.
+Proof.
+intros n [m Hm]; subst; induction m as [| m Hm].
+rewrite sum_1; apply alt_ones_correct_even, Even_0.
+rewrite -(sum_castF (nat_S_double_S _)).
+rewrite 2!sum_ind_l; unfold castF at 1, liftF_S at 1, castF at 1.
+rewrite {1}alt_ones_correct_even; try apply Even_0;
+    rewrite alt_ones_correct_odd; try apply Odd_1.
+rewrite plus_assoc plus_opp_r plus_zero_l.
+rewrite -(sum_eq (@alt_ones ((2 * m)%coq_nat.+1))).
+rewrite -(sum_castF (addn1_sym _)); easy.
+apply alt_ones_shift2.
+Qed.
+
+Lemma alt_ones_3_eq : @alt_ones 3 = tripleF 1 (- (1)) 1.
+Proof.
+apply extF; intros i; destruct (ord3_dec i) as [[Hi | Hi] | Hi]; rewrite Hi.
+rewrite -> alt_ones_correct_even, tripleF_0; try apply Even_0; easy.
+rewrite -> alt_ones_correct_odd, tripleF_1; try apply Odd_1; easy.
+rewrite -> alt_ones_correct_even, tripleF_2; try apply Even_2; easy.
+Qed.
+
+Lemma sum_alt_ones_3 : sum (tripleF 1 (- (1 : K)) 1) = 1.
+Proof. rewrite -alt_ones_3_eq; apply sum_alt_ones_odd, Odd_3. Qed.
+
+Lemma sum_alt_ones_3_invertible : invertible (sum (tripleF 1 (- (1 : K)) 1)).
+Proof. rewrite sum_alt_ones_3; apply invertible_one. Qed.
+
+Definition unit_partition {n} (u : 'K^n) : 'K^n.+1 :=
+  insertF u (1 - sum u) ord0.
+
+Lemma unit_partition_correct_0 :
+  forall {n} (u : 'K^n) i, i = ord0 -> unit_partition u i = 1 - sum u.
+Proof. move=>> ->; unfold unit_partition; rewrite insertF_correct_l; easy. Qed.
+
+Lemma unit_partition_correct_S :
+  forall {n} (u : 'K^n) {i} (Hi : i <> ord0),
+    unit_partition u i = u (lower_S Hi).
+Proof.
+intros n u i Hi; unfold unit_partition; rewrite insertF_correct_rr.
+apply ord_n0_nlt_equiv in Hi; destruct i; simpl in *; auto with zarith.
+intros Hi'; f_equal; apply ord_inj; easy.
+Qed.
+
+Lemma sum_unit_partition : forall {n} (u : 'K^n), sum (unit_partition u) = 1.
+Proof.
+intros; unfold unit_partition; rewrite sum_insertF plus_minus_l; easy.
+Qed.
+
+Lemma unit_partition_1_eq :
+  forall a, unit_partition (singleF a) = coupleF (1 - a) a.
+Proof.
+intros; apply extF; intros i; destruct (ord2_dec i) as [Hi | Hi]; rewrite Hi.
+rewrite -> unit_partition_correct_0, sum_singleF, coupleF_0; easy.
+rewrite unit_partition_correct_S; try easy.
+intros Hi'; rewrite singleF_0 coupleF_1; easy.
+Qed.
+
+Lemma sum_unit_partition_1 :
+  forall (a : K), sum (coupleF (1 - a) a) = 1.
+Proof. intros; rewrite -unit_partition_1_eq; apply sum_unit_partition. Qed.
+
+Lemma sum_unit_partition_1_invertible :
+  forall (a : K), invertible (sum (coupleF (1 - a) a)).
+Proof. intros; rewrite sum_unit_partition_1; apply invertible_one. Qed.
+
+End FF_Ring_Facts1.
+
+
+Section FF_Ring_Facts2.
+
+Definition plusn {K : Ring} (n : nat) (x : K) : K := sum (constF n x).
+Definition plusn1 (K : Ring) (n : nat) : K := plusn n 1.
+
+Context {K : Ring}.
+
+Lemma plusn1_eq : forall n, plusn1 K n = plusn n 1.
+Proof. easy. Qed.
+
+Lemma plusn_0_l : forall x : K, plusn 0 x = 0.
+Proof. intros; apply sum_nil. Qed.
+
+Lemma plusn_0_r : forall n, plusn n (0 : K) = 0.
+Proof. intros; apply sum_zero. Qed.
+
+Lemma plusn_1_l : forall x : K, plusn 1 x = x.
+Proof. intros; apply sum_1. Qed.
+
+Lemma plusn_2_l : forall x : K, plusn 2 x = x + x.
+Proof. intros; apply sum_2. Qed.
+
+Lemma plusn_2_1 : plusn 2 (1 : K) = 2.
+Proof. rewrite plusn_2_l; easy. Qed.
+
+Lemma plusn1_2 : plusn1 K 2 = 2.
+Proof. rewrite plusn1_eq plusn_2_1; easy. Qed.
+
+Lemma plusn_3_l : forall x : K, plusn 3 x = x + x + x.
+Proof. intros; apply sum_3. Qed.
+
+Lemma plusn_ind : forall n (x : K), plusn n.+1 x = x + plusn n x.
+Proof. intros; apply sum_ind_l. Qed.
+
+Lemma plusn_distr_l :
+  forall n1 n2 (x : K), plusn (n1 + n2) x = plusn n1 x + plusn n2 x.
+Proof. intros; unfold plusn; rewrite -concatF_constF sum_concatF; easy. Qed.
+
+Lemma plusn_distr_r :
+  forall n (x y : K), plusn n (x + y) = plusn n x + plusn n y.
+Proof.
+intros n x y; induction n as [| n Hn].
+rewrite !plusn_0_l plus_zero_l; easy.
+rewrite !plusn_ind Hn -(plus_assoc _ (plusn _ _)) (plus_assoc (plusn _ _))
+    (plus_comm (plusn _ _) y) -(plus_assoc y) -(plus_assoc x y); easy.
+Qed.
+
+Lemma plusn_assoc_l :
+  forall n1 n2 (x : K), plusn (n1 * n2)%coq_nat x = plusn n1 (plusn n2 x).
+Proof.
+intros n1 n2 x; induction n1 as [| n1 Hn1].
+rewrite Nat.mul_0_l !plusn_0_l; easy.
+replace (n1.+1 * n2)%coq_nat with ((n1 * n2)%coq_nat + n2)%coq_nat by lia.
+rewrite plusn_distr_l Hn1 plusn_ind plus_comm; easy.
+Qed.
+
+Lemma plusn_assoc_r :
+  forall n1 n2 (x : K), plusn (n1 * n2)%coq_nat x = plusn n2 (plusn n1 x).
+Proof. intros; rewrite -plusn_assoc_l; f_equal; auto with arith. Qed.
+
+Lemma plusn_mult : forall n (x y : K), plusn n (x * y) = plusn n x * y.
+Proof. intros; unfold plusn; rewrite -mult_sum_distr_r; easy. Qed.
+
+Lemma plusn_eq : forall n (x : K), plusn n x = (plusn1 K n) * x.
+Proof. intros n x; rewrite -plusn_mult mult_one_l; easy. Qed.
+
+Lemma plusn_zero_equiv :
+  forall n, plusn1 K n = 0 <-> forall x : K, plusn n x = 0.
+Proof.
+intros; split; intros Hn; unfold plusn1; try easy.
+intros; rewrite plusn_eq Hn mult_zero_l; easy.
+Qed.
+
+End FF_Ring_Facts2.
+
+
+Section Ring_charac.
+
+Lemma ring_charac_EX :
+  forall K : Ring,
+    { n | n <> 0%nat /\ plusn1 K n = 0 /\
+          forall p, p <> 0%nat -> (p < n)%coq_nat -> plusn1 K p <> 0 } +
+    { forall n, n <> 0%nat -> plusn1 K n <> 0 }.
+Proof.
+intros; pose (P n := n <> 0%nat /\ plusn1 K n = 0).
+destruct (LPO P (fun=> classic _)) as [[N HN] | H]; [left | right].
+(* *)
+apply ex_EX; destruct (dec_inh_nat_subset_has_unique_least_element P)
+    as [n [[Hn1 Hn2] _]]; [intros; apply classic | exists N; easy |].
+exists n; repeat split; try apply Hn1.
+intros p Hp0; rewrite contra_not_r_equiv Nat.nlt_ge; intros Hp1.
+apply Hn2; easy.
+(* *)
+intros n; unfold P in H; specialize (H n); rewrite not_and_equiv in H.
+destruct H; easy.
+Qed.
+
+Definition ring_charac (K : Ring) : nat :=
+  match ring_charac_EX K with
+  | inleft H => proj1_sig H
+  | inright _ => 0
+  end.
+
+Definition is_field_not_charac_2 (K : Ring) : Prop :=
+  is_field K /\ ring_charac K <> 2%nat.
+
+Context {K : Ring}.
+
+Lemma ring_charac_0 :
+  ring_charac K = 0%nat <-> forall n, n <> 0%nat -> plusn1 K n <> 0.
+Proof.
+unfold ring_charac;
+    destruct ring_charac_EX as [[n [Hn0 [Hn1 Hn2]]] | ]; simpl; try easy.
+rewrite iff_not_equiv not_all_ex_not_equiv; split; try easy; intros _.
+exists n; rewrite -contra_equiv not_imp_and_equiv; easy.
+Qed.
+
+Lemma ring_charac_S :
+  forall n, ring_charac K = n.+1 <->
+    plusn n.+1 (1 : K) = 0 /\
+    (forall p, p <> 0%nat -> (p < n.+1)%coq_nat -> plusn1 K p <> 0).
+Proof.
+intros n; unfold ring_charac;
+    destruct ring_charac_EX as [[m [Hm0 [Hm1 Hm2]]] | H];
+    simpl; split; try (now intros; subst); intros [Hn1 Hn2].
+(* *)
+destruct (nat_lt_eq_gt_dec m n.+1) as [[H | H] | H]; try easy.
+contradict Hm1; apply Hn2; easy.
+contradict Hn1; apply Hm2; easy.
+(* *)
+contradict Hn1; apply H; easy.
+Qed.
+
+Lemma ring_charac_equiv :
+  forall n, ring_charac K = n <->
+    (forall k, (forall x : K, plusn k x = 0) <-> k mod n = 0%nat).
+Proof.
+intros n; destruct n.
+(* *)
+rewrite ring_charac_0; unfold Nat.modulo; split.
+(* . *)
+intros H k; split; intros Hk.
+destruct k; try easy.
+specialize (H k.+1); specialize (Hk 1); contradict Hk; auto.
+intros; subst; apply plusn_0_l.
+(* . *)
+intros H n Hn; contradict Hn; apply H.
+intros x; rewrite plusn_eq Hn mult_zero_l; easy.
+(* *)
+rewrite ring_charac_S; split.
+(* . *)
+intros [Hn1 Hn2] k; split.
+(* .. *)
+move=> /plusn_zero_equiv Hk0.
+generalize (Nat.div_mod_eq k n.+1); intros Hk1.
+rewrite Hk1 plusn1_eq plusn_distr_l plusn_assoc_r
+    Hn1 plusn_0_r plus_zero_l in Hk0.
+apply NNPP_equiv; contradict Hk0; apply Hn2; try easy.
+apply Nat.mod_upper_bound; easy.
+(* .. *)
+move=> /Nat.div_exact Hk; apply plusn_zero_equiv; rewrite Hk; try lia.
+unfold plusn1; rewrite plusn_assoc_r Hn1 plusn_0_r; easy.
+(* . *)
+intros Hn; split.
+apply plusn_zero_equiv, Hn, Nat.mod_same; easy.
+intros p Hp0 Hp1; contradict Hp0; rewrite -(Nat.mod_small _ _ Hp1).
+apply Hn, plusn_zero_equiv; easy.
+Qed.
+
+Lemma ring_charac_1 : ring_charac K = 1%nat <-> zero_struct K.
+Proof.
+rewrite -one_is_zero_equiv ring_charac_S plusn_1_l; split; try easy.
+intros H; split; try easy.
+intros p Hp0 Hp1; contradict Hp0; apply lt_1; easy.
+Qed.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (30/06/2023): this should be true for any prime number. *)
+Lemma ring_charac_2 :
+  ring_charac K = 2%nat <-> nonzero_struct K /\ (2 : K) = 0.
+Proof.
+split; [intros HK | intros [HK0 HK1]; rewrite -plusn_2_1 in HK1].
+(* *)
+split.
+(* . *)
+apply nonzero_not_zero_struct_equiv; contradict HK.
+apply ring_charac_1 in HK; rewrite HK; easy.
+(* . *)
+rewrite ring_charac_equiv in HK.
+rewrite -plusn1_2; apply plusn_zero_equiv, HK, Nat.mod_same; easy.
+(* *)
+apply ring_charac_equiv; intros k; rewrite -plusn_zero_equiv; split; intros Hk.
+(* . *)
+assert (Hk1 : k mod 2 = 0%nat \/ k mod 2 = 1%nat)
+    by now apply lt_2, Nat.mod_upper_bound.
+destruct Hk1 as [Hk1 | Hk1]; try easy; contradict Hk.
+generalize (Nat.div_mod_eq k 2); intros Hk2; rewrite Hk2.
+rewrite plusn1_eq plusn_distr_l plusn_assoc_r HK1 plusn_0_r plus_zero_l Hk1 plusn_1_l.
+apply one_not_zero; easy.
+(* . *)
+apply Nat.div_exact in Hk; try easy; rewrite Hk.
+rewrite plusn1_eq plusn_assoc_r HK1 plusn_0_r; easy.
+Qed.
+
+Lemma invertible_plusn :
+  forall n, nonzero_struct K -> n <> 0%nat ->
+    invertible (plusn1 K n) -> ring_charac K <> n.
+Proof.
+intros n HK0 Hn HK1; contradict Hn; destruct n; try easy; contradict HK1.
+apply ring_charac_S in Hn; rewrite plusn1_eq (proj1 Hn).
+apply noninvertible_zero; easy.
+Qed.
+
+End Ring_charac.
+
+
+Section FT_Ring_Facts.
+
+(** Properties with the second identity element of rings (one). *)
+
+Context {K : Ring}.
+
+Lemma onesT : forall {m n} (i : 'I_m), (ones : 'K^{m,n}) i = ones.
+Proof. easy. Qed.
+
+Definition alt_onesT {m n} : 'K^{m,n} :=
+  fun i j => if Nat.Even_Odd_dec (i + j) then 1 else - (1 : K).
+
+End FT_Ring_Facts.
+
+
+Section Ring_R_Facts.
+
+Lemma one_eq_R : 1 = 1%R.
+Proof. easy. Qed.
+
+Lemma nonzero_struct_R : nonzero_struct R_AbelianGroup.
+Proof. exists 1; unfold one, zero; simpl; easy. Qed.
+
+Lemma one_not_zero_R : (1 : R_Ring) <> 0.
+Proof. apply one_not_zero, nonzero_struct_R. Qed.
+
+Lemma minus_one_not_zero_R : - (1 : R_Ring) <> 0.
+Proof. apply minus_one_not_zero, nonzero_struct_R. Qed.
+
+Lemma ones_not_zero_R : forall {n}, (ones : 'R_Ring^n.+1) <> 0.
+Proof. intros n; apply (@ones_not_zero _ n), nonzero_struct_R. Qed.
+
+Lemma sum_ones_R : forall {n}, sum (ones : 'R_Ring^n) = INR n.
+Proof.
+intros n; induction n as [|n Hn].
+rewrite sum_nil; easy.
+rewrite sum_ind_l plus_comm Hn S_INR; easy.
+Qed.
+
+Lemma is_inverse_Rinv : forall a, a <> 0 -> is_inverse a (/ a)%R.
+Proof.
+intros; unfold is_inverse, mult; simpl; rewrite -> Rinv_l, Rinv_r; easy.
+Qed.
+
+Lemma inv_eq_R : forall a, / a = (/ a)%R.
+Proof.
+intros; destruct (Req_dec a 0) as [Ha | Ha].
+rewrite Ha inv_0 Rinv_0; easy.
+apply inv_correct, is_inverse_Rinv; easy.
+Qed.
+
+Lemma div_eq_R : forall a b, a / b = (a / b)%R.
+Proof. intros; unfold div, Rdiv; rewrite inv_eq_R; easy. Qed.
+
+Lemma invertible_equiv_R : forall a : R, invertible a <-> a <> 0.
+Proof.
+intros; split; intros Ha.
+contradict Ha; rewrite Ha; apply (noninvertible_zero nonzero_struct_R).
+apply (Invertible _ (/ a)); rewrite inv_eq_R; apply is_inverse_Rinv; easy.
+Qed.
+
+Lemma non_invertible_equiv_R : forall a : R, ~ invertible a <-> a = 0.
+Proof. intros; rewrite -iff_not_r_equiv; apply invertible_equiv_R. Qed.
+
+Lemma invertible_2 : invertible 2%R.
+Proof. apply invertible_equiv_R; easy. Qed.
+
+Lemma mult_comm_R : is_comm R_Ring.
+Proof. intro; unfold mult; simpl; apply Rmult_comm. Qed.
+
+Lemma mult_pos_R : forall a : R, (0 <= a * a)%R.
+Proof. unfold mult; simpl; apply Rle_0_sqr. Qed.
+
+Lemma mult_inv_l_R : forall a : R, a <> 0 -> (/ a)%R * a = 1.
+Proof.
+intros; rewrite -inv_eq_R; apply mult_inv_l; apply invertible_equiv_R; easy.
+Qed.
+
+Lemma mult_inv_r_R : forall a : R, a <> 0 -> a * (/ a)%R = 1.
+Proof.
+intros; rewrite -inv_eq_R; apply mult_inv_r; apply invertible_equiv_R; easy.
+Qed.
+
+Lemma is_inverse_nonzero_l_R : forall a b : R, is_inverse a b -> a <> 0.
+Proof. move=>>; apply (is_inverse_nonzero_l _ _ nonzero_struct_R). Qed.
+
+Lemma is_inverse_nonzero_r_R : forall a b : R, is_inverse a b -> b <> 0.
+Proof. move=>>; apply (is_inverse_nonzero_r _ _ nonzero_struct_R). Qed.
+
+Lemma is_inverse_equiv_R : forall a : R, is_inverse a (/ a)%R <-> a <> 0.
+Proof.
+intros; split; try apply is_inverse_nonzero_l_R.
+split; [apply mult_inv_l_R | apply mult_inv_r_R]; easy.
+Qed.
+
+Lemma has_inv_R : has_inv R_Ring.
+Proof. move=>> /is_inverse_equiv_R H; apply (Invertible _ _ H). Qed.
+
+Lemma is_comm_field_R : is_comm_field R_Ring.
+Proof.
+repeat split. apply mult_comm_R. apply nonzero_struct_R. apply has_inv_R.
+Qed.
+
+End Ring_R_Facts.
+
+
+Section Kron_Def.
+
+(* FIXME: COMMENTS TO BE REMOVED FOR PUBLICATION!
+ TODO FC (09/10/2023): kron should be defined on a semiring, like nat! *)
+
+Definition kron (K : Ring) (i j : nat) : K :=
+  match (Nat.eq_dec i j) with
+    | left _  => one
+    | right _ => zero
+    end.
+
+Definition kronecker := kron R_Ring.
+
+End Kron_Def.
+
+
+Section Kron_Facts.
+
+Variable K : Ring.
+
+Lemma kron_or : forall i j, kron K i j = 0 \/ kron K i j = 1.
+Proof.
+intros i j; unfold kron; destruct (Nat.eq_dec i j); [right | left]; easy.
+Qed.
+
+Lemma kron_is_1 : forall i j, i = j -> kron K i j = 1.
+Proof.
+intros i j; unfold kron; case (Nat.eq_dec i j); easy.
+Qed.
+
+Lemma kron_1 : forall i j, nonzero_struct K -> kron K i j = 1 -> i = j.
+Proof.
+intros i j HK; unfold kron; case (Nat.eq_dec i j); try easy.
+intros _ H; contradict H; apply not_eq_sym, one_not_zero; easy.
+Qed.
+
+Lemma kron_in_equiv : forall i j, nonzero_struct K -> kron K i j = 1 <-> i = j.
+Proof. intros; split; [apply kron_1; easy | apply kron_is_1]. Qed.
+
+Lemma kron_is_0 : forall i j, i <> j -> kron K i j = 0.
+Proof.
+intros i j; unfold kron; case (Nat.eq_dec i j); easy.
+Qed.
+
+Lemma kron_0 : forall i j, nonzero_struct K -> kron K i j = 0 -> i <> j.
+Proof.
+intros i j HK; unfold kron; case (Nat.eq_dec i j); try easy.
+intros _ H; contradict H; apply one_not_zero; easy.
+Qed.
+
+Lemma kron_out_equiv :
+  forall i j, nonzero_struct K -> kron K i j = 0 <-> i <> j.
+Proof. intros; split; [apply kron_0; easy | apply kron_is_0]. Qed.
+
+Lemma kron_sym : forall i j, kron K i j = kron K j i.
+Proof.
+intros i j; unfold kron at 2; destruct (Nat.eq_dec j i).
+apply kron_is_1; easy.
+apply kron_is_0; auto.
+Qed.
+
+Lemma kron_pred_eq :
+  forall i j, (i <> 0)%nat -> (j <> 0)%nat ->
+    kron K (i - 1) (j - 1) = kron K i j.
+Proof.
+intros i j Hi Hj; destruct (Nat.eq_dec i j) as [H | H].
+rewrite kron_is_1; try now rewrite kron_is_1.
+rewrite H; easy.
+rewrite kron_is_0; try now rewrite kron_is_0.
+contradict H.
+apply Nat.pred_inj; try easy.
+rewrite -2!Nat.sub_1_r; easy.
+Qed.
+
+End Kron_Facts.
+
+
+Section FF_Kron_Facts.
+
+(** Properties with the second identity element of rings (one). *)
+
+Context {K : Ring}.
+
+Lemma kron_skipF_diag_l : forall {n} (i : 'I_n.+1), skipF (kron K i) i = 0.
+Proof.
+intros n i; apply skipF_zero_compat; intros j Hj.
+apply kron_is_0; contradict Hj; apply ord_inj; easy.
+Qed.
+
+Lemma kron_skipF_diag_r :
+  forall {n} (j : 'I_n.+1), skipF (fun i : 'I_n.+1 => kron K i j) j = 0.
+Proof.
+intros n j; apply skipF_zero_compat; intros i Hi.
+apply kron_is_0; contradict Hi; apply ord_inj; easy.
+Qed.
+
+Lemma sum_kron_l : forall {n} (j : 'I_n), sum (fun i : 'I_n => kron K i j) = 1.
+Proof.
+intros n j; destruct n; try now destruct j.
+rewrite (sum_one _ j); try now apply kron_is_1.
+apply kron_skipF_diag_r.
+Qed.
+
+Lemma sum_kron_r : forall {n} (i : 'I_n), sum (fun j : 'I_n => kron K i j) = 1.
+Proof.
+intros n j; destruct n; try now destruct j.
+rewrite (sum_one _ j); try now apply kron_is_1.
+apply kron_skipF_diag_l.
+Qed.
+
+End FF_Kron_Facts.
+
+
+Local Open Scope R_scope.
+
+
+Section Kron_R_Facts.
+
+Lemma kronecker_or : forall i j, kronecker i j = 0 \/ kronecker i j = 1.
+Proof. apply kron_or. Qed.
+
+Lemma kronecker_bound : forall i j, 0 <= kronecker i j <= 1.
+Proof.
+intros i j; destruct (kronecker_or i j) as [H | H]; rewrite H; Lra.lra.
+Qed.
+
+Lemma kronecker_is_1 : forall i j, i = j -> kronecker i j = 1.
+Proof. apply kron_is_1. Qed.
+
+Lemma kronecker_1 : forall i j, kronecker i j = 1 -> i = j.
+Proof. move=>>; apply kron_1, nonzero_struct_R. Qed.
+
+Lemma kronecker_in_equiv : forall i j, kronecker i j = 1 <-> i = j.
+Proof. move=>>; apply kron_in_equiv, nonzero_struct_R. Qed.
+
+Lemma kronecker_is_0 : forall i j, i <> j -> kronecker i j = 0.
+Proof. apply kron_is_0. Qed.
+
+Lemma kronecker_0 : forall i j, kronecker i j = 0 -> i <> j.
+Proof. move=>>; apply kron_0, nonzero_struct_R. Qed.
+
+Lemma kronecker_out_equiv : forall i j, kronecker i j = 0 <-> i <> j.
+Proof. move=>>; apply kron_out_equiv, nonzero_struct_R. Qed.
+
+Lemma kronecker_sym : forall i j, kronecker i j = kronecker j i.
+Proof. apply kron_sym. Qed.
+
+Lemma kronecker_pred_eq :
+  forall i j, (i <> 0)%nat -> (j <> 0)%nat ->
+    kronecker (i - 1) (j - 1) = kronecker i j.
+Proof. apply kron_pred_eq. Qed.
+
+Lemma kronecker_skipF_diag_l :
+  forall {n} (i : 'I_n.+1), skipF (kronecker i) i = 0%M.
+Proof. intros; apply kron_skipF_diag_l. Qed.
+
+Lemma kronecker_skipF_diag_r :
+  forall {n} (j : 'I_n.+1), skipF (fun i : 'I_n.+1 => kronecker i j) j = 0%M.
+Proof. intros; apply kron_skipF_diag_r. Qed.
+
+Lemma sum_kronecker_l :
+  forall {n} (j : 'I_n), sum (fun i : 'I_n => kronecker i j) = 1.
+Proof. intros; apply: sum_kron_l. Qed.
+
+Lemma sum_kronecker_r :
+  forall {n} (i : 'I_n), sum (fun j : 'I_n => kronecker i j) = 1.
+Proof. intros; apply: sum_kron_r. Qed.
+
+End Kron_R_Facts.
+
+
+Section FF_Kron_R_Facts.
+
+Lemma itemF_kronecker_eq :
+  forall {d} (x : R) i, itemF d x i = (fun j => x * kronecker i j).
+Proof.
+intros d x i; apply extF; intros j.
+destruct (ord_eq_dec j i) as [Hj | Hj].
+rewrite Hj itemF_correct_l// kronecker_is_1; auto with real.
+rewrite itemF_correct_r// kronecker_is_0; auto with real.
+contradict Hj; apply ord_inj; easy.
+Qed.
+
+End FF_Kron_R_Facts.
+
+
+Local Close Scope R_scope.
+
+
+Section Compatible_Ring_Def.
+
+Context {K : Ring}.
+
+Definition mult_closed (PK : K -> Prop) : Prop :=
+  forall x y, PK x -> PK y -> PK (x * y).
+
+Definition one_closed (PK : K -> Prop) : Prop := PK one.
+
+Definition compatible_r (PK : K -> Prop) : Prop:=
+  compatible_g PK /\ mult_closed PK /\ one_closed PK.
+
+End Compatible_Ring_Def.
+
+
+Section Compatible_Ring_Facts.
+
+Context {K : Ring}.
+
+Lemma compatible_r_g :
+  forall {PK : K -> Prop}, compatible_r PK -> compatible_g PK.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_r_m :
+  forall {PK : K -> Prop}, compatible_r PK -> compatible_m PK.
+Proof. move=>> /compatible_r_g; apply: compatible_g_m. Qed.
+
+Lemma compatible_r_nonempty :
+  forall {PK : K -> Prop}, compatible_r PK -> nonempty PK.
+Proof. move=>> /compatible_r_g; apply compatible_g_nonempty. Qed.
+
+Lemma compatible_r_zero :
+  forall {PK : K -> Prop}, compatible_r PK -> zero_closed PK.
+Proof. move=>> /compatible_r_g; apply: compatible_g_zero. Qed.
+
+Lemma compatible_r_plus :
+  forall {PK : K -> Prop}, compatible_r PK -> plus_closed PK.
+Proof. move=>> /compatible_r_g; apply: compatible_g_plus. Qed.
+
+Lemma compatible_r_sum :
+  forall {PK : K -> Prop}, compatible_r PK -> sum_closed PK.
+Proof.
+intros; apply plus_zero_sum_closed;
+    [apply compatible_r_plus | apply compatible_r_zero]; easy.
+Qed.
+
+Lemma compatible_r_opp :
+  forall {PK : K -> Prop}, compatible_r PK -> opp_closed PK.
+Proof. move=>> /compatible_r_g; apply compatible_g_opp. Qed.
+
+Lemma compatible_r_minus :
+  forall {PK : K -> Prop}, compatible_r PK -> minus_closed PK.
+Proof. move=>> /compatible_r_g; apply compatible_g_minus. Qed.
+
+Lemma compatible_r_one :
+  forall {PK : K -> Prop}, compatible_r PK -> one_closed PK.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_r_mult :
+  forall {PK : K -> Prop}, compatible_r PK -> mult_closed PK.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_r_full : forall {PK : K -> Prop}, full PK -> compatible_r PK.
+Proof.
+intros; split. apply compatible_g_full; easy. split; try easy.
+unfold one_closed; easy. Qed.
+
+Lemma compatible_r_fullset : compatible_r (@fullset K).
+Proof. apply compatible_r_full; easy. Qed.
+
+Lemma compatible_r_inter :
+  forall {PK QK : K -> Prop},
+    compatible_r PK -> compatible_r QK -> compatible_r (inter PK QK).
+Proof.
+intros PK QK [HPK1 [HPK2 HPK3]] [HQK1 [HQK2 HQK3]]; split.
+apply compatible_g_inter, compatible_r_g; easy.
+split; try easy; intros x y [Hx1 Hx2] [Hy1 Hy2]; split; auto.
+Qed.
+
+(* Gostiaux T1, Th 5.12 p. 133. *)
+Lemma compatible_r_inter_any :
+  forall {Idx : Type} {PK : Idx -> K -> Prop},
+    (forall i, compatible_r (PK i)) -> compatible_r (inter_any PK).
+Proof.
+move=>> HPK; split; [ | split].
+apply compatible_g_inter_any; intros; apply compatible_r_g; easy.
+move=>> Hx Hy i; apply HPK; easy.
+intros i; apply (HPK i).
+Qed.
+
+Definition span_r (gen : K -> Prop) := span_struct compatible_r gen.
+
+Lemma span_r_compatible_r : forall gen, compatible_r (span_r gen).
+Proof.
+apply span_struct_compatible; move=>>; apply compatible_r_inter_any.
+Qed.
+
+Lemma span_r_nonempty : forall gen, nonempty (span_r gen).
+Proof. intros; apply compatible_r_nonempty, span_r_compatible_r. Qed.
+
+Lemma span_r_zero : forall gen, zero_closed (span_r gen).
+Proof. intros; apply compatible_r_zero, span_r_compatible_r. Qed.
+
+Lemma span_r_plus : forall gen, plus_closed (span_r gen).
+Proof. intros; apply compatible_r_plus, span_r_compatible_r. Qed.
+
+Lemma span_r_opp : forall gen, opp_closed (span_r gen).
+Proof. intros; apply compatible_r_opp, span_r_compatible_r. Qed.
+
+Lemma span_r_minus : forall gen, minus_closed (span_r gen).
+Proof. intros; apply compatible_r_minus, span_r_compatible_r. Qed.
+
+Lemma span_r_one : forall gen, one_closed (span_r gen).
+Proof. intros; apply compatible_r_one, span_r_compatible_r. Qed.
+
+Lemma span_r_mult : forall gen, mult_closed (span_r gen).
+Proof. intros; apply compatible_r_mult, span_r_compatible_r. Qed.
+
+Lemma span_r_sum : forall gen, sum_closed (span_r gen).
+Proof. intros; apply compatible_r_sum, span_r_compatible_r. Qed.
+
+Lemma span_r_incl : forall gen, incl gen (span_r gen).
+Proof. apply span_struct_incl. Qed.
+
+Lemma span_r_lub :
+  forall gen PK, compatible_r PK -> incl gen PK -> incl (span_r gen) PK.
+Proof. apply span_struct_lub. Qed.
+
+Lemma span_r_full : forall PK, compatible_r PK -> span_r PK = PK.
+Proof. apply span_struct_full. Qed.
+
+End Compatible_Ring_Facts.
+
+
+Section Compatible_Ring_Morphism_Facts.
+
+Context {K1 K2 : Ring}.
+
+Context {PK1 : K1 -> Prop}.
+Context {PK2 : K2 -> Prop}.
+
+Context {f : K1 -> K2}.
+Hypothesis Hf : morphism_r f.
+
+Lemma compatible_r_image : compatible_r PK1 -> compatible_r (image f PK1).
+Proof.
+intros [HPK1a [HPK1b HPK1c]]; split; [| split].
+apply compatible_g_image; [apply morphism_r_g |]; easy.
+intros _ _ [x1 Hx1] [y1 Hy1]; rewrite -(morphism_r_mult _ Hf); apply Im; auto.
+unfold one_closed; rewrite -(morphism_r_one _ Hf); easy.
+Qed.
+
+Lemma compatible_r_preimage : compatible_r PK2 -> compatible_r (preimage f PK2).
+Proof.
+intros [HPK2a [HPK2b HPK2c]]; split; [| split; unfold preimage].
+apply compatible_g_preimage; [apply morphism_r_g |]; easy.
+move=>> Hx1 Hy1; rewrite (morphism_r_mult _ Hf); auto.
+unfold one_closed; rewrite (morphism_r_one _ Hf); easy.
+Qed.
+
+(* Apparently, there is no structure on morphism_r! *)
+
+End Compatible_Ring_Morphism_Facts.
+
+
+Section Sub_Ring_Def.
+
+Context {K : Ring}.
+Context {PK : K -> Prop}.
+
+Definition sub_r (HPK : compatible_r PK) : Type :=
+  sub_g (compatible_r_g HPK).
+Definition mk_sub_r_ (HPK : compatible_r PK) {x} (Hx : PK x) : sub_r HPK :=
+  mk_sub_g_ (compatible_r_g HPK) Hx.
+Definition mk_sub_r {HPK : compatible_r PK} {x} (Hx : PK x) : sub_r HPK :=
+  mk_sub_r_ HPK Hx.
+
+Hypothesis HPK : compatible_r PK.
+
+Definition sub_mult (x y : sub_r HPK) : sub_r HPK :=
+  mk_sub_r (compatible_r_mult HPK _ _ (in_sub x) (in_sub y)).
+
+Definition sub_one : sub_r HPK := mk_sub_r (compatible_r_one HPK).
+
+Lemma sub_mult_assoc:
+  forall x y z, sub_mult x (sub_mult y z) = sub_mult (sub_mult x y) z.
+Proof. intros; apply val_inj, mult_assoc. Qed.
+
+Lemma sub_mult_one_r : forall x, sub_mult x sub_one = x.
+Proof. intros; apply val_inj, mult_one_r. Qed.
+
+Lemma sub_mult_one_l : forall x, sub_mult sub_one x = x.
+Proof. intros; apply val_inj, mult_one_l. Qed.
+
+Lemma sub_mult_distr_r :
+  forall x y z, sub_mult (x + y) z = sub_mult x z + sub_mult y z.
+Proof. intros; apply val_inj, mult_distr_r. Qed.
+
+Lemma sub_mult_distr_l :
+  forall x y z, sub_mult x (y + z) = sub_mult x y + sub_mult x z.
+Proof. intros; apply val_inj, mult_distr_l. Qed.
+
+Definition sub_Ring_mixin :=
+  Ring.Mixin _ _ _
+    sub_mult_assoc sub_mult_one_r sub_mult_one_l
+    sub_mult_distr_r sub_mult_distr_l.
+
+Canonical Structure sub_Ring :=
+  Ring.Pack _ (Ring.Class _ _ sub_Ring_mixin) (sub_r HPK).
+
+Lemma val_one : f_one_compat val.
+Proof. easy. Qed.
+
+Lemma val_mult : f_mult_compat val.
+Proof. easy. Qed.
+
+Lemma val_morphism_r : morphism_r val.
+Proof. easy. Qed.
+
+Lemma mk_sub_one_equiv : forall {x} (Hx : PK x), mk_sub_r Hx = 1 <-> x = 1.
+Proof. intros; split; [move=> /val_eq | intros; apply val_inj]; easy. Qed.
+
+Lemma mk_sub_mult :
+  forall (x y : K) (Hx : PK x) (Hy : PK y),
+    mk_sub Hx * mk_sub Hy = mk_sub (compatible_r_mult HPK _ _ Hx Hy).
+Proof. easy. Qed.
+
+Lemma sub_one_eq : (1 : sub_r HPK) = mk_sub_ _ (1 : K) (compatible_r_one HPK).
+Proof. easy. Qed.
+
+Lemma sub_mult_eq :
+  forall (x y : sub_r HPK),
+    x * y = mk_sub (compatible_r_mult HPK _ _ (in_sub x) (in_sub y)).
+Proof. easy. Qed.
+
+Lemma mk_sub_r_inj :
+  forall {x y} (Hx : PK x) (Hy : PK y),
+    mk_sub_r Hx = mk_sub_r Hy :> sub_r HPK -> x = y.
+Proof. apply mk_sub_g_inj. Qed.
+
+End Sub_Ring_Def.
+
+
+Section Sub_Ring_Morphism_Facts1.
+
+Context {K1 K2 : Ring}.
+
+Context {PK1 : K1 -> Prop}.
+Context {PK2 : K2 -> Prop}.
+Hypothesis HPK1 : compatible_r PK1.
+Hypothesis HPK2 : compatible_r PK2.
+
+Context {f : K1 -> K2}.
+Context {fS : sub_r HPK1 -> sub_r HPK2}.
+Hypothesis HfS : forall x, val (fS x) = f (val x).
+
+Lemma sub_r_fun_rev : funS PK1 PK2 f.
+Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
+
+Lemma sub_r_inj : injS PK1 f -> injective fS.
+Proof. apply sub_g_inj, HfS. Qed.
+
+Lemma sub_r_inj_rev : injective fS -> injS PK1 f.
+Proof. apply sub_g_inj_rev, HfS. Qed.
+
+Lemma sub_r_inj_equiv : injective fS <-> injS PK1 f.
+Proof. apply sub_g_inj_equiv, HfS. Qed.
+
+Lemma sub_r_surj : surjS PK1 PK2 f -> surjective fS.
+Proof. apply sub_g_surj, HfS. Qed.
+
+Lemma sub_r_surj_rev : surjective fS -> surjS PK1 PK2 f.
+Proof. apply sub_g_surj_rev, HfS. Qed.
+
+Lemma sub_r_surj_equiv : surjective fS <-> surjS PK1 PK2 f.
+Proof. apply sub_g_surj_equiv, HfS. Qed.
+
+Lemma sub_r_bij : bijS PK1 PK2 f -> bijective fS.
+Proof. apply sub_g_bij, HfS. Qed.
+
+Lemma sub_r_bij_rev : bijective fS -> bijS PK1 PK2 f.
+Proof. apply sub_g_bij_rev, HfS. Qed.
+
+Lemma sub_r_bij_equiv : bijective fS <-> bijS PK1 PK2 f.
+Proof. apply sub_g_bij_equiv, HfS. Qed.
+
+Lemma sub_r_f_mult_compat : f_mult_compat f -> f_mult_compat fS.
+Proof. intros Hf x1 y1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
+
+Lemma sub_r_f_one_compat : f_one_compat f -> f_one_compat fS.
+Proof. intros Hf; apply val_inj; rewrite HfS Hf; easy. Qed.
+
+Lemma sub_r_morphism : morphism_r f -> morphism_r fS.
+Proof.
+intros [Hf1 [Hf2 Hf3]]; repeat split; [apply (sub_g_morphism HfS Hf1) |
+    apply (sub_r_f_mult_compat Hf2) | apply (sub_r_f_one_compat Hf3)].
+Qed.
+
+End Sub_Ring_Morphism_Facts1.
+
+
+Section Sub_Ring_Morphism_Facts2.
+
+Context {K1 K2 : Ring}.
+
+Context {PK1 : K1 -> Prop}.
+Context {PK2 : K2 -> Prop}.
+Hypothesis HPK1 : compatible_r PK1.
+Hypothesis HPK2 : compatible_r PK2.
+
+Context {f : K1 -> K2}.
+Hypothesis Hf : funS PK1 PK2 f.
+
+Definition fct_sub_r : sub_r HPK1 -> sub_r HPK2 :=
+  fct_sub_g (compatible_r_g HPK1) (compatible_r_g HPK2) Hf.
+
+Lemma fct_sub_r_inj : injS PK1 f -> injective fct_sub_r.
+Proof. apply fct_sub_g_inj. Qed.
+
+Lemma fct_sub_r_inj_rev : injective fct_sub_r -> injS PK1 f.
+Proof. apply fct_sub_g_inj_rev. Qed.
+
+Lemma fct_sub_r_inj_equiv : injective fct_sub_r <-> injS PK1 f.
+Proof. apply fct_sub_g_inj_equiv. Qed.
+
+Lemma fct_sub_r_surj : surjS PK1 PK2 f -> surjective fct_sub_r.
+Proof. apply fct_sub_g_surj. Qed.
+
+Lemma fct_sub_r_surj_rev : surjective fct_sub_r -> surjS PK1 PK2 f.
+Proof. apply fct_sub_g_surj_rev. Qed.
+
+Lemma fct_sub_r_surj_equiv : surjective fct_sub_r <-> surjS PK1 PK2 f.
+Proof. apply fct_sub_g_surj_equiv. Qed.
+
+Lemma fct_sub_r_bij : bijS PK1 PK2 f -> bijective fct_sub_r.
+Proof. apply fct_sub_g_bij. Qed.
+
+Lemma fct_sub_r_bij_rev : bijective fct_sub_r -> bijS PK1 PK2 f.
+Proof. apply fct_sub_g_bij_rev. Qed.
+
+Lemma fct_sub_r_bij_equiv : bijective fct_sub_r <-> bijS PK1 PK2 f.
+Proof. apply fct_sub_g_bij_equiv. Qed.
+
+Lemma fct_sub_r_f_mult_compat : f_mult_compat f -> f_mult_compat fct_sub_r.
+Proof. apply sub_r_f_mult_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_r_f_one_compat : f_one_compat f -> f_one_compat fct_sub_r.
+Proof. apply sub_r_f_one_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_r_morphism : morphism_r f -> morphism_r fct_sub_r.
+Proof. apply sub_r_morphism, fct_sub_correct. Qed.
+
+End Sub_Ring_Morphism_Facts2.