diff --git a/FEM/Algebra/AffineSpace.v b/FEM/Algebra/AffineSpace.v
index 9eff43c0765cd2a384a2ecb489030e42b66c8086..861c1f325316d6c0e6015d1de7529be4cc4fa334 100644
--- a/FEM/Algebra/AffineSpace.v
+++ b/FEM/Algebra/AffineSpace.v
@@ -24,10 +24,10 @@ From LM Require Import linear_map.
 Close Scope R_scope.
 Set Warnings "notation-overridden".
 
-From Lebesgue Require Import Subset Subset_charac Function.
+From Lebesgue Require Import Subset Subset_dec Subset_charac Subset_any Function.
 
 From FEM Require Import Compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl Group_compl Ring_compl.
+From FEM Require Import Sub_struct Monoid_compl Group_compl Ring_compl.
 From FEM Require Import ModuleSpace_compl ModuleSpace_R_compl.
 
 
@@ -348,7 +348,7 @@ Lemma fct_vect_l_bij_ex :
   forall (fA : T -> E) (fu : T -> V), exists! fB, fct_vect fA fB = fu.
 Proof.
 intros fA fu.
-destruct (unique_choice _ _ (fun x B => fA x --> B = fu x)) as [fB HfB].
+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.
@@ -652,7 +652,7 @@ Qed.
 
 Lemma vectF_l_bij_ex : forall {n} O u, exists! (A : 'E^n), vectF O A = u.
 Proof.
-intros n O u; destruct (unique_choice _ _ _ (vect_l_bij_ex O)) as [A HA].
+intros n O u; 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.
@@ -2250,3 +2250,1105 @@ 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.
+
+
+(* 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.
+
+Lemma PE_as_inhabited : inhabited (sub_struct HPE).
+Proof. apply sub_struct_inhabited; exists O; easy. Qed.
+
+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 _ _ _ PE_as_inhabited _ 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/Algebra.v b/FEM/Algebra/Algebra.v
index 0f6fc8ff169a11718823e8ba2d9d982904042fa0..e58d71b129f46e5df7d658e6d1b259145715ac87 100644
--- a/FEM/Algebra/Algebra.v
+++ b/FEM/Algebra/Algebra.v
@@ -15,6 +15,6 @@ COPYING file for more details.
 *)
 
 From FEM Require Export nat_compl ord_compl Finite_family Finite_table.
-From FEM Require Export Monoid_compl Group_compl Ring_compl.
+From FEM Require Export Sub_struct Monoid_compl Group_compl Ring_compl.
 From FEM Require Export ModuleSpace_compl ModuleSpace_R_compl matrix.
-From FEM Require Export AffineSpace Sub_struct MonoidComp Finite_dim Finite_dim_R.
+From FEM Require Export MonoidComp AffineSpace Finite_dim Finite_dim_R.
diff --git a/FEM/Algebra/Group_compl.v b/FEM/Algebra/Group_compl.v
index a16900e2289c90d4de20a78546437b473658b2bb..dacfe44c948c8b931f313e59343d4d613db0ea9a 100644
--- a/FEM/Algebra/Group_compl.v
+++ b/FEM/Algebra/Group_compl.v
@@ -24,10 +24,10 @@ From LM Require Import linear_map.
 Close Scope R_scope.
 Set Warnings "notation-overridden".
 
-From Lebesgue Require Import Subset Subset_charac Function.
+From Lebesgue Require Import Subset Subset_charac Subset_any Function.
 
 From FEM Require Import Compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl.
+From FEM Require Import Sub_struct Monoid_compl.
 
 
 Declare Scope Group_scope.
@@ -329,3 +329,737 @@ Lemma constF_opp : forall n (x : G), constF n (- x) = - constF n x.
 Proof. easy. Qed.
 
 End FF_Group_Facts1.
+
+
+Section Compatible_Group_Def.
+
+Context {G : AbelianGroup}.
+
+Definition opp_closed (PG : G -> Prop) : Prop := forall x, PG x -> PG (- x).
+
+Definition minus_closed (PG : G -> Prop) : Prop :=
+  forall x y, PG x -> PG y -> PG (x - y).
+
+End Compatible_Group_Def.
+
+
+Section Compatible_Group_Facts.
+
+Context {G : AbelianGroup}.
+
+Lemma plus_opp_minus_closed :
+  forall {PG : G -> Prop}, plus_closed PG -> opp_closed PG -> minus_closed PG.
+Proof. intros PG H1 H2 x; unfold minus; auto. Qed.
+
+Lemma minus_zero_closed :
+  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> zero_closed PG.
+Proof.
+move=>> [x Hx]; unfold zero_closed; rewrite -(minus_eq_zero x); auto.
+Qed.
+
+Lemma plus_opp_zero_closed :
+  forall {PG : G -> Prop},
+    nonempty PG -> plus_closed PG -> opp_closed PG -> zero_closed PG.
+Proof. intros; apply minus_zero_closed, plus_opp_minus_closed; easy. Qed.
+
+Lemma minus_opp_closed :
+  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> opp_closed PG.
+Proof.
+intros PG HPG0 HPG1 x; rewrite -(minus_zero_l x).
+apply HPG1, minus_zero_closed; easy.
+Qed.
+
+Lemma minus_plus_closed :
+  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> plus_closed PG.
+Proof.
+intros PG HPG0 HPG1 x y Hx Hy.
+rewrite -minus_opp; apply HPG1; try easy.
+apply minus_opp_closed; easy.
+Qed.
+
+Lemma minus_plus_opp_closed_equiv :
+  forall {PG : G -> Prop},
+    nonempty PG -> minus_closed PG <-> plus_closed PG /\ opp_closed PG.
+Proof.
+intros; split.
+intros; split; [apply minus_plus_closed | apply minus_opp_closed]; easy.
+intros; apply plus_opp_minus_closed; easy.
+Qed.
+
+(* TO BE REMOVED FOR PUBLICATION!
+ TODO FC (09/06/2023):
+   compatible_g should be compatible_m /\ opp_closed.
+   Then, compatible_g <-> nonempty /\ plus_closed /\ opp_closed
+                      <-> nonempty /\ minus_closed. *)
+
+Lemma minus_closed_compatible_g :
+  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> compatible_g PG.
+Proof. intros; split; auto. Qed.
+
+Lemma plus_opp_closed_compatible_g :
+  forall {PG : G -> Prop},
+    nonempty PG -> plus_closed PG -> opp_closed PG -> compatible_g PG.
+Proof. intros; split; auto. Qed.
+
+Lemma compatible_g_nonempty :
+  forall {PG : G -> Prop}, compatible_g PG -> nonempty PG.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_g_zero :
+  forall {PG : G -> Prop}, compatible_g PG -> zero_closed PG.
+Proof. apply compatible_g_zero. Qed.
+
+Lemma compatible_g_plus :
+  forall {PG : G -> Prop}, compatible_g PG -> plus_closed PG.
+Proof. intros; do 2 intro; apply compatible_g_plus; easy. Qed.
+
+Lemma compatible_g_opp :
+  forall {PG : G -> Prop}, compatible_g PG -> opp_closed PG.
+Proof. intros; intro; apply compatible_g_opp; easy. Qed.
+
+Lemma compatible_g_minus :
+  forall {PG : G -> Prop}, compatible_g PG -> minus_closed PG.
+Proof. intros PG H x; apply H. Qed.
+
+Lemma compatible_g_m :
+  forall {PG : G -> Prop}, compatible_g PG -> compatible_m PG.
+Proof.
+intros; split; [apply compatible_g_plus | apply compatible_g_zero]; easy.
+Qed.
+
+Lemma compatible_m_g :
+  forall {PG : G -> Prop},
+    opp_closed PG -> compatible_m PG -> compatible_g PG.
+Proof.
+intros; apply plus_opp_closed_compatible_g;
+  [apply compatible_m_nonempty | apply compatible_m_plus | ]; easy.
+Qed.
+
+Lemma compatible_g_m_equiv :
+  forall {PG : G -> Prop},
+    opp_closed PG -> compatible_g PG <-> compatible_m PG.
+Proof. intros; split; [apply compatible_g_m | apply compatible_m_g; easy]. Qed.
+
+Lemma compatible_g_zero_sub_struct : compatible_g (@zero_sub_struct G).
+Proof.
+split; try now exists 0.
+intros x y Hx Hy; rewrite Hx Hy opp_zero plus_zero_l; easy.
+Qed.
+
+Lemma compatible_g_full : forall {PG : G -> Prop}, full PG -> compatible_g PG.
+Proof. intros; split; try exists 0; easy. Qed.
+
+Lemma compatible_g_fullset : compatible_g (@fullset G).
+Proof. apply compatible_g_full; easy. Qed.
+
+Lemma compatible_g_inter :
+  forall {PG QG : G -> Prop},
+    compatible_g PG -> compatible_g QG -> compatible_g (inter PG QG).
+Proof.
+intros PG QG HPG HQG; split.
+destruct HPG as [HPG _], HQG as [HQG _];
+    intros x y [Hx1 Hx2] [Hy1 Hy2]; split; auto.
+exists 0; split; apply compatible_g_zero; easy.
+Qed.
+
+(* Gostiaux T1, Th 4.6 p. 89. *)
+Lemma compatible_g_inter_any :
+  forall {Idx : Type} {PG : Idx -> G -> Prop},
+    (forall i, compatible_g (PG i)) -> compatible_g (inter_any PG).
+Proof.
+intros Idx PG HPG; split.
+intros x y Hx Hy i; apply HPG; easy.
+exists 0; intros i; apply compatible_g_zero; easy.
+Qed.
+
+(* Gostiaux T1, Th 4.10 pp. 89-91. *)
+Definition span_g (gen : G -> Prop) := span_struct compatible_g gen.
+
+Lemma span_g_compatible_g : forall gen, compatible_g (span_g gen).
+Proof.
+apply span_struct_compatible; move=>>; apply compatible_g_inter_any.
+Qed.
+
+Lemma span_g_nonempty : forall gen, nonempty (span_g gen).
+Proof. intros; apply compatible_g_nonempty, span_g_compatible_g. Qed.
+
+Lemma span_g_zero : forall gen, zero_closed (span_g gen).
+Proof. intros; apply compatible_g_zero, span_g_compatible_g. Qed.
+
+Lemma span_g_plus : forall gen, plus_closed (span_g gen).
+Proof. intros; apply compatible_g_plus, span_g_compatible_g. Qed.
+
+Lemma span_g_opp : forall gen, opp_closed (span_g gen).
+Proof. intros; apply compatible_g_opp, span_g_compatible_g. Qed.
+
+Lemma span_g_minus : forall gen, minus_closed (span_g gen).
+Proof. intros; apply compatible_g_minus, span_g_compatible_g. Qed.
+
+Lemma span_g_incl : forall gen, incl gen (span_g gen).
+Proof. apply span_struct_incl. Qed.
+
+(* Gostiaux T1, Def 4.9 p. 90. *)
+Lemma span_g_lub :
+  forall gen PG, compatible_g PG -> incl gen PG -> incl (span_g gen) PG.
+Proof. apply span_struct_lub. Qed.
+
+Lemma span_g_full : forall PG, compatible_g PG -> span_g PG = PG.
+Proof. apply span_struct_full. Qed.
+
+End Compatible_Group_Facts.
+
+
+Section Compatible_Group_Morphism_Facts1.
+
+Context {G1 G2 : AbelianGroup}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_g f.
+
+Lemma compatible_g_image : compatible_g PG1 -> compatible_g (image f PG1).
+Proof.
+intros [HPG1 [x01 Hx01]]; split; try now exists (f x01).
+intros _ _ [x1 Hx1] [y1 Hy1]; rewrite -(morphism_g_opp f Hf) -Hf.
+apply Im; auto.
+Qed.
+
+Lemma compatible_g_preimage : compatible_g PG2 -> compatible_g (preimage f PG2).
+Proof.
+intros HPG2; split; unfold preimage.
+destruct HPG2 as [HPG2 _]; intros; rewrite Hf morphism_g_opp; auto.
+apply compatible_g_zero in HPG2; exists 0; rewrite morphism_g_zero; easy.
+Qed.
+
+Lemma compatible_g_morphism_g : compatible_g (@morphism_g G1 G2).
+Proof.
+split; [apply morphism_g_fct_minus | exists 0; apply morphism_g_fct_zero].
+Qed.
+
+End Compatible_Group_Morphism_Facts1.
+
+
+Section Compatible_Group_Morphism_Def.
+
+Context {G1 G2 : AbelianGroup}.
+
+Variable PG1 : G1 -> Prop.
+
+Variable f : G1 -> G2.
+
+Definition morphism_g_sub : Prop := f_plus_compat_sub PG1 f.
+
+End Compatible_Group_Morphism_Def.
+
+
+Section Compatible_Group_Morphism_Facts2a.
+
+Context {G1 G2 : AbelianGroup}.
+
+Context {PG1 : G1 -> Prop}.
+Hypothesis HPG1 : compatible_g PG1.
+
+Context {PG2 : G2 -> Prop}.
+Hypothesis HPG2 : compatible_g PG2.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_g f.
+
+Lemma KerS_compatible_g : compatible_g (KerS PG1 f).
+Proof.
+apply compatible_g_inter, compatible_g_preimage,
+    compatible_g_zero_sub_struct; easy.
+Qed.
+
+Lemma RgS_gen_compatible_g : compatible_g (RgS_gen PG1 PG2 f).
+Proof. apply compatible_g_inter, compatible_g_image; easy. Qed.
+
+Lemma RgS_compatible_g : compatible_g (RgS PG1 f).
+Proof. apply compatible_g_image; easy. Qed.
+
+Lemma KerS_g_zero_equiv :
+  KerS PG1 f = zero_sub_struct <-> incl (KerS PG1 f) zero_sub_struct.
+Proof.
+apply KerS_m_zero_equiv; [apply compatible_g_m | apply morphism_g_m]; easy.
+Qed.
+
+Lemma gmS_injS : incl (KerS PG1 f) zero_sub_struct -> injS PG1 f.
+Proof.
+intros Hf1 x1 y1 Hx1 Hy1; rewrite -2!minus_zero_equiv -(morphism_g_minus _ Hf).
+intros H1; apply Hf1, KerS_equiv; split; [apply HPG1 |]; easy.
+Qed.
+
+Lemma gmS_injS_rev : injS PG1 f -> KerS PG1 f = zero_sub_struct.
+Proof.
+apply mmS_injS_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
+Qed.
+
+Lemma gmS_injS_equiv : injS PG1 f <-> KerS PG1 f = zero_sub_struct.
+Proof.
+split; [apply gmS_injS_rev | rewrite KerS_g_zero_equiv; apply gmS_injS].
+Qed.
+
+Lemma gmS_bijS_gen :
+  funS PG1 PG2 f -> incl (KerS PG1 f) zero_sub_struct -> incl PG2 (RgS PG1 f) ->
+  bijS PG1 PG2 f.
+Proof.
+intros; apply bijS_equiv; [apply inhabited_g | repeat split];
+    [| apply gmS_injS | apply surjS_correct]; easy.
+Qed.
+
+Lemma gmS_bijS_gen_rev :
+  bijS PG1 PG2 f ->
+  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
+Proof.
+apply mmS_bijS_gen_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
+Qed.
+
+Lemma gmS_bijS_gen_equiv :
+  bijS PG1 PG2 f <->
+  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
+Proof.
+split; [apply gmS_bijS_gen_rev | rewrite KerS_g_zero_equiv RgS_gen_full_equiv].
+intros; apply gmS_bijS_gen; easy.
+Qed.
+
+Lemma gmS_bijS : incl (KerS PG1 f) zero_sub_struct -> bijS PG1 (RgS PG1 f) f.
+Proof.
+rewrite bijS_RgS_equiv; [| apply inhabited_g]; apply gmS_injS; easy.
+Qed.
+
+Lemma gmS_bijS_rev : bijS PG1 (RgS PG1 f) f -> KerS PG1 f = zero_sub_struct.
+Proof.
+apply mmS_bijS_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
+Qed.
+
+Lemma gmS_bijS_equiv : bijS PG1 (RgS PG1 f) f <-> KerS PG1 f = zero_sub_struct.
+Proof.
+split; [apply gmS_bijS_rev | rewrite KerS_g_zero_equiv//; apply gmS_bijS].
+Qed.
+
+Lemma gmS_bijS_injS_equiv : bijS PG1 (RgS PG1 f) f <-> injS PG1 f.
+Proof. apply mmS_bijS_injS_equiv. Qed.
+
+End Compatible_Group_Morphism_Facts2a.
+
+
+Section Compatible_Group_Morphism_Facts2b.
+
+Context {G1 G2 : AbelianGroup}.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_g f.
+
+Lemma Ker_compatible_g : compatible_g (Ker f).
+Proof.
+rewrite -KerS_full; apply (KerS_compatible_g compatible_g_fullset Hf).
+Qed.
+
+Lemma Rg_compatible_g : compatible_g (Rg f).
+Proof.
+rewrite -RgS_full; apply (RgS_compatible_g compatible_g_fullset Hf).
+Qed.
+
+Lemma Ker_g_zero_equiv :
+  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
+Proof. apply Ker_m_zero_equiv, morphism_g_m; easy. Qed.
+
+Lemma gm_inj : incl (Ker f) zero_sub_struct -> injective f.
+Proof.
+rewrite -KerS_full inj_S_equiv; apply (gmS_injS compatible_g_fullset Hf).
+Qed.
+
+Lemma gm_inj_rev : injective f -> Ker f = zero_sub_struct.
+Proof. apply mm_inj_rev, morphism_g_m; easy. Qed.
+
+Lemma gm_inj_equiv : injective f <-> Ker f = zero_sub_struct.
+Proof.
+rewrite -KerS_full inj_S_equiv; apply (gmS_injS_equiv compatible_g_fullset Hf).
+Qed.
+
+Lemma gm_bij :
+  incl (Ker f) zero_sub_struct -> incl fullset (Rg f) -> bijective f.
+Proof.
+rewrite -KerS_full -RgS_full bij_S_equiv; [| apply inhabited_g].
+apply (gmS_bijS_gen compatible_g_fullset Hf), funS_full.
+Qed.
+
+Lemma gm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
+Proof. apply mm_bij_rev, morphism_g_m; easy. Qed.
+
+Lemma gm_bij_equiv : bijective f <-> Ker f = zero_sub_struct /\ Rg f = fullset.
+Proof.
+rewrite -KerS_full -RgS_gen_full bij_S_equiv; [| apply inhabited_g].
+rewrite (gmS_bijS_gen_equiv compatible_g_fullset Hf); easy.
+Qed.
+
+End Compatible_Group_Morphism_Facts2b.
+
+
+Section Compatible_Group_Morphism_Facts3.
+
+Context {E1 E2 E3 : AbelianGroup}.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : morphism_g f.
+
+Context {g : E2 -> E3}.
+Hypothesis Hg : morphism_g g.
+
+Lemma Ker_g_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
+Proof. apply Ker_m_comp_r; apply morphism_g_m; easy. Qed.
+
+Lemma Ker_g_comp_r_alt :
+  forall h x2, injective g -> h = g \o f ->
+    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
+Proof. apply Ker_m_comp_r_alt; apply morphism_g_m; easy. Qed.
+
+End Compatible_Group_Morphism_Facts3.
+
+
+Section Compatible_Group_Morphism_Facts4a.
+
+Context {G1 G2 : AbelianGroup}.
+
+Variable PG1 : G1 -> Prop.
+
+Lemma morphism_g_sub_id : morphism_g_sub PG1 ssrfun.id.
+Proof. easy. Qed.
+
+Variable f : G1 -> G2.
+
+Lemma morphism_g_is_sub : morphism_g f -> morphism_g_sub PG1 f.
+Proof. easy. Qed.
+
+End Compatible_Group_Morphism_Facts4a.
+
+
+Section Compatible_Group_Morphism_Facts4b.
+
+Context {G1 G2 G3 : AbelianGroup}.
+
+Context {PG1 : G1 -> Prop}.
+Variable PG2 : G2 -> Prop.
+
+Variable f : G1 -> G2.
+Variable g : G2 -> G3.
+
+Hypothesis Hf : funS PG1 PG2 f.
+
+Lemma morphism_g_comp_sub :
+  morphism_g_sub PG1 f -> morphism_g_sub PG2 g ->
+  morphism_g_sub PG1 (g \o f).
+Proof. apply f_plus_compat_comp_sub; easy. Qed.
+
+End Compatible_Group_Morphism_Facts4b.
+
+
+Section Sub_Group_Def.
+
+Context {G : AbelianGroup}.
+Context {PG : G -> Prop}.
+
+Definition sub_g (HPG : compatible_g PG) : Type :=
+  sub_struct (compatible_g_m HPG).
+Definition mk_sub_g_ (HPG : compatible_g PG) {x} (Hx : PG x) : sub_g HPG :=
+  mk_sub_ (compatible_g_m HPG) x Hx.
+Definition mk_sub_g {HPG : compatible_g PG} {x} (Hx : PG x) : sub_g HPG :=
+  mk_sub_g_ HPG Hx.
+
+Hypothesis HPG : compatible_g PG.
+
+Definition sub_opp (x : sub_g HPG) : sub_g HPG :=
+  mk_sub_g (compatible_g_opp HPG _ (in_sub x)).
+
+Lemma sub_plus_opp_r : forall x, x + (sub_opp x) = 0.
+Proof. intros; apply val_inj, plus_opp_r. Qed.
+
+Definition sub_AbelianGroup_mixin := AbelianGroup.Mixin _ _ sub_plus_opp_r.
+
+Canonical Structure sub_AbelianGroup :=
+  AbelianGroup.Pack _
+    (AbelianGroup.Class _ _ sub_AbelianGroup_mixin) (sub_g HPG).
+
+Lemma val_opp : f_opp_compat val.
+Proof. easy. Qed.
+
+Lemma val_minus : f_minus_compat val.
+Proof. easy. Qed.
+
+Lemma val_morphism_g : morphism_g val.
+Proof. easy. Qed.
+
+Lemma mk_sub_g_zero : mk_sub_g (compatible_g_zero HPG) = (0 : sub_g HPG).
+Proof. apply val_inj; easy. Qed.
+
+Lemma mk_sub_g_zero_equiv :
+  forall {x} (Hx : PG x), mk_sub_g Hx = 0 :> sub_g HPG <-> x = 0.
+Proof. apply mk_sub_zero_equiv. Qed.
+
+Lemma mk_sub_opp :
+  forall (x : G) (Hx : PG x), - mk_sub Hx = mk_sub (compatible_g_opp HPG _ Hx).
+Proof. easy. Qed.
+
+Lemma mk_sub_minus :
+  forall (x y : G) (Hx : PG x) (Hy : PG y),
+    mk_sub Hx - mk_sub Hy = mk_sub (compatible_g_minus HPG _ _ Hx Hy).
+Proof. intros; apply val_inj; easy. Qed.
+
+Lemma sub_opp_eq :
+  forall (x : sub_g HPG), - x = mk_sub (compatible_g_opp HPG _ (in_sub x)).
+Proof. easy. Qed.
+
+Lemma sub_minus_eq :
+  forall (x y : sub_g HPG),
+    x - y = mk_sub (compatible_g_minus HPG _ _ (in_sub x) (in_sub y)).
+Proof. intros; apply mk_sub_minus. Qed.
+
+Lemma mk_sub_g_inj :
+  forall {x y} (Hx : PG x) (Hy : PG y),
+    mk_sub_g Hx = mk_sub_g Hy :> sub_g HPG -> x = y.
+Proof. apply mk_sub_inj. Qed.
+
+End Sub_Group_Def.
+
+
+Section Sub_Group_Facts1.
+
+Context {G : AbelianGroup}.
+Context {PGa PGb : G -> Prop}.
+Hypothesis HPG : incl PGa PGb.
+Hypothesis HPGb : compatible_g PGb.
+Let PGb_g := sub_g HPGb.
+Let PGa' : PGb_g -> Prop := preimage val PGa.
+
+Lemma RgS_val_g_eq : RgS PGa' val = PGa.
+Proof.
+apply RgS_preimage_equiv; intros x Hx; apply Rg_ex;
+    exists (mk_sub (HPG _ Hx)); easy.
+Qed.
+
+Lemma preimage_val_compatible_g : compatible_g PGa -> compatible_g PGa'.
+Proof. intros; apply compatible_g_preimage; easy. Qed.
+
+Lemma preimage_val_compatible_g_rev : compatible_g PGa' -> compatible_g PGa.
+Proof. intros; rewrite -RgS_val_g_eq; apply RgS_compatible_g; easy. Qed.
+
+Lemma preimage_val_compatible_g_equiv : compatible_g PGa' <-> compatible_g PGa.
+Proof.
+split; [apply preimage_val_compatible_g_rev | apply preimage_val_compatible_g].
+Qed.
+
+End Sub_Group_Facts1.
+
+
+Section Sub_Group_Facts2.
+
+Context {G : AbelianGroup}.
+Context {PGb : G -> Prop}.
+Hypothesis HPGb : compatible_g PGb.
+Let PGb_g := sub_g HPGb.
+Context {PGa : PGb_g -> Prop}.
+Let PGa' := RgS PGa val.
+
+Lemma preimage_val_g_eq : preimage val PGa' = PGa.
+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_g : compatible_g PGa -> compatible_g PGa'.
+Proof. intros; apply RgS_compatible_g; easy. Qed.
+
+Lemma RgS_val_compatible_g_rev : compatible_g PGa' -> compatible_g PGa.
+Proof.
+intros; rewrite -preimage_val_g_eq; apply compatible_g_preimage; easy.
+Qed.
+
+Lemma RgS_val_compatible_g_equiv : compatible_g PGa' <-> compatible_g PGa.
+Proof.
+split; [apply RgS_val_compatible_g_rev | apply RgS_val_compatible_g].
+Qed.
+
+End Sub_Group_Facts2.
+
+
+Section Sub_Group_Morphism_Facts1.
+
+Context {G1 G2 : AbelianGroup}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+Context {HPG1 : compatible_g PG1}.
+Context {HPG2 : compatible_g PG2}.
+
+Context {f : G1 -> G2}.
+Context {fS : sub_g HPG1 -> sub_g HPG2}.
+Hypothesis HfS : forall x, val (fS x) = f (val x).
+
+Lemma sub_g_fun_rev : funS PG1 PG2 f.
+Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
+
+Lemma sub_g_inj : injS PG1 f -> injective fS.
+Proof. apply sub_m_inj, HfS. Qed.
+
+Lemma sub_g_inj_rev : injective fS -> injS PG1 f.
+Proof. apply sub_m_inj_rev, HfS. Qed.
+
+Lemma sub_g_inj_equiv : injective fS <-> injS PG1 f.
+Proof. apply sub_m_inj_equiv, HfS. Qed.
+
+Lemma sub_g_surj : surjS PG1 PG2 f -> surjective fS.
+Proof. apply sub_m_surj, HfS. Qed.
+
+Lemma sub_g_surj_rev : surjective fS -> surjS PG1 PG2 f.
+Proof. apply sub_m_surj_rev, HfS. Qed.
+
+Lemma sub_g_surj_equiv : surjective fS <-> surjS PG1 PG2 f.
+Proof. apply sub_m_surj_equiv, HfS. Qed.
+
+Lemma sub_g_bij : bijS PG1 PG2 f -> bijective fS.
+Proof. apply sub_m_bij, HfS. Qed.
+
+Lemma sub_g_bij_rev : bijective fS -> bijS PG1 PG2 f.
+Proof. apply sub_m_bij_rev, HfS. Qed.
+
+Lemma sub_g_bij_equiv : bijective fS <-> bijS PG1 PG2 f.
+Proof. apply sub_m_bij_equiv, HfS. Qed.
+
+Lemma sub_g_f_opp_compat : f_opp_compat f -> f_opp_compat fS.
+Proof. intros Hf x1; apply val_inj; rewrite HfS Hf -HfS; easy. Qed.
+
+Lemma sub_g_f_minus_compat : f_minus_compat f -> f_minus_compat fS.
+Proof. intros Hf x1 y1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
+
+Lemma sub_g_morphism : morphism_g f -> morphism_g fS.
+Proof. apply sub_m_f_plus_compat, HfS. Qed.
+
+Lemma Ker_sub_g_KerS_zero_equiv :
+  morphism_g f ->
+  Ker fS = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
+Proof. move=> /morphism_g_m; apply Ker_sub_zero_equiv, HfS. Qed.
+
+Lemma KerS_g_zero_equiv_alt :
+  morphism_g f ->
+  KerS PG1 f = zero_sub_struct <->
+  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
+Proof.
+move=>>; rewrite subset_ext_equiv; split; [intros [Hf1 _] | intros Hf1; split].
+intros [x1 Hx1] Hx1a; apply val_inj, Hf1; easy.
+intros x1 [Hx1 Hx1a]; apply (mk_sub_g_inj HPG1 Hx1 (compatible_g_zero HPG1)).
+rewrite mk_sub_g_zero; apply Hf1; easy.
+move=> x1 ->; apply compatible_g_zero, KerS_compatible_g; easy.
+Qed.
+
+Lemma Ker_sub_g_zero_equiv :
+  morphism_g f ->
+  Ker fS = zero_sub_struct <->
+  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros; rewrite Ker_sub_g_KerS_zero_equiv//; apply KerS_g_zero_equiv_alt; easy.
+Qed.
+
+Lemma gmS_injS_sub_equiv :
+  morphism_g f -> injS PG1 f <-> Ker fS = zero_sub_struct.
+Proof.
+intros Hf; rewrite gmS_injS_equiv//; split;
+    intros Hf1; apply subset_ext; intros x1.
+rewrite (Ker_sub_KerS_equiv HfS) Hf1; apply val_zero_equiv.
+destruct (classic_dec (PG1 x1)) as [Hx1 | Hx1a].
+rewrite (KerS_Ker_sub_equiv HfS Hx1) Hf1; apply mk_sub_zero_equiv.
+split; intros Hx1b; contradict Hx1a;
+    [apply Hx1b | rewrite Hx1b; apply compatible_g_zero, HPG1].
+Qed.
+
+Lemma gmS_injS_sub_equiv_alt :
+  morphism_g f ->
+  injS PG1 f <-> forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros H; rewrite (gmS_injS_sub_equiv H); apply (Ker_sub_g_zero_equiv H).
+Qed.
+
+Lemma gmS_bijS_sub_equiv :
+  morphism_g f -> bijS PG1 (RgS PG1 f) f <-> Ker fS = zero_sub_struct.
+Proof. rewrite gmS_bijS_injS_equiv; apply gmS_injS_sub_equiv. Qed.
+
+Lemma gmS_bijS_sub_equiv_alt :
+  morphism_g f ->
+  bijS PG1 (RgS PG1 f) f <-> forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
+Proof. rewrite gmS_bijS_injS_equiv; apply gmS_injS_sub_equiv_alt. Qed.
+
+End Sub_Group_Morphism_Facts1.
+
+
+Section Sub_Group_Morphism_Facts2.
+
+Context {G1 G2 : AbelianGroup}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+Hypothesis HPG1 : compatible_g PG1.
+Hypothesis HPG2 : compatible_g PG2.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : funS PG1 PG2 f.
+
+Definition fct_sub_g : sub_g HPG1 -> sub_g HPG2 :=
+  fct_sub (compatible_g_m HPG1) (compatible_g_m HPG2) Hf.
+
+Lemma fct_sub_g_inj : injS PG1 f -> injective fct_sub_g.
+Proof. apply fct_sub_m_inj. Qed.
+
+Lemma fct_sub_g_inj_rev : injective fct_sub_g -> injS PG1 f.
+Proof. apply fct_sub_m_inj_rev. Qed.
+
+Lemma fct_sub_g_inj_equiv : injective fct_sub_g <-> injS PG1 f.
+Proof. apply fct_sub_m_inj_equiv. Qed.
+
+Lemma fct_sub_g_surj : surjS PG1 PG2 f -> surjective fct_sub_g.
+Proof. apply fct_sub_m_surj. Qed.
+
+Lemma fct_sub_g_surj_rev : surjective fct_sub_g -> surjS PG1 PG2 f.
+Proof. apply fct_sub_m_surj_rev. Qed.
+
+Lemma fct_sub_g_surj_equiv : surjective fct_sub_g <-> surjS PG1 PG2 f.
+Proof. apply fct_sub_m_surj_equiv. Qed.
+
+Lemma fct_sub_g_bij : bijS PG1 PG2 f -> bijective fct_sub_g.
+Proof. apply fct_sub_m_bij. Qed.
+
+Lemma fct_sub_g_bij_rev : bijective fct_sub_g -> bijS PG1 PG2 f.
+Proof. apply fct_sub_m_bij_rev. Qed.
+
+Lemma fct_sub_g_bij_equiv : bijective fct_sub_g <-> bijS PG1 PG2 f.
+Proof. apply fct_sub_m_bij_equiv. Qed.
+
+Lemma fct_sub_g_f_opp_compat : f_opp_compat f -> f_opp_compat fct_sub_g.
+Proof. apply sub_g_f_opp_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_g_f_minus_compat : f_minus_compat f -> f_minus_compat fct_sub_g.
+Proof. apply sub_g_f_minus_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_g_morphism : morphism_g f -> morphism_g fct_sub_g.
+Proof. apply sub_g_morphism, fct_sub_correct. Qed.
+
+Lemma Ker_fct_sub_g_KerS_zero_equiv :
+  morphism_g f ->
+  Ker fct_sub_g = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
+Proof. apply Ker_sub_g_KerS_zero_equiv, fct_sub_correct. Qed.
+
+Lemma Ker_fct_sub_g_zero_equiv :
+  morphism_g f ->
+  Ker fct_sub_g = zero_sub_struct <->
+  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
+Proof. apply Ker_sub_g_zero_equiv, fct_sub_correct. Qed.
+
+Lemma gmS_injS_fct_sub_equiv :
+  morphism_g f -> injS PG1 f <-> Ker fct_sub_g = zero_sub_struct.
+Proof. apply gmS_injS_sub_equiv, fct_sub_correct. Qed.
+
+Lemma gmS_bijS_fct_sub_equiv :
+  morphism_g f -> bijS PG1 (RgS PG1 f) f <-> Ker fct_sub_g = zero_sub_struct.
+Proof. apply gmS_bijS_sub_equiv, fct_sub_correct. Qed.
+
+End Sub_Group_Morphism_Facts2.
diff --git a/FEM/Algebra/ModuleSpace_R_compl.v b/FEM/Algebra/ModuleSpace_R_compl.v
index 68ed188b98cfb98413f9b685ada511043b5c376a..1d6d245c576190a3c2620fadf6087a08b9940a9b 100644
--- a/FEM/Algebra/ModuleSpace_R_compl.v
+++ b/FEM/Algebra/ModuleSpace_R_compl.v
@@ -29,7 +29,8 @@ Set Warnings "notation-overridden".
 From Lebesgue Require Import Subset Subset_charac Function.
 
 From FEM Require Import Compl nat_compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl Group_compl Ring_compl ModuleSpace_compl.
+From FEM Require Import Sub_struct Monoid_compl Group_compl Ring_compl.
+From FEM Require Import ModuleSpace_compl.
 
 
 (** Results needing a commutative Ring.
@@ -623,3 +624,39 @@ Canonical Structure Rn_PreHilbert {n} :=
   PreHilbert.Pack 'R^n (PreHilbert.Class _ _ Rn_PreHilbert_mixin) 'R^n.
 
 End Dot_Product_R_Facts.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_R_Facts1.
+
+Variable E1 E2 : ModuleSpace R_Ring.
+
+Lemma compatible_ms_linear_mapping : compatible_ms (@is_linear_mapping _ E1 E2).
+Proof.
+split; [apply compatible_g_is_linear_mapping |].
+intros; apply is_linear_mapping_f_scal; easy.
+Qed.
+
+(* Lm is the vector subspace of linear mappings. *)
+Definition Lm := sub_ms compatible_ms_linear_mapping.
+
+End Compatible_ModuleSpace_Linear_Mapping_R_Facts1.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_R_Facts2.
+
+Variable E : ModuleSpace R_Ring.
+
+(* TO BE REMOVED FOR PUBLICATION!
+ TODO FC (25/11/2023): define ring of endomorphisms (with + and \o). *)
+
+
+
+Definition Endom : (E -> E) -> Prop := fun f => is_linear_mapping f.
+
+(* TO BE REMOVED FOR PUBLICATION!
+ TODO FC (25/11/2023): define non-Abelian group of automorphisms (with \o). *)
+
+Definition Autom : (E -> E) -> Prop :=
+  fun f => is_linear_mapping f /\ bijective f.
+
+End Compatible_ModuleSpace_Linear_Mapping_R_Facts2.
diff --git a/FEM/Algebra/ModuleSpace_compl.v b/FEM/Algebra/ModuleSpace_compl.v
index 0328ae9226dcd7a5f26a6158ed2017085d254d21..bc6afa3363c95cd57d393b4084b8da09ecd4509e 100644
--- a/FEM/Algebra/ModuleSpace_compl.v
+++ b/FEM/Algebra/ModuleSpace_compl.v
@@ -26,10 +26,10 @@ From LM Require Import hilbert.
 Close Scope R_scope.
 Set Warnings "notation-overridden".
 
-From Lebesgue Require Import Subset Subset_charac Function.
+From Lebesgue Require Import Subset Subset_charac Subset_any Function.
 
 From FEM Require Import Compl nat_compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl Group_compl Ring_compl.
+From FEM Require Import Sub_struct Monoid_compl Group_compl Ring_compl.
 
 
 (** Results needing a commutative Ring are only stated for
@@ -1689,3 +1689,906 @@ Lemma dot_product_zero_r : forall {n} (u : 'K^n), u â‹… 0 = 0.
 Proof. intros; apply: comb_lin_zero_r. Qed.
 
 End Dot_Product_Facts.
+
+
+Section Compatible_ModuleSpace_Def.
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Definition scal_closed (PE : E -> Prop) : Prop :=
+  forall a u, PE u -> PE (scal a u).
+
+Definition scal_rev_closed (PE : E -> Prop) : Prop :=
+  forall a u, invertible a -> PE (scal a u) -> PE u.
+
+Definition comb_lin_closed (PE : E -> Prop) : Prop :=
+  forall n L (B : 'E^n), inclF B PE -> PE (comb_lin L B).
+
+End Compatible_ModuleSpace_Def.
+
+
+Section Compatible_ModuleSpace_Facts.
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+
+Lemma scal_scal_rev_closed :
+  forall (PE : E -> Prop), scal_closed PE -> scal_rev_closed PE.
+Proof.
+intros PE HPE a u [b [Hb _]]; rewrite -{2}(scal_one u) -Hb -scal_assoc; auto.
+Qed.
+
+Lemma scal_rev_scal_closed :
+  forall (PE : E -> Prop),
+    has_inv K -> zero_closed PE -> scal_rev_closed PE -> scal_closed PE.
+Proof.
+move=>> HK HPE0 HPE a u Hu; destruct (classic (a = 0)) as [Ha | Ha].
+rewrite Ha scal_zero_l; easy. specialize (HK a Ha).
+generalize HK; intros [b Hb]; generalize Hb; intros [Hb1 _].
+rewrite -(scal_one u) -Hb1 -scal_assoc in Hu; apply (HPE b); try easy.
+apply (is_inverse_invertible_r a _ Hb).
+Qed.
+
+Lemma scal_closed_equiv :
+  forall (PE : E -> Prop),
+    has_inv K -> zero_closed PE -> scal_closed PE <-> scal_rev_closed PE.
+Proof.
+move=>> HK HPE; split.
+apply scal_scal_rev_closed.
+apply scal_rev_scal_closed; easy.
+Qed.
+
+Lemma scal_zero_closed :
+  forall {PE : E -> Prop}, nonempty PE -> scal_closed PE -> zero_closed PE.
+Proof.
+intros PE [x Hx]; unfold zero_closed; rewrite -(scal_zero_l x); auto.
+Qed.
+
+Lemma scal_opp_closed :
+  forall {PE : E -> Prop}, scal_closed PE -> opp_closed PE.
+Proof. intros PE H u; rewrite <- scal_opp_one; apply H. Qed.
+
+Lemma comb_lin_zero_closed :
+  forall {PE : E -> Prop}, comb_lin_closed PE -> zero_closed PE.
+Proof.
+move=>> HPE.
+destruct (hat0F_is_nonempty K) as [L], (hat0F_is_nonempty E) as [B].
+specialize (HPE 0%nat L B); rewrite comb_lin_nil in HPE.
+apply HPE, inclF_nil.
+Qed.
+
+Lemma plus_scal_comb_lin_closed :
+  forall {PE : E -> Prop},
+    nonempty PE -> plus_closed PE -> scal_closed PE -> comb_lin_closed PE.
+Proof.
+intros PE HPE0 HPE1 HPE2 n L B HB; induction n as [| n Hn].
+rewrite comb_lin_nil; apply scal_zero_closed; easy.
+rewrite comb_lin_ind_l; apply HPE1. apply HPE2, HB. apply Hn; intro; apply HB.
+Qed.
+
+Lemma comb_lin_plus_scal_closed :
+  forall {PE : E -> Prop},
+    comb_lin_closed PE -> plus_closed PE /\ scal_closed PE.
+Proof.
+intros PE HPE; split.
+(* *)
+intros u v Hu Hv; rewrite -sum_coupleF -comb_lin_ones_l.
+apply HPE; intros i; destruct (ord2_dec i);
+    [rewrite coupleF_l | rewrite coupleF_r]; easy.
+(* *)
+intros a u Hu; replace (scal a u) with (comb_lin (singleF a) (singleF u)).
+apply HPE; intro; rewrite singleF_0; easy.
+rewrite comb_lin_1 2!singleF_0; easy.
+Qed.
+
+Lemma comb_lin_plus_scal_closed_equiv :
+  forall {PE : E -> Prop},
+    nonempty PE -> comb_lin_closed PE <-> plus_closed PE /\ scal_closed PE.
+Proof.
+intros; split.
+apply comb_lin_plus_scal_closed.
+intros; apply plus_scal_comb_lin_closed; easy.
+Qed.
+
+(* TO BE REMOVED FOR PUBLICATION!
+ TODO FC (09/06/2023):
+   compatible_ms should be compatible_g /\ scal_closed.
+   Then, compatible_ms <-> nonempty /\ plus_closed /\ scal_closed
+                       <-> comb_lin_closed. *)
+
+Lemma plus_scal_closed_compatible_ms :
+  forall {PE : E -> Prop},
+    nonempty PE -> plus_closed PE -> scal_closed PE -> compatible_ms PE.
+Proof.
+intros; split; auto.
+apply plus_opp_closed_compatible_g; try apply scal_opp_closed; easy.
+Qed.
+
+Lemma comb_lin_closed_compatible_ms :
+  forall {PE : E -> Prop}, comb_lin_closed PE -> compatible_ms PE.
+Proof.
+move=>> HPE; apply plus_scal_closed_compatible_ms.
+exists 0; apply comb_lin_zero_closed; easy.
+1,2: apply comb_lin_plus_scal_closed; easy.
+Qed.
+
+Lemma compatible_ms_g :
+  forall {PE : E -> Prop}, compatible_ms PE -> compatible_g PE.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_g_ms :
+  forall {PE : E -> Prop},
+    scal_closed PE -> compatible_g PE -> compatible_ms PE.
+Proof.
+intros; apply plus_scal_closed_compatible_ms;
+    [apply compatible_g_nonempty | apply: compatible_g_plus |]; easy.
+Qed.
+
+Lemma compatible_ms_g_equiv :
+  forall {PE : E -> Prop},
+    scal_closed PE -> compatible_ms PE <-> compatible_g PE.
+Proof.
+intros; split; [apply compatible_ms_g | apply compatible_g_ms; easy].
+Qed.
+
+Lemma compatible_ms_m :
+  forall {PE : E -> Prop}, compatible_ms PE -> compatible_m PE.
+Proof. move=>> /compatible_ms_g; apply: compatible_g_m. Qed.
+
+Lemma compatible_m_ms :
+  forall {PE : E -> Prop},
+    scal_closed PE -> compatible_m PE -> compatible_ms PE.
+Proof.
+intros; apply compatible_g_ms;
+    [| apply compatible_m_g; [apply scal_opp_closed |]]; easy.
+Qed.
+
+Lemma compatible_ms_m_equiv :
+  forall {PE : E -> Prop},
+    scal_closed PE -> compatible_ms PE <-> compatible_m PE.
+Proof.
+intros; split; [apply compatible_ms_m | apply compatible_m_ms; easy].
+Qed.
+
+Lemma compatible_ms_nonempty :
+  forall {PE : E -> Prop}, compatible_ms PE -> nonempty PE.
+Proof. move=>> /compatible_ms_g; apply compatible_g_nonempty. Qed.
+
+Lemma compatible_ms_zero :
+  forall {PE : E -> Prop}, compatible_ms PE -> zero_closed PE.
+Proof. move=>> /compatible_ms_g; apply: compatible_g_zero. Qed.
+
+Lemma compatible_ms_plus :
+  forall {PE : E -> Prop}, compatible_ms PE -> plus_closed PE.
+Proof. move=>> /compatible_ms_g; apply: compatible_g_plus. Qed.
+
+Lemma compatible_ms_opp :
+  forall {PE : E -> Prop}, compatible_ms PE -> opp_closed PE.
+Proof. move=>> /compatible_ms_g; apply compatible_g_opp. Qed.
+
+Lemma compatible_ms_minus :
+  forall {PE : E -> Prop}, compatible_ms PE -> minus_closed PE.
+Proof. move=>> /compatible_ms_g; apply compatible_g_minus. Qed.
+
+Lemma compatible_ms_scal :
+  forall {PE : E -> Prop}, compatible_ms PE -> scal_closed PE.
+Proof. move=>> [_ H]; move=>>; apply H. Qed.
+
+Lemma compatible_ms_scal_rev :
+  forall {PE : E -> Prop}, compatible_ms PE -> scal_rev_closed PE.
+Proof. move=>> /compatible_ms_scal; apply scal_scal_rev_closed. Qed.
+
+Lemma compatible_ms_comb_lin :
+  forall {PE : E -> Prop}, compatible_ms PE -> comb_lin_closed PE.
+Proof.
+intros PE HPE n L B HB; induction n as [| n Hn].
+(* *)
+rewrite comb_lin_nil.
+apply compatible_ms_zero; easy.
+(* *)
+rewrite comb_lin_ind_l.
+apply compatible_ms_plus; try easy.
+apply HPE, HB.
+apply Hn.
+intros i; apply HB.
+Qed.
+
+Lemma compatible_ms_plus_equiv :
+  forall {PE : E -> Prop},
+    compatible_ms PE -> forall u v, PE u -> PE (u + v) <-> PE v.
+Proof.
+intros PE HPE u v Hu; split; [rewrite -{2}(minus_plus_l u v) minus_sym | ].
+1,2: apply compatible_ms_plus; try easy.
+apply compatible_ms_opp; easy.
+Qed.
+
+Lemma compatible_ms_zero_sub_struct : compatible_ms (@zero_sub_struct E).
+Proof.
+split; try apply: compatible_g_zero_sub_struct.
+intros u a Hu; rewrite Hu scal_zero_r; easy.
+Qed.
+
+Lemma compatible_ms_full :
+  forall {PE : E -> Prop}, full PE -> compatible_ms PE.
+Proof. intros; split; try apply compatible_g_full; easy. Qed.
+
+Lemma compatible_ms_fullset : compatible_ms (@fullset E).
+Proof. apply compatible_ms_full; easy. Qed.
+
+Lemma compatible_ms_inter :
+  forall {PE QE : E -> Prop},
+    compatible_ms PE -> compatible_ms QE -> compatible_ms (inter PE QE).
+Proof.
+intros PE QE [HPE1 HPE2] [HQE1 HQE2]; split.
+apply compatible_g_inter; easy.
+intros u a [Hu1 Hu2]; split; auto.
+Qed.
+
+(* Gostiaux T1, Th 6.15 p. 167. *)
+Lemma compatible_ms_inter_any :
+  forall {Idx : Type} {PE : Idx -> E -> Prop},
+    (forall i, compatible_ms (PE i)) -> compatible_ms (inter_any PE).
+Proof.
+intros Idx PE HPE; split.
+apply compatible_g_inter_any; intros; apply compatible_ms_g; easy.
+intros x l Hx i; apply HPE; easy.
+Qed.
+
+(* Gostiaux T1, Th 6.17 p. 168. *)
+Definition span_ms (gen : E -> Prop) := span_struct compatible_ms gen.
+
+Lemma span_ms_compatible_ms : forall gen, compatible_ms (span_ms gen).
+Proof.
+apply span_struct_compatible; move=>>; apply compatible_ms_inter_any.
+Qed.
+
+Lemma span_ms_nonempty : forall (gen : E -> Prop), nonempty (span_ms gen).
+Proof. intros; apply compatible_ms_nonempty, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_zero : forall (gen : E -> Prop), zero_closed (span_ms gen).
+Proof. intros; apply compatible_ms_zero; apply span_ms_compatible_ms. Qed.
+
+Lemma span_ms_plus : forall (gen : E -> Prop), plus_closed (span_ms gen).
+Proof. intros; apply compatible_ms_plus, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_opp : forall (gen : E -> Prop), opp_closed (span_ms gen).
+Proof. intros; apply compatible_ms_opp, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_minus : forall (gen : E -> Prop), minus_closed (span_ms gen).
+Proof. intros; apply compatible_ms_minus, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_scal : forall (gen : E -> Prop), scal_closed (span_ms gen).
+Proof. intros; apply compatible_ms_scal, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_comb_lin :
+  forall (gen : E -> Prop), comb_lin_closed (span_ms gen).
+Proof. intros; apply compatible_ms_comb_lin, span_ms_compatible_ms. Qed.
+
+Lemma span_ms_incl : forall gen, incl gen (span_ms gen).
+Proof. apply span_struct_incl. Qed.
+
+(* Gostiaux T1, Def 6.16 p. 167. *)
+Lemma span_ms_lub :
+  forall gen PE, compatible_ms PE -> incl gen PE -> incl (span_ms gen) PE.
+Proof. apply span_struct_lub. Qed.
+
+Lemma span_ms_full : forall PE, compatible_ms PE -> span_ms PE = PE.
+Proof. apply span_struct_full. Qed.
+
+End Compatible_ModuleSpace_Facts.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts1.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : is_linear_mapping f.
+
+Lemma compatible_ms_image : compatible_ms PE1 -> compatible_ms (image f PE1).
+Proof.
+destruct Hf as [_ Hf1]; intros [HPE1a HPE1b]; split.
+apply compatible_g_image; [apply is_linear_mapping_morphism_g |]; easy.
+intros _ a [u Hu]; rewrite -Hf1; apply Im; auto.
+Qed.
+
+Lemma compatible_ms_preimage :
+  compatible_ms PE2 -> compatible_ms (preimage f PE2).
+Proof.
+destruct Hf as [_ Hf1]; intros [HPE2a HPE2b]; split.
+apply compatible_g_preimage; [apply is_linear_mapping_morphism_g |]; easy.
+intros u a Hu; unfold preimage; rewrite Hf1; auto.
+Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts1.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Def.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Variable PE1 : E1 -> Prop.
+
+Variable f : E1 -> E2.
+
+Definition f_scal_compat_sub : Prop :=
+  forall a x1 , PE1 x1 -> f (scal a x1) = scal a (f x1).
+
+Definition is_linear_mapping_sub : Prop :=
+  f_plus_compat_sub PE1 f /\ f_scal_compat_sub.
+
+End Compatible_ModuleSpace_Linear_Mapping_Def.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts2a.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Context {PE1 : E1 -> Prop}.
+Hypothesis HPE1 : compatible_ms PE1.
+
+Context {PE2 : E2 -> Prop}.
+Hypothesis HPE2 : compatible_ms PE2.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : is_linear_mapping f.
+
+Lemma KerS_compatible_ms : compatible_ms (KerS PE1 f).
+Proof.
+apply compatible_ms_inter, compatible_ms_preimage,
+    compatible_ms_zero_sub_struct; easy.
+Qed.
+
+Lemma RgS_gen_compatible_ms : compatible_ms (RgS_gen PE1 PE2 f).
+Proof. apply compatible_ms_inter, compatible_ms_image; easy. Qed.
+
+Lemma RgS_compatible_ms : compatible_ms (RgS PE1 f).
+Proof. apply compatible_ms_image; easy. Qed.
+
+Lemma KerS_zero_equiv :
+  KerS PE1 f = zero_sub_struct <-> incl (KerS PE1 f) zero_sub_struct.
+Proof.
+apply: KerS_g_zero_equiv;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_injS : incl (KerS PE1 f) zero_sub_struct -> injS PE1 f.
+Proof.
+apply: gmS_injS;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_injS_rev : injS PE1 f -> KerS PE1 f = zero_sub_struct.
+Proof.
+apply: gmS_injS_rev;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_injS_equiv : injS PE1 f <-> KerS PE1 f = zero_sub_struct.
+Proof.
+apply: gmS_injS_equiv;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_gen :
+  funS PE1 PE2 f -> incl (KerS PE1 f) zero_sub_struct -> incl PE2 (RgS PE1 f) ->
+  bijS PE1 PE2 f.
+Proof.
+apply: gmS_bijS_gen;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_gen_rev :
+  bijS PE1 PE2 f ->
+  funS PE1 PE2 f /\ KerS PE1 f = zero_sub_struct /\ RgS_gen PE1 PE2 f = PE2.
+Proof.
+apply: gmS_bijS_gen_rev;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_gen_equiv :
+  bijS PE1 PE2 f <->
+  funS PE1 PE2 f /\ KerS PE1 f = zero_sub_struct /\ RgS_gen PE1 PE2 f = PE2.
+Proof.
+apply: gmS_bijS_gen_equiv;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS : incl (KerS PE1 f) zero_sub_struct -> bijS PE1 (RgS PE1 f) f.
+Proof.
+apply: gmS_bijS;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_rev : bijS PE1 (RgS PE1 f) f -> KerS PE1 f = zero_sub_struct.
+Proof.
+apply: gmS_bijS_rev;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_equiv : bijS PE1 (RgS PE1 f) f <-> KerS PE1 f = zero_sub_struct.
+Proof.
+apply: gmS_bijS_equiv;
+    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
+Qed.
+
+Lemma lmS_bijS_injS_equiv : bijS PE1 (RgS PE1 f) f <-> injS PE1 f.
+Proof. apply gmS_bijS_injS_equiv. Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts2a.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts2b.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : is_linear_mapping f.
+
+Lemma Ker_compatible_ms : compatible_ms (Ker f).
+Proof.
+rewrite -KerS_full; apply (KerS_compatible_ms compatible_ms_fullset Hf).
+Qed.
+
+Lemma Rg_compatible_ms : compatible_ms (Rg f).
+Proof.
+rewrite -RgS_full; apply (RgS_compatible_ms compatible_ms_fullset Hf).
+Qed.
+
+Lemma Ker_zero_equiv :
+  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
+Proof.
+apply: Ker_g_zero_equiv; apply is_linear_mapping_morphism_g; easy.
+Qed.
+
+Lemma lm_inj : incl (Ker f) zero_sub_struct -> injective f.
+Proof. apply: gm_inj; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma lm_inj_rev : injective f -> Ker f = zero_sub_struct.
+Proof. apply: gm_inj_rev; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma lm_inj_equiv : injective f <-> Ker f = zero_sub_struct.
+Proof. apply: gm_inj_equiv; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma lm_bij :
+  incl (Ker f) zero_sub_struct -> incl fullset (Rg f) -> bijective f.
+Proof. apply: gm_bij; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma lm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
+Proof. apply: gm_bij_rev; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma lm_bij_equiv : bijective f <-> Ker f = zero_sub_struct /\ Rg f = fullset.
+Proof. apply: gm_bij_equiv; apply is_linear_mapping_morphism_g; easy. Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts2b.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts3.
+
+Context {K : Ring}.
+Context {E1 E2 E3 : ModuleSpace K}.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : is_linear_mapping f.
+
+Context {g : E2 -> E3}.
+Hypothesis Hg : is_linear_mapping g.
+
+Lemma Ker_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
+Proof. apply: Ker_g_comp_r; apply is_linear_mapping_morphism_g; easy. Qed.
+
+Lemma Ker_comp_r_alt :
+  forall h x2, injective g -> h = g \o f ->
+    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
+Proof.
+apply: Ker_g_comp_r_alt; apply is_linear_mapping_morphism_g; easy.
+Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts3.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts4a.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Variable PE1 : E1 -> Prop.
+
+Lemma f_scal_compat_sub_id : f_scal_compat_sub PE1 ssrfun.id.
+Proof. easy. Qed.
+
+Lemma is_linear_mapping_sub_id : is_linear_mapping_sub PE1 ssrfun.id.
+Proof. easy. Qed.
+
+Variable f : E1 -> E2.
+
+Lemma f_scal_compat_is_sub : f_scal_compat f -> f_scal_compat_sub PE1 f.
+Proof. easy. Qed.
+
+Lemma is_linear_mapping_is_sub :
+  is_linear_mapping f -> is_linear_mapping_sub PE1 f.
+Proof. intros [Hfa Hfb]; easy. Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts4a.
+
+
+Section Compatible_ModuleSpace_Linear_Mapping_Facts4b.
+
+Context {K : Ring}.
+Context {E1 E2 E3 : ModuleSpace K}.
+
+Context {PE1 : E1 -> Prop}.
+Variable PE2 : E2 -> Prop.
+
+Variable f : E1 -> E2.
+Variable g : E2 -> E3.
+
+Hypothesis Hf : funS PE1 PE2 f.
+
+Lemma f_scal_compat_comp_sub :
+  f_scal_compat_sub PE1 f -> f_scal_compat_sub PE2 g ->
+  f_scal_compat_sub PE1 (g \o f).
+Proof.
+intros Hf1 Hg1 a x1 Hx1; rewrite comp_correct Hf1// Hg1//; apply Hf; easy.
+Qed.
+
+Lemma is_linear_mapping_comp_sub :
+  is_linear_mapping_sub PE1 f -> is_linear_mapping_sub PE2 g ->
+  is_linear_mapping_sub PE1 (g \o f).
+Proof.
+intros [Hfa Hfb] [Hga Hgb]; split;
+    [apply (f_plus_compat_comp_sub PE2) | apply f_scal_compat_comp_sub]; easy.
+Qed.
+
+End Compatible_ModuleSpace_Linear_Mapping_Facts4b.
+
+
+Section Sub_ModuleSpace_Def.
+
+Context {K : Ring}.
+Context {E : ModuleSpace K}.
+Context {PE : E -> Prop}.
+
+Definition sub_ms (HPE : compatible_ms PE) : Type :=
+  sub_g (compatible_ms_g HPE).
+Definition mk_sub_ms_ (HPE : compatible_ms PE) {x} (Hx : PE x) : sub_ms HPE :=
+  mk_sub_g_ (compatible_ms_g HPE) Hx.
+Definition mk_sub_ms {HPE : compatible_ms PE} {x} (Hx : PE x) : sub_ms HPE :=
+  mk_sub_ms_ HPE Hx.
+
+Hypothesis HPE : compatible_ms PE.
+
+Definition sub_scal a (x : sub_ms HPE) : sub_ms HPE :=
+  mk_sub_ms (compatible_ms_scal HPE a (val x) (in_sub x)).
+
+Lemma sub_scal_assoc :
+  forall a b x, sub_scal a (sub_scal b x) = sub_scal (a * b) x.
+Proof. intros; apply val_inj, scal_assoc. Qed.
+
+Lemma sub_scal_one_l : forall x, sub_scal 1 x = x.
+Proof. intros; apply val_inj, scal_one. Qed.
+
+Lemma sub_scal_distr_l :
+  forall a x y, sub_scal a (x + y) = sub_scal a x + sub_scal a y.
+Proof. intros; apply val_inj, scal_distr_l. Qed.
+
+Lemma sub_scal_distr_r  :
+  forall a b x, sub_scal (a + b) x = sub_scal a x + sub_scal b x.
+Proof. intros; apply val_inj, scal_distr_r. Qed.
+
+Definition sub_ModuleSpace_mixin :=
+  ModuleSpace.Mixin _ _ _
+    sub_scal_assoc sub_scal_one_l sub_scal_distr_l sub_scal_distr_r.
+
+Canonical Structure sub_ModuleSpace :=
+  ModuleSpace.Pack _ _
+    (ModuleSpace.Class _ _ _ sub_ModuleSpace_mixin) (sub_ms HPE).
+
+Lemma val_scal : f_scal_compat val.
+Proof. easy. Qed.
+
+Lemma val_comb_lin : f_comb_lin_compat val.
+Proof.
+intros n; induction n as [| n Hn].
+intros; rewrite 2!comb_lin_nil; easy.
+intros; rewrite 2!comb_lin_ind_l; simpl; f_equal; apply Hn.
+Qed.
+
+Lemma val_lin_map : is_linear_mapping val.
+Proof. easy. Qed.
+
+Lemma mk_sub_scal :
+  forall a (x : E) (Hx : PE x),
+    scal a (mk_sub_ms Hx) = mk_sub_ms (compatible_ms_scal HPE a _ Hx).
+Proof. easy. Qed.
+
+Lemma mk_sub_comb_lin :
+  forall {n} L (B : 'E^n) (HB : inclF B PE),
+    comb_lin L (fun i => mk_sub_ms (HB i)) =
+      mk_sub_ms (compatible_ms_comb_lin HPE _ L _ HB).
+Proof. intros; apply val_inj, val_comb_lin. Qed.
+
+Lemma sub_scal_eq :
+  forall a (x : sub_ms HPE),
+    scal a x = mk_sub_ms (compatible_ms_scal HPE a _ (in_sub x)).
+Proof. easy. Qed.
+
+Lemma sub_comb_lin_eq :
+  forall {n} L (B : '(sub_ms HPE)^n),
+    comb_lin L B =
+      mk_sub_ms (compatible_ms_comb_lin HPE _ L _ (fun i => in_sub (B i))).
+Proof. intros; apply mk_sub_comb_lin. Qed.
+
+Lemma mk_sub_ms_inj :
+  forall {x y} (Hx : PE x) (Hy : PE y),
+    mk_sub_ms Hx = mk_sub_ms Hy :> sub_ms HPE -> x = y.
+Proof. apply mk_sub_g_inj. Qed.
+
+End Sub_ModuleSpace_Def.
+
+
+(* TO BE REMOVED FOR PUBLICATION!
+ TODO FC (30/05/2023):
+  Define sum of sub_ms (Gostiaux T1, Th 6.20 p. 169),
+         direct sum, supplementary (Gostiaux T1, S 6.5 pp. 182-187). *)
+
+
+Section Sub_ModuleSpace_Facts1.
+
+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 RgS_val_ms_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_ms : compatible_ms PEa -> compatible_ms PEa'.
+Proof. intros; apply compatible_ms_preimage; easy. Qed.
+
+Lemma preimage_val_compatible_ms_rev : compatible_ms PEa' -> compatible_ms PEa.
+Proof. intros; rewrite -RgS_val_ms_eq; apply RgS_compatible_ms; easy. Qed.
+
+Lemma preimage_val_compatible_ms_equiv : compatible_ms PEa' <-> compatible_ms PEa.
+Proof.
+split; [apply preimage_val_compatible_ms_rev | apply preimage_val_compatible_ms].
+Qed.
+
+End Sub_ModuleSpace_Facts1.
+
+
+Section Sub_ModuleSpace_Facts2.
+
+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_val_ms_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_ms : compatible_ms PEa -> compatible_ms PEa'.
+Proof. intros; apply RgS_compatible_ms; easy. Qed.
+
+Lemma RgS_val_compatible_ms_rev : compatible_ms PEa' -> compatible_ms PEa.
+Proof.
+intros; rewrite -preimage_val_ms_eq; apply compatible_ms_preimage; easy.
+Qed.
+
+Lemma RgS_val_compatible_ms_equiv : compatible_ms PEa' <-> compatible_ms PEa.
+Proof.
+split; [apply RgS_val_compatible_ms_rev | apply RgS_val_compatible_ms].
+Qed.
+
+End Sub_ModuleSpace_Facts2.
+
+
+Section Sub_ModuleSpace_Linear_Mapping_Facts1.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+Hypothesis HPE1 : compatible_ms PE1.
+Hypothesis HPE2 : compatible_ms PE2.
+
+Context {f : E1 -> E2}.
+Context {fS : sub_ms HPE1 -> sub_ms HPE2}.
+Hypothesis HfS : forall x, val (fS x) = f (val x).
+
+Lemma sub_ms_fun_rev : funS PE1 PE2 f.
+Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
+
+Lemma sub_ms_inj : injS PE1 f -> injective fS.
+Proof. apply sub_g_inj, HfS. Qed.
+
+Lemma sub_ms_inj_rev : injective fS -> injS PE1 f.
+Proof. apply sub_g_inj_rev, HfS. Qed.
+
+Lemma sub_ms_inj_equiv : injective fS <-> injS PE1 f.
+Proof. apply sub_g_inj_equiv, HfS. Qed.
+
+Lemma sub_ms_surj : surjS PE1 PE2 f -> surjective fS.
+Proof. apply sub_g_surj, HfS. Qed.
+
+Lemma sub_ms_surj_rev : surjective fS -> surjS PE1 PE2 f.
+Proof. apply sub_g_surj_rev, HfS. Qed.
+
+Lemma sub_ms_surj_equiv : surjective fS <-> surjS PE1 PE2 f.
+Proof. apply sub_g_surj_equiv, HfS. Qed.
+
+Lemma sub_ms_bij : bijS PE1 PE2 f -> bijective fS.
+Proof. apply sub_g_bij, HfS. Qed.
+
+Lemma sub_ms_bij_rev : bijective fS -> bijS PE1 PE2 f.
+Proof. apply sub_g_bij_rev, HfS. Qed.
+
+Lemma sub_ms_bij_equiv : bijective fS <-> bijS PE1 PE2 f.
+Proof. apply sub_g_bij_equiv, HfS. Qed.
+
+Lemma sub_ms_f_scal_compat : f_scal_compat f -> f_scal_compat fS.
+Proof. intros Hf a x1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
+
+Lemma sub_ms_f_comb_lin_compat : f_comb_lin_compat f -> f_comb_lin_compat fS.
+Proof.
+intros Hf n L B; apply val_inj; rewrite HfS !val_comb_lin Hf; f_equal.
+apply extF; intros i; rewrite !mapF_correct; easy.
+Qed.
+
+Lemma sub_ms_linear_mapping : is_linear_mapping f -> is_linear_mapping fS.
+Proof.
+intros [Hf1 Hf2]; split; [apply (sub_g_morphism HfS Hf1) |
+    move=>>; apply sub_ms_f_scal_compat; move=>>; easy].
+Qed.
+
+Lemma Ker_sub_ms_KerS_zero_equiv :
+  is_linear_mapping f ->
+  Ker fS = zero_sub_struct <-> KerS PE1 f = zero_sub_struct.
+Proof.
+intros Hf; apply (Ker_sub_g_KerS_zero_equiv HfS),
+    is_linear_mapping_morphism_g, Hf.
+Qed.
+
+Lemma KerS_ms_zero_equiv_alt :
+  is_linear_mapping f ->
+  KerS PE1 f = zero_sub_struct <->
+  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros H; apply: KerS_g_zero_equiv_alt; apply is_linear_mapping_morphism_g, H.
+Qed.
+
+Lemma Ker_sub_ms_zero_equiv :
+  is_linear_mapping f ->
+  Ker fS = zero_sub_struct <->
+  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros Hf; apply (Ker_sub_g_zero_equiv HfS), is_linear_mapping_morphism_g, Hf.
+Qed.
+
+Lemma lmS_injS_sub_equiv :
+  is_linear_mapping f -> injS PE1 f <-> Ker fS = zero_sub_struct.
+Proof.
+intros Hf; apply (gmS_injS_sub_equiv HfS), is_linear_mapping_morphism_g, Hf.
+Qed.
+
+Lemma lmS_injS_sub_equiv_alt :
+  is_linear_mapping f ->
+  injS PE1 f <-> forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros Hf; apply: (gmS_injS_sub_equiv_alt HfS);
+    apply is_linear_mapping_morphism_g, Hf.
+Qed.
+
+Lemma lmS_bijS_sub_equiv :
+  is_linear_mapping f -> bijS PE1 (RgS PE1 f) f <-> Ker fS = zero_sub_struct.
+Proof.
+intros Hf; apply (gmS_bijS_sub_equiv HfS), is_linear_mapping_morphism_g, Hf.
+Qed.
+
+Lemma lmS_bijS_sub_equiv_alt :
+  is_linear_mapping f ->
+  bijS PE1 (RgS PE1 f) f <-> forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
+Proof.
+intros Hf; apply: (gmS_bijS_sub_equiv_alt HfS);
+    apply is_linear_mapping_morphism_g, Hf.
+Qed.
+
+End Sub_ModuleSpace_Linear_Mapping_Facts1.
+
+
+Section Sub_ModuleSpace_Linear_Mapping_Facts2.
+
+Context {K : Ring}.
+Context {E1 E2 : ModuleSpace K}.
+
+Context {PE1 : E1 -> Prop}.
+Context {PE2 : E2 -> Prop}.
+Hypothesis HPE1 : compatible_ms PE1.
+Hypothesis HPE2 : compatible_ms PE2.
+
+Context {f : E1 -> E2}.
+Hypothesis Hf : funS PE1 PE2 f.
+
+Definition fct_sub_ms : sub_ms HPE1 -> sub_ms HPE2 :=
+  fct_sub_g (compatible_ms_g HPE1) (compatible_ms_g HPE2) Hf.
+
+Lemma fct_sub_ms_inj : injS PE1 f -> injective fct_sub_ms.
+Proof. apply fct_sub_g_inj. Qed.
+
+Lemma fct_sub_ms_inj_rev : injective fct_sub_ms -> injS PE1 f.
+Proof. apply fct_sub_g_inj_rev. Qed.
+
+Lemma fct_sub_ms_inj_equiv : injective fct_sub_ms <-> injS PE1 f.
+Proof. apply fct_sub_g_inj_equiv. Qed.
+
+Lemma fct_sub_ms_surj : surjS PE1 PE2 f -> surjective fct_sub_ms.
+Proof. apply fct_sub_g_surj. Qed.
+
+Lemma fct_sub_ms_surj_rev : surjective fct_sub_ms -> surjS PE1 PE2 f.
+Proof. apply fct_sub_g_surj_rev. Qed.
+
+Lemma fct_sub_ms_surj_equiv : surjective fct_sub_ms <-> surjS PE1 PE2 f.
+Proof. apply fct_sub_g_surj_equiv. Qed.
+
+Lemma fct_sub_ms_bij : bijS PE1 PE2 f -> bijective fct_sub_ms.
+Proof. apply fct_sub_g_bij. Qed.
+
+Lemma fct_sub_ms_bij_rev : bijective fct_sub_ms -> bijS PE1 PE2 f.
+Proof. apply fct_sub_g_bij_rev. Qed.
+
+Lemma fct_sub_ms_bij_equiv : bijective fct_sub_ms <-> bijS PE1 PE2 f.
+Proof. apply fct_sub_g_bij_equiv. Qed.
+
+Lemma fct_sub_ms_f_scal_compat : f_scal_compat f -> f_scal_compat fct_sub_ms.
+Proof. apply sub_ms_f_scal_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_ms_f_comb_lin_compat :
+  f_comb_lin_compat f -> f_comb_lin_compat fct_sub_ms.
+Proof. apply sub_ms_f_comb_lin_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_ms_linear_mapping :
+  is_linear_mapping f -> is_linear_mapping fct_sub_ms.
+Proof. apply sub_ms_linear_mapping, fct_sub_correct. Qed.
+
+Lemma Ker_fct_sub_ms_KerS_zero_equiv :
+  is_linear_mapping f ->
+  Ker fct_sub_ms = zero_sub_struct <-> KerS PE1 f = zero_sub_struct.
+Proof. apply Ker_sub_ms_KerS_zero_equiv, fct_sub_correct. Qed.
+
+Lemma Ker_fct_sub_ms_zero_equiv :
+  is_linear_mapping f ->
+  Ker fct_sub_ms = zero_sub_struct <->
+  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
+Proof. apply Ker_sub_ms_zero_equiv, fct_sub_correct. Qed.
+
+Lemma lmS_injS_fct_sub_equiv :
+  is_linear_mapping f -> injS PE1 f <-> Ker fct_sub_ms = zero_sub_struct.
+Proof. apply lmS_injS_sub_equiv, fct_sub_correct. Qed.
+
+Lemma lmS_bijS_fct_sub_equiv :
+  is_linear_mapping f ->
+  bijS PE1 (RgS PE1 f) f <-> Ker fct_sub_ms = zero_sub_struct.
+Proof. apply lmS_bijS_sub_equiv, fct_sub_correct. Qed.
+
+End Sub_ModuleSpace_Linear_Mapping_Facts2.
diff --git a/FEM/Algebra/MonoidComp.v b/FEM/Algebra/MonoidComp.v
index 4cb14ecef8e0b2a161c6de2311fa26078abbbc74..b575712692dc4736b7bc12a986a97332d7bd4322 100644
--- a/FEM/Algebra/MonoidComp.v
+++ b/FEM/Algebra/MonoidComp.v
@@ -27,7 +27,7 @@ Set Warnings "-notation-overridden".
 From Lebesgue Require Import logic_compl Function Subset.
 
 From FEM Require Import Compl nat_compl ord_compl Finite_family.
-From FEM Require Import Monoid_compl Group_compl ModuleSpace_compl Sub_struct.
+From FEM Require Import Sub_struct Monoid_compl Group_compl ModuleSpace_compl.
 
 (** Support for compositional abelian monoids. *)
 
diff --git a/FEM/Algebra/Monoid_compl.v b/FEM/Algebra/Monoid_compl.v
index 970a19944ffff14576d00892e607f2205c4b02dd..0c07b3de4376e34c9d030071af119186514bd6d8 100644
--- a/FEM/Algebra/Monoid_compl.v
+++ b/FEM/Algebra/Monoid_compl.v
@@ -25,9 +25,10 @@ From LM Require Import linear_map.
 Close Scope R_scope.
 Set Warnings "notation-overridden".
 
-From Lebesgue Require Import Subset Function.
+From Lebesgue Require Import Subset Subset_any Function.
 
 From FEM Require Import Compl nat_compl ord_compl Finite_family Finite_table.
+From FEM Require Import Sub_struct.
 
 
 Declare Scope Monoid_scope.
@@ -3293,3 +3294,756 @@ rewrite concatnF_ind_l sum_castF sum_concatF sum_ind_l IHn; easy.
 Qed.
 
 End ConcatnF_Monoid_Facts.
+
+
+Section Compatible_Monoid_Def.
+
+Context {G : AbelianMonoid}.
+
+Definition plus_closed (PG : G -> Prop) : Prop :=
+  forall x y, PG x -> PG y -> PG (x + y).
+
+Definition zero_closed (PG : G -> Prop) : Prop := PG 0.
+
+Definition sum_closed (PG : G -> Prop) : Prop :=
+  forall n (u : 'G^n), inclF u PG -> PG (sum u).
+
+Definition compatible_m (PG : G -> Prop) : Prop :=
+  plus_closed PG /\ zero_closed PG.
+
+End Compatible_Monoid_Def.
+
+
+Section Compatible_Monoid_Facts.
+
+Context {G : AbelianMonoid}.
+
+Lemma plus_zero_sum_closed :
+  forall {PG : G -> Prop},
+    plus_closed PG -> zero_closed PG -> sum_closed PG.
+Proof.
+intros PG HPG1 HPG2 n u Hu; induction n as [| n Hn].
+rewrite sum_nil; easy.
+rewrite sum_ind_l; apply HPG1; try easy; apply Hn; intro; apply Hu.
+Qed.
+
+Lemma sum_plus_zero_closed :
+  forall {PG : G -> Prop},
+    sum_closed PG -> plus_closed PG /\ zero_closed PG.
+Proof.
+intros PG HPG; split.
+(* *)
+intros u v Hu Hv; replace (u + v) with (sum (coupleF u v)).
+apply HPG; intros i; destruct (ord2_dec i);
+    [rewrite coupleF_l | rewrite coupleF_r]; easy.
+rewrite sum_coupleF; easy.
+(* *)
+unfold zero_closed; rewrite -(sum_nil 0); apply HPG; intros [i Hi]; easy.
+Qed.
+
+Lemma sum_plus_zero_closed_equiv :
+  forall {PG : G -> Prop},
+    sum_closed PG <-> plus_closed PG /\ zero_closed PG.
+Proof.
+intros; split.
+apply sum_plus_zero_closed.
+intros; apply plus_zero_sum_closed; easy.
+Qed.
+
+Lemma sum_closed_compatible_m :
+  forall {PG : G -> Prop}, sum_closed PG -> compatible_m PG.
+Proof. move=>> /sum_plus_zero_closed_equiv //. Qed.
+
+Lemma compatible_m_nonempty :
+  forall {PG : G -> Prop}, compatible_m PG -> nonempty PG.
+Proof. move=>> H; exists 0; apply H. Qed.
+
+Lemma compatible_m_zero :
+  forall {PG : G -> Prop}, compatible_m PG -> zero_closed PG.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_m_plus :
+  forall {PG : G -> Prop}, compatible_m PG -> plus_closed PG.
+Proof. move=>> H; apply H. Qed.
+
+Lemma compatible_m_sum :
+  forall {PG : G -> Prop}, compatible_m PG -> sum_closed PG.
+Proof.
+intros; apply plus_zero_sum_closed;
+    [apply compatible_m_plus | apply compatible_m_zero]; easy.
+Qed.
+
+Definition zero_sub_struct : G -> Prop := singleton 0.
+
+Lemma zero_sub_struct_correct :
+  forall {PG : G -> Prop}, PG = zero_sub_struct -> forall x, PG x -> x = 0.
+Proof. move=>> -> //. Qed.
+
+Lemma nonzero_sub_struct_ex :
+  forall {PG : G -> Prop},
+    nonempty PG -> PG <> zero_sub_struct <-> exists x, PG x /\ x <> 0.
+Proof.
+intros PG [x0 H0]; rewrite -iff_not_r_equiv not_ex_all_not_equiv; split.
+intros H x; rewrite H; easy.
+intros H; apply subset_ext; intros x; split; intros Hx.
+specialize (H x); rewrite -imp_and_equiv in H; apply H; easy.
+specialize (H x0); rewrite -imp_and_equiv in H; rewrite Hx -(H H0); easy.
+Qed.
+
+Lemma compatible_m_zero_sub_struct : compatible_m zero_sub_struct.
+Proof.
+split; try easy; intros x y Hx Hy; rewrite Hx Hy plus_zero_r; easy.
+Qed.
+
+Lemma compatible_m_full : forall {PG : G -> Prop}, full PG -> compatible_m PG.
+Proof. intros; split; try unfold zero_closed; easy. Qed.
+
+Lemma compatible_m_fullset : compatible_m (@fullset G).
+Proof. easy. Qed.
+
+Lemma compatible_m_inter :
+  forall {PG QG : G -> Prop},
+    compatible_m PG -> compatible_m QG -> compatible_m (inter PG QG).
+Proof.
+intros PG QG HPG HQG; split.
+destruct HPG as [HPG _], HQG as [HQG _];
+    intros x y [Hx1 Hx2] [Hy1 Hy2]; split; auto.
+split; [apply HPG | apply HQG].
+Qed.
+
+Lemma compatible_m_inter_any :
+  forall {Idx : Type} {PG : Idx -> G -> Prop},
+    (forall i, compatible_m (PG i)) -> compatible_m (inter_any PG).
+Proof.
+intros Idx PG HPG; split.
+intros x y Hx Hy i; apply HPG; easy.
+intros i; apply compatible_m_zero; easy.
+Qed.
+
+Definition span_m (gen : G -> Prop) := span_struct compatible_m gen.
+
+Lemma span_m_compatible_m : forall gen, compatible_m (span_m gen).
+Proof.
+apply span_struct_compatible; move=>>; apply compatible_m_inter_any.
+Qed.
+
+Lemma span_m_nonempty : forall gen, nonempty (span_m gen).
+Proof. intros; apply compatible_m_nonempty, span_m_compatible_m. Qed.
+
+Lemma span_m_zero : forall gen, zero_closed (span_m gen).
+Proof. intros; apply compatible_m_zero, span_m_compatible_m. Qed.
+
+Lemma span_m_plus : forall gen, plus_closed (span_m gen).
+Proof. intros; apply compatible_m_plus, span_m_compatible_m. Qed.
+
+Lemma span_m_sum : forall gen, sum_closed (span_m gen).
+Proof. intros; apply compatible_m_sum, span_m_compatible_m. Qed.
+
+Lemma span_m_incl : forall gen, incl gen (span_m gen).
+Proof. apply span_struct_incl. Qed.
+
+Lemma span_m_lub :
+  forall gen PG, compatible_m PG -> incl gen PG -> incl (span_m gen) PG.
+Proof. apply span_struct_lub. Qed.
+
+Lemma span_m_full : forall PG, compatible_m PG -> span_m PG = PG.
+Proof. apply span_struct_full. Qed.
+
+End Compatible_Monoid_Facts.
+
+
+Section Compatible_Monoid_Morphism_Facts1.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_m f.
+
+Lemma compatible_m_image : compatible_m PG1 -> compatible_m (image f PG1).
+Proof.
+destruct Hf as [Hfa Hfb]; intros [HPG1a HPG1b]; split.
+intros _ _ [x1 Hx1] [y1 Hy1]; rewrite -Hfa; apply Im; auto.
+unfold zero_closed; rewrite -Hfb; apply Im; easy.
+Qed.
+
+Lemma compatible_m_preimage : compatible_m PG2 -> compatible_m (preimage f PG2).
+Proof.
+destruct Hf as [Hfa Hfb]; intros [HPG2a HPG2b]; split; unfold preimage.
+do 2 intro; rewrite Hfa; auto.
+unfold zero_closed; rewrite Hfb; easy.
+Qed.
+
+Lemma compatible_m_morphism_m : compatible_m (@morphism_m G1 G2).
+Proof. split; [apply: morphism_m_fct_plus | apply morphism_m_fct_zero]. Qed.
+
+End Compatible_Monoid_Morphism_Facts1.
+
+
+Section Compatible_Monoid_Morphism_Def.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Variable PG1 : G1 -> Prop.
+
+Variable f : G1 -> G2.
+
+Definition Ker : G1 -> Prop := preimage f zero_sub_struct.
+Definition KerS : G1 -> Prop := inter PG1 Ker.
+
+Definition f_plus_compat_sub : Prop :=
+  forall x1 y1, PG1 x1 -> PG1 y1 -> f (x1 + y1) = f x1 + f y1.
+
+Definition f_zero_compat_sub : Prop := PG1 0 -> f 0 = 0.
+
+Definition morphism_m_sub : Prop := f_plus_compat_sub /\ f_zero_compat_sub.
+
+End Compatible_Monoid_Morphism_Def.
+
+
+Section Compatible_Monoid_Morphism_Facts2.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+
+Context {f : G1 -> G2}.
+
+Lemma Ker_correct : forall {x1}, f x1 = 0 -> Ker f x1.
+Proof. easy. Qed.
+
+Lemma Ker_equiv : forall {x1}, Ker f x1 <-> f x1 = 0.
+Proof. easy. Qed.
+
+Lemma KerS_correct : forall {x1}, PG1 x1 -> f x1 = 0 -> KerS PG1 f x1.
+Proof. easy. Qed.
+
+Lemma KerS_equiv : forall {x1}, KerS PG1 f x1 <-> PG1 x1 /\ f x1 = 0.
+Proof. easy. Qed.
+
+Lemma KerS_incl :
+  forall (PG1 : G1 -> Prop) (f : G1 -> G2), incl (KerS PG1 f) PG1.
+Proof. move=>> [Hx _]; easy. Qed.
+
+Lemma KerS_full : KerS fullset f = Ker f.
+Proof. apply inter_full_l. Qed.
+
+End Compatible_Monoid_Morphism_Facts2.
+
+
+Section Compatible_Monoid_Morphism_Facts3a.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Hypothesis HPG1 : compatible_m PG1.
+
+Context {PG2 : G2 -> Prop}.
+Hypothesis HPG2 : compatible_m PG2.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_m f.
+
+Lemma KerS_compatible_m : compatible_m (KerS PG1 f).
+Proof.
+apply compatible_m_inter, compatible_m_preimage,
+    compatible_m_zero_sub_struct; easy.
+Qed.
+
+Lemma RgS_gen_compatible_m : compatible_m (RgS_gen PG1 PG2 f).
+Proof. apply compatible_m_inter, compatible_m_image; easy. Qed.
+
+Lemma RgS_compatible_m : compatible_m (RgS PG1 f).
+Proof. apply compatible_m_image; easy. Qed.
+
+Lemma KerS_m_zero_equiv :
+  KerS PG1 f = zero_sub_struct <-> incl (KerS PG1 f) zero_sub_struct.
+Proof.
+split; intros Hf1; [rewrite Hf1 | apply incl_antisym]; try easy.
+move=>> ->; split; [apply compatible_m_zero | apply morphism_m_zero]; easy.
+Qed.
+
+Lemma mmS_injS_rev : injS PG1 f -> KerS PG1 f = zero_sub_struct.
+Proof.
+rewrite KerS_m_zero_equiv//; intros Hf1 x1 [Hx1a Hx1b]; apply Hf1;
+  [| apply compatible_m_zero | rewrite morphism_m_zero]; easy.
+Qed.
+
+Lemma mmS_bijS_gen_rev :
+  bijS PG1 PG2 f ->
+  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
+Proof.
+rewrite bijS_equiv; [| apply inhabited_m].
+assert (H0 : forall P Q1 Q2 R1 R2 : Prop,
+    (Q1 -> R1) -> (Q2 -> R2) -> P /\ Q1 /\ Q2 -> P /\ R1 /\ R2) by tauto.
+apply H0; [apply mmS_injS_rev | apply surjS_rev].
+Qed.
+
+Lemma mmS_bijS_rev : bijS PG1 (RgS PG1 f) f -> KerS PG1 f = zero_sub_struct.
+Proof.
+rewrite bijS_equiv; [intros; apply mmS_injS_rev; easy | apply inhabited_m].
+Qed.
+
+Lemma mmS_bijS_injS_equiv : bijS PG1 (RgS PG1 f) f <-> injS PG1 f.
+Proof. apply bijS_RgS_equiv, inhabited_m. Qed.
+
+End Compatible_Monoid_Morphism_Facts3a.
+
+
+Section Compatible_Monoid_Morphism_Facts3b.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : morphism_m f.
+
+Lemma Ker_compatible_m : compatible_m (Ker f).
+Proof.
+rewrite -KerS_full; apply (KerS_compatible_m compatible_m_fullset Hf).
+Qed.
+
+Lemma Rg_compatible_m : compatible_m (Rg f).
+Proof.
+rewrite -RgS_full; apply (RgS_compatible_m compatible_m_fullset Hf).
+Qed.
+
+Lemma Ker_m_zero_equiv :
+  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
+Proof.
+rewrite -KerS_full; apply (KerS_m_zero_equiv compatible_m_fullset Hf).
+Qed.
+
+Lemma mm_inj_rev : injective f -> Ker f = zero_sub_struct.
+Proof.
+rewrite -KerS_full inj_S_equiv; apply (mmS_injS_rev compatible_m_fullset Hf).
+Qed.
+
+Lemma mm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
+Proof.
+rewrite -KerS_full -RgS_gen_full bij_S_equiv; [| apply inhabited_m].
+apply (mmS_bijS_gen_rev compatible_m_fullset Hf).
+Qed.
+
+End Compatible_Monoid_Morphism_Facts3b.
+
+
+Section Compatible_Monoid_Morphism_Facts4a.
+
+Context {E1 E2 E3 : AbelianMonoid}.
+Context {PE2 : E2 -> Prop}.
+Context {f : E1 -> E2}.
+Context {g : E2 -> E3}.
+
+Lemma KerS_comp : KerS (preimage f PE2) (g \o f) = preimage f (KerS PE2 g).
+Proof. easy. Qed.
+
+End Compatible_Monoid_Morphism_Facts4a.
+
+
+Section Compatible_Monoid_Morphism_Facts4b.
+
+Context {E1 E2 E3 : AbelianMonoid}.
+Context {f : E1 -> E2}.
+Context {g : E2 -> E3}.
+
+Lemma Ker_comp : Ker (g \o f) = preimage f (Ker g).
+Proof. easy. Qed.
+
+Lemma Ker_comp_l : surjective f -> RgS (Ker (g \o f)) f = Ker g.
+Proof.
+intros; rewrite Ker_comp RgS_preimage; apply inter_left.
+rewrite surj_rev; easy.
+Qed.
+
+Lemma Ker_comp_l_alt :
+  forall h x2, surjective f -> h = g \o f ->
+    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
+Proof. move=>> Hf1 ->; rewrite -RgS_ex Ker_comp_l; easy. Qed.
+
+Hypothesis Hf : morphism_m f.
+Hypothesis Hg : morphism_m g.
+
+Lemma Ker_m_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
+Proof.
+intros; rewrite Ker_comp RgS_preimage; apply inter_left.
+rewrite mm_inj_rev//; move=>> ->; rewrite -(morphism_m_zero f); easy.
+Qed.
+
+Lemma Ker_m_comp_r_alt :
+  forall h x2, injective g -> h = g \o f ->
+    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
+Proof. move=>> Hg1 ->; rewrite -RgS_ex Ker_m_comp_r; easy. Qed.
+
+End Compatible_Monoid_Morphism_Facts4b.
+
+
+Section Compatible_Monoid_Morphism_Facts5a.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Variable PG1 : G1 -> Prop.
+
+Lemma f_plus_compat_sub_id : f_plus_compat_sub PG1 ssrfun.id.
+Proof. easy. Qed.
+
+Lemma f_zero_compat_sub_id : f_zero_compat_sub PG1 ssrfun.id.
+Proof. easy. Qed.
+
+Lemma morphism_m_sub_id : morphism_m_sub PG1 ssrfun.id.
+Proof. easy. Qed.
+
+Variable f : G1 -> G2.
+
+Lemma f_plus_compat_is_sub : f_plus_compat f -> f_plus_compat_sub PG1 f.
+Proof. easy. Qed.
+
+Lemma f_zero_compat_is_sub : f_zero_compat f -> f_zero_compat_sub PG1 f.
+Proof. easy. Qed.
+
+Lemma morphism_m_is_sub : morphism_m f -> morphism_m_sub PG1 f.
+Proof. intros [Hfa Hfb]; easy. Qed.
+
+End Compatible_Monoid_Morphism_Facts5a.
+
+
+Section Compatible_Monoid_Morphism_Facts5b.
+
+Context {G1 G2 G3 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Variable PG2 : G2 -> Prop.
+
+Variable f : G1 -> G2.
+Variable g : G2 -> G3.
+
+Hypothesis Hf : funS PG1 PG2 f.
+
+Lemma f_plus_compat_comp_sub :
+  f_plus_compat_sub PG1 f -> f_plus_compat_sub PG2 g ->
+  f_plus_compat_sub PG1 (g \o f).
+Proof.
+intros Hf1 Hg1 x1 y1 Hx1 Hy1;
+    rewrite !comp_correct Hf1// Hg1//; apply Hf; easy.
+Qed.
+
+Lemma f_zero_compat_comp_sub :
+  f_zero_compat_sub PG1 f -> f_zero_compat_sub PG2 g ->
+  f_zero_compat_sub PG1 (g \o f).
+Proof.
+intros Hf1 Hg1 H0;
+    rewrite comp_correct Hf1// Hg1// -Hf1//; apply Hf; easy.
+Qed.
+
+Lemma morphism_m_comp_sub :
+  morphism_m_sub PG1 f -> morphism_m_sub PG2 g ->
+  morphism_m_sub PG1 (g \o f).
+Proof.
+intros [Hfa Hfb] [Hga Hgb]; split;
+    [apply f_plus_compat_comp_sub | apply f_zero_compat_comp_sub]; easy.
+Qed.
+
+End Compatible_Monoid_Morphism_Facts5b.
+
+
+Section Sub_Monoid_Def.
+
+Context {G : AbelianMonoid}.
+Context {PG : G -> Prop}.
+Hypothesis HPG : compatible_m PG.
+
+Definition sub_plus (x y : sub_struct HPG) : sub_struct HPG :=
+  mk_sub (compatible_m_plus HPG _ _ (in_sub x) (in_sub y)).
+
+Definition sub_zero : sub_struct HPG := mk_sub_ HPG _ (compatible_m_zero HPG).
+
+Lemma sub_plus_comm : forall x y, sub_plus x y = sub_plus y x.
+Proof. intros; apply val_inj, plus_comm. Qed.
+
+Lemma sub_plus_assoc:
+  forall x y z, sub_plus x (sub_plus y z) = sub_plus (sub_plus x y) z.
+Proof. intros; apply val_inj, plus_assoc. Qed.
+
+Lemma sub_plus_zero_r : forall x, sub_plus x sub_zero = x.
+Proof. intros; apply val_inj, plus_zero_r. Qed.
+
+Definition sub_AbelianMonoid_mixin :=
+  AbelianMonoid.Mixin _ _ _ sub_plus_comm sub_plus_assoc sub_plus_zero_r.
+
+Canonical Structure sub_AbelianMonoid :=
+  AbelianMonoid.Pack _ sub_AbelianMonoid_mixin (sub_struct HPG).
+
+Lemma val_zero : f_zero_compat val.
+Proof. easy. Qed.
+
+Lemma val_zero_equiv : forall (x : sub_struct HPG), val x = 0 <-> x = 0.
+Proof.
+intros; split; [rewrite -val_zero; apply val_inj | intros; subst; easy].
+Qed.
+
+Lemma val_plus : f_plus_compat val.
+Proof. easy. Qed.
+
+Lemma val_sum : f_sum_compat val.
+Proof.
+intros n; induction n as [| n Hn].
+intros; rewrite 2!sum_nil; easy.
+intros; rewrite 2!sum_ind_l; simpl; f_equal; apply Hn.
+Qed.
+
+Lemma val_morphism_m : morphism_m val.
+Proof. easy. Qed.
+
+Lemma mk_sub_zero : sub_zero = (0 : sub_struct HPG).
+Proof. easy. Qed.
+
+Lemma mk_sub_zero_equiv : forall {x} (Hx : PG x), mk_sub Hx = 0 <-> x = 0.
+Proof. intros; split; [move=> /val_eq | intros; apply val_inj]; easy. Qed.
+
+Lemma mk_sub_plus :
+  forall (x y : G) (Hx : PG x) (Hy : PG y),
+    mk_sub Hx + mk_sub Hy = mk_sub (compatible_m_plus HPG _ _ Hx Hy).
+Proof. easy. Qed.
+
+Lemma mk_sub_sum :
+  forall {n} (u : 'G^n) (Hu : inclF u PG),
+    sum (fun i => mk_sub (Hu i)) = mk_sub (compatible_m_sum HPG _ _ Hu).
+Proof. intros; apply val_inj, val_sum. Qed.
+
+Lemma sub_zero_eq :
+  (0 : sub_struct HPG) = mk_sub_ _ (0 : G) (compatible_m_zero HPG).
+Proof. easy. Qed.
+
+Lemma sub_plus_eq :
+  forall (x y : sub_struct HPG),
+    x + y = mk_sub (compatible_m_plus HPG _ _ (in_sub x) (in_sub y)).
+Proof. intros; apply val_inj; easy. Qed.
+
+Lemma sub_sum_eq :
+  forall {n} (u : '(sub_struct HPG)^n),
+    sum u = mk_sub (compatible_m_sum HPG _ _ (fun i => in_sub (u i))).
+Proof. intros; apply val_inj, val_sum. Qed.
+
+End Sub_Monoid_Def.
+
+
+Section Sub_Monoid_Facts1.
+
+Context {G : AbelianMonoid}.
+Context {PGa PGb : G -> Prop}.
+Hypothesis HPG : incl PGa PGb.
+Hypothesis HPGb : compatible_m PGb.
+Let PGb_m := sub_struct HPGb.
+Let PGa' : PGb_m -> Prop := preimage val PGa.
+
+Lemma RgS_val_m_eq : RgS PGa' val = PGa.
+Proof.
+apply RgS_preimage_equiv; intros x Hx; apply Rg_ex;
+    exists (mk_sub (HPG _ Hx)); easy.
+Qed.
+
+Lemma preimage_val_compatible_m : compatible_m PGa -> compatible_m PGa'.
+Proof. intros; apply compatible_m_preimage; easy. Qed.
+
+Lemma preimage_val_compatible_m_rev : compatible_m PGa' -> compatible_m PGa.
+Proof. intros; rewrite -RgS_val_m_eq; apply RgS_compatible_m; easy. Qed.
+
+Lemma preimage_val_compatible_m_equiv : compatible_m PGa' <-> compatible_m PGa.
+Proof.
+split; [apply preimage_val_compatible_m_rev | apply preimage_val_compatible_m].
+Qed.
+
+End Sub_Monoid_Facts1.
+
+
+Section Sub_Monoid_Facts2.
+
+Context {G : AbelianMonoid}.
+Context {PGb : G -> Prop}.
+Hypothesis HPGb : compatible_m PGb.
+Let PGb_m := sub_struct HPGb.
+Context {PGa : PGb_m -> Prop}.
+Let PGa' := RgS PGa val.
+
+Lemma preimage_val_m_eq : preimage val PGa' = PGa.
+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_m : compatible_m PGa -> compatible_m PGa'.
+Proof. intros; apply RgS_compatible_m; easy. Qed.
+
+Lemma RgS_val_compatible_m_rev : compatible_m PGa' -> compatible_m PGa.
+Proof.
+intros; rewrite -preimage_val_m_eq; apply compatible_m_preimage; easy.
+Qed.
+
+Lemma RgS_val_compatible_m_equiv : compatible_m PGa' <-> compatible_m PGa.
+Proof.
+split; [apply RgS_val_compatible_m_rev | apply RgS_val_compatible_m].
+Qed.
+
+End Sub_Monoid_Facts2.
+
+
+Section Sub_Monoid_Morphism_Facts2a.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+Context {HPG1 : compatible_m PG1}.
+Context {HPG2 : compatible_m PG2}.
+
+Context {f : G1 -> G2}.
+Context {fS : sub_struct HPG1 -> sub_struct HPG2}.
+Hypothesis HfS : forall x, val (fS x) = f (val x).
+
+Lemma sub_m_fun_rev : funS PG1 PG2 f.
+Proof. apply (sub_fun_rev HPG1 HPG2 HfS). Qed.
+
+Lemma sub_m_inj : injS PG1 f -> injective fS.
+Proof. apply sub_inj, HfS. Qed.
+
+Lemma sub_m_inj_rev : injective fS -> injS PG1 f.
+Proof. apply sub_inj_rev, HfS. Qed.
+
+Lemma sub_m_inj_equiv : injective fS <-> injS PG1 f.
+Proof. apply sub_inj_equiv, HfS. Qed.
+
+Lemma sub_m_surj : surjS PG1 PG2 f -> surjective fS.
+Proof. apply sub_surj, HfS. Qed.
+
+Lemma sub_m_surj_rev : surjective fS -> surjS PG1 PG2 f.
+Proof. apply sub_surj_rev, HfS. Qed.
+
+Lemma sub_m_surj_equiv : surjective fS <-> surjS PG1 PG2 f.
+Proof. apply sub_surj_equiv, HfS. Qed.
+
+Lemma sub_m_bij : bijS PG1 PG2 f -> bijective fS.
+Proof. apply sub_bij, HfS; apply inhabited_m. Qed.
+
+Lemma sub_m_bij_rev : bijective fS -> bijS PG1 PG2 f.
+Proof. apply sub_bij_rev, HfS; apply inhabited_m. Qed.
+
+Lemma sub_m_bij_equiv : bijective fS <-> bijS PG1 PG2 f.
+Proof. apply sub_bij_equiv, HfS; apply inhabited_m. Qed.
+
+Lemma sub_m_f_plus_compat : f_plus_compat f -> f_plus_compat fS.
+Proof. intros Hf x1 y1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
+
+Lemma sub_m_f_zero_compat : f_zero_compat f -> f_zero_compat fS.
+Proof. intros Hf; apply val_inj; rewrite HfS Hf; easy. Qed.
+
+Lemma sub_m_f_sum_compat : f_sum_compat f -> f_sum_compat fS.
+Proof.
+intros Hf n u1; apply val_inj; rewrite HfS !val_sum Hf; f_equal.
+apply extF; intros i; rewrite !mapF_correct; easy.
+Qed.
+
+Lemma sub_m_morphism : morphism_m f -> morphism_m fS.
+Proof.
+intros [Hf1 Hf2]; split;
+    [apply (sub_m_f_plus_compat Hf1) | apply (sub_m_f_zero_compat Hf2)].
+Qed.
+
+Lemma Ker_sub_KerS_equiv :
+  forall (x1 : sub_struct HPG1), Ker fS x1 <-> KerS PG1 f (val x1).
+Proof.
+intros; rewrite Ker_equiv KerS_equiv -HfS; split; intros Hx1.
+split; [apply in_sub | apply val_zero_equiv; easy].
+apply val_inj; easy.
+Qed.
+
+Lemma KerS_Ker_sub_equiv :
+  forall {x1} (Hx1 : PG1 x1), KerS PG1 f x1 <-> Ker fS (mk_sub Hx1).
+Proof. intros; rewrite Ker_sub_KerS_equiv; easy. Qed.
+
+Lemma Ker_sub_zero_equiv :
+  morphism_m f ->
+  Ker fS = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
+Proof.
+intros Hf; split; intros H; apply subset_ext_equiv; split.
+(* *)
+intros x1; destruct (classic_dec (PG1 x1)) as [Hx1 | Hx1].
+move=> /(KerS_Ker_sub_equiv Hx1); rewrite H; move=> /mk_sub_zero_equiv; easy.
+intros [Hx1a _]; easy.
+(* *)
+move=> x1 ->; apply compatible_m_zero, KerS_compatible_m; easy.
+move=> x1 /Ker_sub_KerS_equiv Hx1; rewrite H in Hx1; apply val_inj; easy.
+move=> x1 ->; apply (compatible_m_zero (Ker_compatible_m (sub_m_morphism Hf))).
+Qed.
+
+Lemma Rg_sub_RgS_equiv :
+  forall (x2 : sub_struct HPG2), Rg fS x2 <-> RgS PG1 f (val x2).
+Proof.
+intros x2; split; [intros [x1 Hx1]; rewrite HfS; apply Im, in_sub |].
+intros Hx2; inversion Hx2 as [x1 Hx1a Hx1b].
+apply Rg_ex; exists (mk_sub Hx1a); apply val_inj; rewrite HfS; easy.
+Qed.
+
+Lemma RgS_Rg_sub_equiv :
+  forall {x2} (Hx2 : PG2 x2), RgS PG1 f x2 <-> Rg fS (mk_sub Hx2).
+Proof. intros; rewrite Rg_sub_RgS_equiv; easy. Qed.
+
+End Sub_Monoid_Morphism_Facts2a.
+
+
+Section Sub_Monoid_Morphism_Facts2b.
+
+Context {G1 G2 : AbelianMonoid}.
+
+Context {PG1 : G1 -> Prop}.
+Context {PG2 : G2 -> Prop}.
+Hypothesis HPG1 : compatible_m PG1.
+Hypothesis HPG2 : compatible_m PG2.
+
+Context {f : G1 -> G2}.
+Hypothesis Hf : funS PG1 PG2 f.
+
+Definition fct_sub_m : sub_struct HPG1 -> sub_struct HPG2 :=
+  fct_sub HPG1 HPG2 Hf.
+
+Lemma fct_sub_m_inj : injS PG1 f -> injective fct_sub_m.
+Proof. apply fct_sub_inj. Qed.
+
+Lemma fct_sub_m_inj_rev : injective fct_sub_m -> injS PG1 f.
+Proof. apply fct_sub_inj_rev. Qed.
+
+Lemma fct_sub_m_inj_equiv : injective fct_sub_m <-> injS PG1 f.
+Proof. apply fct_sub_inj_equiv. Qed.
+
+Lemma fct_sub_m_surj : surjS PG1 PG2 f -> surjective fct_sub_m.
+Proof. apply fct_sub_surj. Qed.
+
+Lemma fct_sub_m_surj_rev : surjective fct_sub_m -> surjS PG1 PG2 f.
+Proof. apply fct_sub_surj_rev. Qed.
+
+Lemma fct_sub_m_surj_equiv : surjective fct_sub_m <-> surjS PG1 PG2 f.
+Proof. apply fct_sub_surj_equiv. Qed.
+
+Lemma fct_sub_m_bij : bijS PG1 PG2 f -> bijective fct_sub_m.
+Proof. apply fct_sub_bij, inhabited_m. Qed.
+
+Lemma fct_sub_m_bij_rev : bijective fct_sub_m -> bijS PG1 PG2 f.
+Proof. apply fct_sub_bij_rev, inhabited_m. Qed.
+
+Lemma fct_sub_m_bij_equiv : bijective fct_sub_m <-> bijS PG1 PG2 f.
+Proof. apply fct_sub_bij_equiv, inhabited_m. Qed.
+
+Lemma fct_sub_m_f_plus_compat : f_plus_compat f -> f_plus_compat fct_sub_m.
+Proof. apply sub_m_f_plus_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_m_f_zero_compat : f_zero_compat f -> f_zero_compat fct_sub_m.
+Proof. apply sub_m_f_zero_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_m_f_sum_compat : f_sum_compat f -> f_sum_compat fct_sub_m.
+Proof. apply sub_m_f_sum_compat, fct_sub_correct. Qed.
+
+Lemma fct_sub_m_morphism : morphism_m f -> morphism_m fct_sub_m.
+Proof. apply sub_m_morphism, fct_sub_correct. Qed.
+
+End Sub_Monoid_Morphism_Facts2b.
diff --git a/FEM/Algebra/Ring_compl.v b/FEM/Algebra/Ring_compl.v
index 7daed2f709c6b7f2c8741e510d649dfe64738cff..5c79c3aec060a6b67fd637594fa33347f0e803ac 100644
--- a/FEM/Algebra/Ring_compl.v
+++ b/FEM/Algebra/Ring_compl.v
@@ -24,13 +24,14 @@ From LM Require Import linear_map.
 Close Scope R_scope.
 Set Warnings "notation-overridden".
 
-From Lebesgue Require Import nat_compl Subset Subset_charac Function.
+From Lebesgue Require Import nat_compl Subset Subset_charac Subset_any Function.
 
 From FEM Require Import Compl nat_compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl Group_compl.
+From FEM Require Import Sub_struct Monoid_compl Group_compl.
 
 
-(** Results needing commutativity of multiplication are only stated on R. *)
+(** Results needing a commutative Ring are only stated for
+ the ring of real numbers R_Ring. *)
 
 
 Section Ring_Def0.
@@ -1140,3 +1141,361 @@ 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.
diff --git a/FEM/Algebra/Sub_struct.v b/FEM/Algebra/Sub_struct.v
index a2d8fc465d6219627b413db38796a59c31536261..f87848da7a6a89288bf88f84d99ce0a61e0c34a0 100644
--- a/FEM/Algebra/Sub_struct.v
+++ b/FEM/Algebra/Sub_struct.v
@@ -24,20 +24,9 @@ Set Warnings "notation-overridden".
 From LM Require Import linear_map.
 Close Scope R_scope.
 
-From Lebesgue Require Import Subset Subset_dec Subset_any Function.
-From Lebesgue Require Import Subset_system_def Subset_system_base.
+From Lebesgue Require Import Subset Subset_any Function.
 
 From FEM Require Import Compl ord_compl Finite_family Finite_table.
-From FEM Require Import Monoid_compl Group_compl Ring_compl.
-From FEM Require Import ModuleSpace_compl ModuleSpace_R_compl AffineSpace.
-
-
-(** Results needing a commutative Ring are only stated for
- the ring of real numbers R_Ring. *)
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (05/03/2024):
- only keep here first sections on generic stuff, dispatch the rest in dedicated files. *)
 
 
 Section Sub_Algebraic_Structure_Def.
@@ -312,3898 +301,3 @@ Lemma fct_sub_bij_equiv : bijective fct_sub <-> bijS P1 P2 f.
 Proof. apply sub_bij_equiv, fct_sub_correct; easy. Qed.
 
 End Sub_Fct2.
-
-
-Local Open Scope Monoid_scope.
-
-
-Section Compatible_Monoid_Def.
-
-Context {G : AbelianMonoid}.
-
-Definition plus_closed (PG : G -> Prop) : Prop :=
-  forall x y, PG x -> PG y -> PG (x + y).
-
-Definition zero_closed (PG : G -> Prop) : Prop := PG 0.
-
-Definition sum_closed (PG : G -> Prop) : Prop :=
-  forall n (u : 'G^n), inclF u PG -> PG (sum u).
-
-Definition compatible_m (PG : G -> Prop) : Prop :=
-  plus_closed PG /\ zero_closed PG.
-
-End Compatible_Monoid_Def.
-
-
-Section Compatible_Monoid_Facts.
-
-Context {G : AbelianMonoid}.
-
-Lemma plus_zero_sum_closed :
-  forall {PG : G -> Prop},
-    plus_closed PG -> zero_closed PG -> sum_closed PG.
-Proof.
-intros PG HPG1 HPG2 n u Hu; induction n as [| n Hn].
-rewrite sum_nil; easy.
-rewrite sum_ind_l; apply HPG1; try easy; apply Hn; intro; apply Hu.
-Qed.
-
-Lemma sum_plus_zero_closed :
-  forall {PG : G -> Prop},
-    sum_closed PG -> plus_closed PG /\ zero_closed PG.
-Proof.
-intros PG HPG; split.
-(* *)
-intros u v Hu Hv; replace (u + v) with (sum (coupleF u v)).
-apply HPG; intros i; destruct (ord2_dec i);
-    [rewrite coupleF_l | rewrite coupleF_r]; easy.
-rewrite sum_coupleF; easy.
-(* *)
-unfold zero_closed; rewrite -(sum_nil 0); apply HPG; intros [i Hi]; easy.
-Qed.
-
-Lemma sum_plus_zero_closed_equiv :
-  forall {PG : G -> Prop},
-    sum_closed PG <-> plus_closed PG /\ zero_closed PG.
-Proof.
-intros; split.
-apply sum_plus_zero_closed.
-intros; apply plus_zero_sum_closed; easy.
-Qed.
-
-Lemma sum_closed_compatible_m :
-  forall {PG : G -> Prop}, sum_closed PG -> compatible_m PG.
-Proof. move=>> /sum_plus_zero_closed_equiv //. Qed.
-
-Lemma compatible_m_nonempty :
-  forall {PG : G -> Prop}, compatible_m PG -> nonempty PG.
-Proof. move=>> H; exists 0; apply H. Qed.
-
-Lemma compatible_m_zero :
-  forall {PG : G -> Prop}, compatible_m PG -> zero_closed PG.
-Proof. move=>> H; apply H. Qed.
-
-Lemma compatible_m_plus :
-  forall {PG : G -> Prop}, compatible_m PG -> plus_closed PG.
-Proof. move=>> H; apply H. Qed.
-
-Lemma compatible_m_sum :
-  forall {PG : G -> Prop}, compatible_m PG -> sum_closed PG.
-Proof.
-intros; apply plus_zero_sum_closed;
-    [apply compatible_m_plus | apply compatible_m_zero]; easy.
-Qed.
-
-Definition zero_sub_struct : G -> Prop := singleton 0.
-
-Lemma zero_sub_struct_correct :
-  forall {PG : G -> Prop}, PG = zero_sub_struct -> forall x, PG x -> x = 0.
-Proof. move=>> -> //. Qed.
-
-Lemma nonzero_sub_struct_ex :
-  forall {PG : G -> Prop},
-    nonempty PG -> PG <> zero_sub_struct <-> exists x, PG x /\ x <> 0.
-Proof.
-intros PG [x0 H0]; rewrite -iff_not_r_equiv not_ex_all_not_equiv; split.
-intros H x; rewrite H; easy.
-intros H; apply subset_ext; intros x; split; intros Hx.
-specialize (H x); rewrite -imp_and_equiv in H; apply H; easy.
-specialize (H x0); rewrite -imp_and_equiv in H; rewrite Hx -(H H0); easy.
-Qed.
-
-Lemma compatible_m_zero_sub_struct : compatible_m zero_sub_struct.
-Proof.
-split; try easy; intros x y Hx Hy; rewrite Hx Hy plus_zero_r; easy.
-Qed.
-
-Lemma compatible_m_full : forall {PG : G -> Prop}, full PG -> compatible_m PG.
-Proof. intros; split; try unfold zero_closed; easy. Qed.
-
-Lemma compatible_m_fullset : compatible_m (@fullset G).
-Proof. easy. Qed.
-
-Lemma compatible_m_inter :
-  forall {PG QG : G -> Prop},
-    compatible_m PG -> compatible_m QG -> compatible_m (inter PG QG).
-Proof.
-intros PG QG HPG HQG; split.
-destruct HPG as [HPG _], HQG as [HQG _];
-    intros x y [Hx1 Hx2] [Hy1 Hy2]; split; auto.
-split; [apply HPG | apply HQG].
-Qed.
-
-Lemma compatible_m_inter_any :
-  forall {Idx : Type} {PG : Idx -> G -> Prop},
-    (forall i, compatible_m (PG i)) -> compatible_m (inter_any PG).
-Proof.
-intros Idx PG HPG; split.
-intros x y Hx Hy i; apply HPG; easy.
-intros i; apply compatible_m_zero; easy.
-Qed.
-
-Definition span_m (gen : G -> Prop) := span_struct compatible_m gen.
-
-Lemma span_m_compatible_m : forall gen, compatible_m (span_m gen).
-Proof.
-apply span_struct_compatible; move=>>; apply compatible_m_inter_any.
-Qed.
-
-Lemma span_m_nonempty : forall gen, nonempty (span_m gen).
-Proof. intros; apply compatible_m_nonempty, span_m_compatible_m. Qed.
-
-Lemma span_m_zero : forall gen, zero_closed (span_m gen).
-Proof. intros; apply compatible_m_zero, span_m_compatible_m. Qed.
-
-Lemma span_m_plus : forall gen, plus_closed (span_m gen).
-Proof. intros; apply compatible_m_plus, span_m_compatible_m. Qed.
-
-Lemma span_m_sum : forall gen, sum_closed (span_m gen).
-Proof. intros; apply compatible_m_sum, span_m_compatible_m. Qed.
-
-Lemma span_m_incl : forall gen, incl gen (span_m gen).
-Proof. apply span_struct_incl. Qed.
-
-Lemma span_m_lub :
-  forall gen PG, compatible_m PG -> incl gen PG -> incl (span_m gen) PG.
-Proof. apply span_struct_lub. Qed.
-
-Lemma span_m_full : forall PG, compatible_m PG -> span_m PG = PG.
-Proof. apply span_struct_full. Qed.
-
-End Compatible_Monoid_Facts.
-
-
-Section Compatible_Monoid_Morphism_Facts1.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_m f.
-
-Lemma compatible_m_image : compatible_m PG1 -> compatible_m (image f PG1).
-Proof.
-destruct Hf as [Hfa Hfb]; intros [HPG1a HPG1b]; split.
-intros _ _ [x1 Hx1] [y1 Hy1]; rewrite -Hfa; apply Im; auto.
-unfold zero_closed; rewrite -Hfb; apply Im; easy.
-Qed.
-
-Lemma compatible_m_preimage : compatible_m PG2 -> compatible_m (preimage f PG2).
-Proof.
-destruct Hf as [Hfa Hfb]; intros [HPG2a HPG2b]; split; unfold preimage.
-do 2 intro; rewrite Hfa; auto.
-unfold zero_closed; rewrite Hfb; easy.
-Qed.
-
-Lemma compatible_m_morphism_m : compatible_m (@morphism_m G1 G2).
-Proof. split; [apply: morphism_m_fct_plus | apply morphism_m_fct_zero]. Qed.
-
-End Compatible_Monoid_Morphism_Facts1.
-
-
-Section Compatible_Monoid_Morphism_Def.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Variable PG1 : G1 -> Prop.
-
-Variable f : G1 -> G2.
-
-Definition Ker : G1 -> Prop := preimage f zero_sub_struct.
-Definition KerS : G1 -> Prop := inter PG1 Ker.
-
-Definition f_plus_compat_sub : Prop :=
-  forall x1 y1, PG1 x1 -> PG1 y1 -> f (x1 + y1) = f x1 + f y1.
-
-Definition f_zero_compat_sub : Prop := PG1 0 -> f 0 = 0.
-
-Definition morphism_m_sub : Prop := f_plus_compat_sub /\ f_zero_compat_sub.
-
-End Compatible_Monoid_Morphism_Def.
-
-
-Section Compatible_Monoid_Morphism_Facts2.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-
-Context {f : G1 -> G2}.
-
-Lemma Ker_correct : forall {x1}, f x1 = 0 -> Ker f x1.
-Proof. easy. Qed.
-
-Lemma Ker_equiv : forall {x1}, Ker f x1 <-> f x1 = 0.
-Proof. easy. Qed.
-
-Lemma KerS_correct : forall {x1}, PG1 x1 -> f x1 = 0 -> KerS PG1 f x1.
-Proof. easy. Qed.
-
-Lemma KerS_equiv : forall {x1}, KerS PG1 f x1 <-> PG1 x1 /\ f x1 = 0.
-Proof. easy. Qed.
-
-Lemma KerS_incl :
-  forall (PG1 : G1 -> Prop) (f : G1 -> G2), incl (KerS PG1 f) PG1.
-Proof. move=>> [Hx _]; easy. Qed.
-
-Lemma KerS_full : KerS fullset f = Ker f.
-Proof. apply inter_full_l. Qed.
-
-End Compatible_Monoid_Morphism_Facts2.
-
-
-Section Compatible_Monoid_Morphism_Facts3a.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Hypothesis HPG1 : compatible_m PG1.
-
-Context {PG2 : G2 -> Prop}.
-Hypothesis HPG2 : compatible_m PG2.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_m f.
-
-Lemma KerS_compatible_m : compatible_m (KerS PG1 f).
-Proof.
-apply compatible_m_inter, compatible_m_preimage,
-    compatible_m_zero_sub_struct; easy.
-Qed.
-
-Lemma RgS_gen_compatible_m : compatible_m (RgS_gen PG1 PG2 f).
-Proof. apply compatible_m_inter, compatible_m_image; easy. Qed.
-
-Lemma RgS_compatible_m : compatible_m (RgS PG1 f).
-Proof. apply compatible_m_image; easy. Qed.
-
-Lemma KerS_m_zero_equiv :
-  KerS PG1 f = zero_sub_struct <-> incl (KerS PG1 f) zero_sub_struct.
-Proof.
-split; intros Hf1; [rewrite Hf1 | apply incl_antisym]; try easy.
-move=>> ->; split; [apply compatible_m_zero | apply morphism_m_zero]; easy.
-Qed.
-
-Lemma mmS_injS_rev : injS PG1 f -> KerS PG1 f = zero_sub_struct.
-Proof.
-rewrite KerS_m_zero_equiv//; intros Hf1 x1 [Hx1a Hx1b]; apply Hf1;
-  [| apply compatible_m_zero | rewrite morphism_m_zero]; easy.
-Qed.
-
-Lemma mmS_bijS_gen_rev :
-  bijS PG1 PG2 f ->
-  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
-Proof.
-rewrite bijS_equiv; [| apply inhabited_m].
-assert (H0 : forall P Q1 Q2 R1 R2 : Prop,
-    (Q1 -> R1) -> (Q2 -> R2) -> P /\ Q1 /\ Q2 -> P /\ R1 /\ R2) by tauto.
-apply H0; [apply mmS_injS_rev | apply surjS_rev].
-Qed.
-
-Lemma mmS_bijS_rev : bijS PG1 (RgS PG1 f) f -> KerS PG1 f = zero_sub_struct.
-Proof.
-rewrite bijS_equiv; [intros; apply mmS_injS_rev; easy | apply inhabited_m].
-Qed.
-
-Lemma mmS_bijS_injS_equiv : bijS PG1 (RgS PG1 f) f <-> injS PG1 f.
-Proof. apply bijS_RgS_equiv, inhabited_m. Qed.
-
-End Compatible_Monoid_Morphism_Facts3a.
-
-
-Section Compatible_Monoid_Morphism_Facts3b.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_m f.
-
-Lemma Ker_compatible_m : compatible_m (Ker f).
-Proof.
-rewrite -KerS_full; apply (KerS_compatible_m compatible_m_fullset Hf).
-Qed.
-
-Lemma Rg_compatible_m : compatible_m (Rg f).
-Proof.
-rewrite -RgS_full; apply (RgS_compatible_m compatible_m_fullset Hf).
-Qed.
-
-Lemma Ker_m_zero_equiv :
-  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
-Proof.
-rewrite -KerS_full; apply (KerS_m_zero_equiv compatible_m_fullset Hf).
-Qed.
-
-Lemma mm_inj_rev : injective f -> Ker f = zero_sub_struct.
-Proof.
-rewrite -KerS_full inj_S_equiv; apply (mmS_injS_rev compatible_m_fullset Hf).
-Qed.
-
-Lemma mm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
-Proof.
-rewrite -KerS_full -RgS_gen_full bij_S_equiv; [| apply inhabited_m].
-apply (mmS_bijS_gen_rev compatible_m_fullset Hf).
-Qed.
-
-End Compatible_Monoid_Morphism_Facts3b.
-
-
-Section Compatible_Monoid_Morphism_Facts4a.
-
-Context {E1 E2 E3 : AbelianMonoid}.
-Context {PE2 : E2 -> Prop}.
-Context {f : E1 -> E2}.
-Context {g : E2 -> E3}.
-
-Lemma KerS_comp : KerS (preimage f PE2) (g \o f) = preimage f (KerS PE2 g).
-Proof. easy. Qed.
-
-End Compatible_Monoid_Morphism_Facts4a.
-
-
-Section Compatible_Monoid_Morphism_Facts4b.
-
-Context {E1 E2 E3 : AbelianMonoid}.
-Context {f : E1 -> E2}.
-Context {g : E2 -> E3}.
-
-Lemma Ker_comp : Ker (g \o f) = preimage f (Ker g).
-Proof. easy. Qed.
-
-Lemma Ker_comp_l : surjective f -> RgS (Ker (g \o f)) f = Ker g.
-Proof.
-intros; rewrite Ker_comp RgS_preimage; apply inter_left.
-rewrite surj_rev; easy.
-Qed.
-
-Lemma Ker_comp_l_alt :
-  forall h x2, surjective f -> h = g \o f ->
-    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
-Proof. move=>> Hf1 ->; rewrite -RgS_ex Ker_comp_l; easy. Qed.
-
-Hypothesis Hf : morphism_m f.
-Hypothesis Hg : morphism_m g.
-
-Lemma Ker_m_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
-Proof.
-intros; rewrite Ker_comp RgS_preimage; apply inter_left.
-rewrite mm_inj_rev//; move=>> ->; rewrite -(morphism_m_zero f); easy.
-Qed.
-
-Lemma Ker_m_comp_r_alt :
-  forall h x2, injective g -> h = g \o f ->
-    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
-Proof. move=>> Hg1 ->; rewrite -RgS_ex Ker_m_comp_r; easy. Qed.
-
-End Compatible_Monoid_Morphism_Facts4b.
-
-
-Section Compatible_Monoid_Morphism_Facts5a.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Variable PG1 : G1 -> Prop.
-
-Lemma f_plus_compat_sub_id : f_plus_compat_sub PG1 ssrfun.id.
-Proof. easy. Qed.
-
-Lemma f_zero_compat_sub_id : f_zero_compat_sub PG1 ssrfun.id.
-Proof. easy. Qed.
-
-Lemma morphism_m_sub_id : morphism_m_sub PG1 ssrfun.id.
-Proof. easy. Qed.
-
-Variable f : G1 -> G2.
-
-Lemma f_plus_compat_is_sub : f_plus_compat f -> f_plus_compat_sub PG1 f.
-Proof. easy. Qed.
-
-Lemma f_zero_compat_is_sub : f_zero_compat f -> f_zero_compat_sub PG1 f.
-Proof. easy. Qed.
-
-Lemma morphism_m_is_sub : morphism_m f -> morphism_m_sub PG1 f.
-Proof. intros [Hfa Hfb]; easy. Qed.
-
-End Compatible_Monoid_Morphism_Facts5a.
-
-
-Section Compatible_Monoid_Morphism_Facts5b.
-
-Context {G1 G2 G3 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Variable PG2 : G2 -> Prop.
-
-Variable f : G1 -> G2.
-Variable g : G2 -> G3.
-
-Hypothesis Hf : funS PG1 PG2 f.
-
-Lemma f_plus_compat_comp_sub :
-  f_plus_compat_sub PG1 f -> f_plus_compat_sub PG2 g ->
-  f_plus_compat_sub PG1 (g \o f).
-Proof.
-intros Hf1 Hg1 x1 y1 Hx1 Hy1;
-    rewrite !comp_correct Hf1// Hg1//; apply Hf; easy.
-Qed.
-
-Lemma f_zero_compat_comp_sub :
-  f_zero_compat_sub PG1 f -> f_zero_compat_sub PG2 g ->
-  f_zero_compat_sub PG1 (g \o f).
-Proof.
-intros Hf1 Hg1 H0;
-    rewrite comp_correct Hf1// Hg1// -Hf1//; apply Hf; easy.
-Qed.
-
-Lemma morphism_m_comp_sub :
-  morphism_m_sub PG1 f -> morphism_m_sub PG2 g ->
-  morphism_m_sub PG1 (g \o f).
-Proof.
-intros [Hfa Hfb] [Hga Hgb]; split;
-    [apply f_plus_compat_comp_sub | apply f_zero_compat_comp_sub]; easy.
-Qed.
-
-End Compatible_Monoid_Morphism_Facts5b.
-
-
-Section Sub_Monoid_Def.
-
-Context {G : AbelianMonoid}.
-Context {PG : G -> Prop}.
-Hypothesis HPG : compatible_m PG.
-
-Definition sub_plus (x y : sub_struct HPG) : sub_struct HPG :=
-  mk_sub (compatible_m_plus HPG _ _ (in_sub x) (in_sub y)).
-
-Definition sub_zero : sub_struct HPG := mk_sub_ HPG _ (compatible_m_zero HPG).
-
-Lemma sub_plus_comm : forall x y, sub_plus x y = sub_plus y x.
-Proof. intros; apply val_inj, plus_comm. Qed.
-
-Lemma sub_plus_assoc:
-  forall x y z, sub_plus x (sub_plus y z) = sub_plus (sub_plus x y) z.
-Proof. intros; apply val_inj, plus_assoc. Qed.
-
-Lemma sub_plus_zero_r : forall x, sub_plus x sub_zero = x.
-Proof. intros; apply val_inj, plus_zero_r. Qed.
-
-Definition sub_AbelianMonoid_mixin :=
-  AbelianMonoid.Mixin _ _ _ sub_plus_comm sub_plus_assoc sub_plus_zero_r.
-
-Canonical Structure sub_AbelianMonoid :=
-  AbelianMonoid.Pack _ sub_AbelianMonoid_mixin (sub_struct HPG).
-
-Lemma val_zero : f_zero_compat val.
-Proof. easy. Qed.
-
-Lemma val_zero_equiv : forall (x : sub_struct HPG), val x = 0 <-> x = 0.
-Proof.
-intros; split; [rewrite -val_zero; apply val_inj | intros; subst; easy].
-Qed.
-
-Lemma val_plus : f_plus_compat val.
-Proof. easy. Qed.
-
-Lemma val_sum : f_sum_compat val.
-Proof.
-intros n; induction n as [| n Hn].
-intros; rewrite 2!sum_nil; easy.
-intros; rewrite 2!sum_ind_l; simpl; f_equal; apply Hn.
-Qed.
-
-Lemma val_morphism_m : morphism_m val.
-Proof. easy. Qed.
-
-Lemma mk_sub_zero : sub_zero = (0 : sub_struct HPG).
-Proof. easy. Qed.
-
-Lemma mk_sub_zero_equiv : forall {x} (Hx : PG x), mk_sub Hx = 0 <-> x = 0.
-Proof. intros; split; [move=> /val_eq | intros; apply val_inj]; easy. Qed.
-
-Lemma mk_sub_plus :
-  forall (x y : G) (Hx : PG x) (Hy : PG y),
-    mk_sub Hx + mk_sub Hy = mk_sub (compatible_m_plus HPG _ _ Hx Hy).
-Proof. easy. Qed.
-
-Lemma mk_sub_sum :
-  forall {n} (u : 'G^n) (Hu : inclF u PG),
-    sum (fun i => mk_sub (Hu i)) = mk_sub (compatible_m_sum HPG _ _ Hu).
-Proof. intros; apply val_inj, val_sum. Qed.
-
-Lemma sub_zero_eq :
-  (0 : sub_struct HPG) = mk_sub_ _ (0 : G) (compatible_m_zero HPG).
-Proof. easy. Qed.
-
-Lemma sub_plus_eq :
-  forall (x y : sub_struct HPG),
-    x + y = mk_sub (compatible_m_plus HPG _ _ (in_sub x) (in_sub y)).
-Proof. intros; apply val_inj; easy. Qed.
-
-Lemma sub_sum_eq :
-  forall {n} (u : '(sub_struct HPG)^n),
-    sum u = mk_sub (compatible_m_sum HPG _ _ (fun i => in_sub (u i))).
-Proof. intros; apply val_inj, val_sum. Qed.
-
-End Sub_Monoid_Def.
-
-
-Section Sub_Monoid_Facts1.
-
-Context {G : AbelianMonoid}.
-Context {PGa PGb : G -> Prop}.
-Hypothesis HPG : incl PGa PGb.
-Hypothesis HPGb : compatible_m PGb.
-Let PGb_m := sub_struct HPGb.
-Let PGa' : PGb_m -> Prop := preimage val PGa.
-
-Lemma RgS_val_m_eq : RgS PGa' val = PGa.
-Proof.
-apply RgS_preimage_equiv; intros x Hx; apply Rg_ex;
-    exists (mk_sub (HPG _ Hx)); easy.
-Qed.
-
-Lemma preimage_val_compatible_m : compatible_m PGa -> compatible_m PGa'.
-Proof. intros; apply compatible_m_preimage; easy. Qed.
-
-Lemma preimage_val_compatible_m_rev : compatible_m PGa' -> compatible_m PGa.
-Proof. intros; rewrite -RgS_val_m_eq; apply RgS_compatible_m; easy. Qed.
-
-Lemma preimage_val_compatible_m_equiv : compatible_m PGa' <-> compatible_m PGa.
-Proof.
-split; [apply preimage_val_compatible_m_rev | apply preimage_val_compatible_m].
-Qed.
-
-End Sub_Monoid_Facts1.
-
-
-Section Sub_Monoid_Facts2.
-
-Context {G : AbelianMonoid}.
-Context {PGb : G -> Prop}.
-Hypothesis HPGb : compatible_m PGb.
-Let PGb_m := sub_struct HPGb.
-Context {PGa : PGb_m -> Prop}.
-Let PGa' := RgS PGa val.
-
-Lemma preimage_val_m_eq : preimage val PGa' = PGa.
-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_m : compatible_m PGa -> compatible_m PGa'.
-Proof. intros; apply RgS_compatible_m; easy. Qed.
-
-Lemma RgS_val_compatible_m_rev : compatible_m PGa' -> compatible_m PGa.
-Proof.
-intros; rewrite -preimage_val_m_eq; apply compatible_m_preimage; easy.
-Qed.
-
-Lemma RgS_val_compatible_m_equiv : compatible_m PGa' <-> compatible_m PGa.
-Proof.
-split; [apply RgS_val_compatible_m_rev | apply RgS_val_compatible_m].
-Qed.
-
-End Sub_Monoid_Facts2.
-
-
-Section Sub_Monoid_Morphism_Facts2a.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-Context {HPG1 : compatible_m PG1}.
-Context {HPG2 : compatible_m PG2}.
-
-Context {f : G1 -> G2}.
-Context {fS : sub_struct HPG1 -> sub_struct HPG2}.
-Hypothesis HfS : forall x, val (fS x) = f (val x).
-
-Lemma sub_m_fun_rev : funS PG1 PG2 f.
-Proof. apply (sub_fun_rev HPG1 HPG2 HfS). Qed.
-
-Lemma sub_m_inj : injS PG1 f -> injective fS.
-Proof. apply sub_inj, HfS. Qed.
-
-Lemma sub_m_inj_rev : injective fS -> injS PG1 f.
-Proof. apply sub_inj_rev, HfS. Qed.
-
-Lemma sub_m_inj_equiv : injective fS <-> injS PG1 f.
-Proof. apply sub_inj_equiv, HfS. Qed.
-
-Lemma sub_m_surj : surjS PG1 PG2 f -> surjective fS.
-Proof. apply sub_surj, HfS. Qed.
-
-Lemma sub_m_surj_rev : surjective fS -> surjS PG1 PG2 f.
-Proof. apply sub_surj_rev, HfS. Qed.
-
-Lemma sub_m_surj_equiv : surjective fS <-> surjS PG1 PG2 f.
-Proof. apply sub_surj_equiv, HfS. Qed.
-
-Lemma sub_m_bij : bijS PG1 PG2 f -> bijective fS.
-Proof. apply sub_bij, HfS; apply inhabited_m. Qed.
-
-Lemma sub_m_bij_rev : bijective fS -> bijS PG1 PG2 f.
-Proof. apply sub_bij_rev, HfS; apply inhabited_m. Qed.
-
-Lemma sub_m_bij_equiv : bijective fS <-> bijS PG1 PG2 f.
-Proof. apply sub_bij_equiv, HfS; apply inhabited_m. Qed.
-
-Lemma sub_m_f_plus_compat : f_plus_compat f -> f_plus_compat fS.
-Proof. intros Hf x1 y1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
-
-Lemma sub_m_f_zero_compat : f_zero_compat f -> f_zero_compat fS.
-Proof. intros Hf; apply val_inj; rewrite HfS Hf; easy. Qed.
-
-Lemma sub_m_f_sum_compat : f_sum_compat f -> f_sum_compat fS.
-Proof.
-intros Hf n u1; apply val_inj; rewrite HfS !val_sum Hf; f_equal.
-apply extF; intros i; rewrite !mapF_correct; easy.
-Qed.
-
-Lemma sub_m_morphism : morphism_m f -> morphism_m fS.
-Proof.
-intros [Hf1 Hf2]; split;
-    [apply (sub_m_f_plus_compat Hf1) | apply (sub_m_f_zero_compat Hf2)].
-Qed.
-
-Lemma Ker_sub_KerS_equiv :
-  forall (x1 : sub_struct HPG1), Ker fS x1 <-> KerS PG1 f (val x1).
-Proof.
-intros; rewrite Ker_equiv KerS_equiv -HfS; split; intros Hx1.
-split; [apply in_sub | apply val_zero_equiv; easy].
-apply val_inj; easy.
-Qed.
-
-Lemma KerS_Ker_sub_equiv :
-  forall {x1} (Hx1 : PG1 x1), KerS PG1 f x1 <-> Ker fS (mk_sub Hx1).
-Proof. intros; rewrite Ker_sub_KerS_equiv; easy. Qed.
-
-Lemma Ker_sub_zero_equiv :
-  morphism_m f ->
-  Ker fS = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
-Proof.
-intros Hf; split; intros H; apply subset_ext_equiv; split.
-(* *)
-intros x1; destruct (classic_dec (PG1 x1)) as [Hx1 | Hx1].
-move=> /(KerS_Ker_sub_equiv Hx1); rewrite H; move=> /mk_sub_zero_equiv; easy.
-intros [Hx1a _]; easy.
-(* *)
-move=> x1 ->; apply compatible_m_zero, KerS_compatible_m; easy.
-move=> x1 /Ker_sub_KerS_equiv Hx1; rewrite H in Hx1; apply val_inj; easy.
-move=> x1 ->; apply (compatible_m_zero (Ker_compatible_m (sub_m_morphism Hf))).
-Qed.
-
-Lemma Rg_sub_RgS_equiv :
-  forall (x2 : sub_struct HPG2), Rg fS x2 <-> RgS PG1 f (val x2).
-Proof.
-intros x2; split; [intros [x1 Hx1]; rewrite HfS; apply Im, in_sub |].
-intros Hx2; inversion Hx2 as [x1 Hx1a Hx1b].
-apply Rg_ex; exists (mk_sub Hx1a); apply val_inj; rewrite HfS; easy.
-Qed.
-
-Lemma RgS_Rg_sub_equiv :
-  forall {x2} (Hx2 : PG2 x2), RgS PG1 f x2 <-> Rg fS (mk_sub Hx2).
-Proof. intros; rewrite Rg_sub_RgS_equiv; easy. Qed.
-
-End Sub_Monoid_Morphism_Facts2a.
-
-
-Section Sub_Monoid_Morphism_Facts2b.
-
-Context {G1 G2 : AbelianMonoid}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-Hypothesis HPG1 : compatible_m PG1.
-Hypothesis HPG2 : compatible_m PG2.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : funS PG1 PG2 f.
-
-Definition fct_sub_m : sub_struct HPG1 -> sub_struct HPG2 :=
-  fct_sub HPG1 HPG2 Hf.
-
-Lemma fct_sub_m_inj : injS PG1 f -> injective fct_sub_m.
-Proof. apply fct_sub_inj. Qed.
-
-Lemma fct_sub_m_inj_rev : injective fct_sub_m -> injS PG1 f.
-Proof. apply fct_sub_inj_rev. Qed.
-
-Lemma fct_sub_m_inj_equiv : injective fct_sub_m <-> injS PG1 f.
-Proof. apply fct_sub_inj_equiv. Qed.
-
-Lemma fct_sub_m_surj : surjS PG1 PG2 f -> surjective fct_sub_m.
-Proof. apply fct_sub_surj. Qed.
-
-Lemma fct_sub_m_surj_rev : surjective fct_sub_m -> surjS PG1 PG2 f.
-Proof. apply fct_sub_surj_rev. Qed.
-
-Lemma fct_sub_m_surj_equiv : surjective fct_sub_m <-> surjS PG1 PG2 f.
-Proof. apply fct_sub_surj_equiv. Qed.
-
-Lemma fct_sub_m_bij : bijS PG1 PG2 f -> bijective fct_sub_m.
-Proof. apply fct_sub_bij, inhabited_m. Qed.
-
-Lemma fct_sub_m_bij_rev : bijective fct_sub_m -> bijS PG1 PG2 f.
-Proof. apply fct_sub_bij_rev, inhabited_m. Qed.
-
-Lemma fct_sub_m_bij_equiv : bijective fct_sub_m <-> bijS PG1 PG2 f.
-Proof. apply fct_sub_bij_equiv, inhabited_m. Qed.
-
-Lemma fct_sub_m_f_plus_compat : f_plus_compat f -> f_plus_compat fct_sub_m.
-Proof. apply sub_m_f_plus_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_m_f_zero_compat : f_zero_compat f -> f_zero_compat fct_sub_m.
-Proof. apply sub_m_f_zero_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_m_f_sum_compat : f_sum_compat f -> f_sum_compat fct_sub_m.
-Proof. apply sub_m_f_sum_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_m_morphism : morphism_m f -> morphism_m fct_sub_m.
-Proof. apply sub_m_morphism, fct_sub_correct. Qed.
-
-End Sub_Monoid_Morphism_Facts2b.
-
-
-Local Open Scope Group_scope.
-
-
-Section Compatible_Group_Def.
-
-Context {G : AbelianGroup}.
-
-Definition opp_closed (PG : G -> Prop) : Prop := forall x, PG x -> PG (- x).
-
-Definition minus_closed (PG : G -> Prop) : Prop :=
-  forall x y, PG x -> PG y -> PG (x - y).
-
-End Compatible_Group_Def.
-
-
-Section Compatible_Group_Facts.
-
-Context {G : AbelianGroup}.
-
-Lemma plus_opp_minus_closed :
-  forall {PG : G -> Prop}, plus_closed PG -> opp_closed PG -> minus_closed PG.
-Proof. intros PG H1 H2 x; unfold minus; auto. Qed.
-
-Lemma minus_zero_closed :
-  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> zero_closed PG.
-Proof.
-move=>> [x Hx]; unfold zero_closed; rewrite -(minus_eq_zero x); auto.
-Qed.
-
-Lemma plus_opp_zero_closed :
-  forall {PG : G -> Prop},
-    nonempty PG -> plus_closed PG -> opp_closed PG -> zero_closed PG.
-Proof. intros; apply minus_zero_closed, plus_opp_minus_closed; easy. Qed.
-
-Lemma minus_opp_closed :
-  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> opp_closed PG.
-Proof.
-intros PG HPG0 HPG1 x; rewrite -(minus_zero_l x).
-apply HPG1, minus_zero_closed; easy.
-Qed.
-
-Lemma minus_plus_closed :
-  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> plus_closed PG.
-Proof.
-intros PG HPG0 HPG1 x y Hx Hy.
-rewrite -minus_opp; apply HPG1; try easy.
-apply minus_opp_closed; easy.
-Qed.
-
-Lemma minus_plus_opp_closed_equiv :
-  forall {PG : G -> Prop},
-    nonempty PG -> minus_closed PG <-> plus_closed PG /\ opp_closed PG.
-Proof.
-intros; split.
-intros; split; [apply minus_plus_closed | apply minus_opp_closed]; easy.
-intros; apply plus_opp_minus_closed; easy.
-Qed.
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (09/06/2023):
-   compatible_g should be compatible_m /\ opp_closed.
-   Then, compatible_g <-> nonempty /\ plus_closed /\ opp_closed
-                      <-> nonempty /\ minus_closed. *)
-
-Lemma minus_closed_compatible_g :
-  forall {PG : G -> Prop}, nonempty PG -> minus_closed PG -> compatible_g PG.
-Proof. intros; split; auto. Qed.
-
-Lemma plus_opp_closed_compatible_g :
-  forall {PG : G -> Prop},
-    nonempty PG -> plus_closed PG -> opp_closed PG -> compatible_g PG.
-Proof. intros; split; auto. Qed.
-
-Lemma compatible_g_nonempty :
-  forall {PG : G -> Prop}, compatible_g PG -> nonempty PG.
-Proof. move=>> H; apply H. Qed.
-
-Lemma compatible_g_zero :
-  forall {PG : G -> Prop}, compatible_g PG -> zero_closed PG.
-Proof. apply compatible_g_zero. Qed.
-
-Lemma compatible_g_plus :
-  forall {PG : G -> Prop}, compatible_g PG -> plus_closed PG.
-Proof. intros; do 2 intro; apply compatible_g_plus; easy. Qed.
-
-Lemma compatible_g_opp :
-  forall {PG : G -> Prop}, compatible_g PG -> opp_closed PG.
-Proof. intros; intro; apply compatible_g_opp; easy. Qed.
-
-Lemma compatible_g_minus :
-  forall {PG : G -> Prop}, compatible_g PG -> minus_closed PG.
-Proof. intros PG H x; apply H. Qed.
-
-Lemma compatible_g_m :
-  forall {PG : G -> Prop}, compatible_g PG -> compatible_m PG.
-Proof.
-intros; split; [apply compatible_g_plus | apply compatible_g_zero]; easy.
-Qed.
-
-Lemma compatible_m_g :
-  forall {PG : G -> Prop},
-    opp_closed PG -> compatible_m PG -> compatible_g PG.
-Proof.
-intros; apply plus_opp_closed_compatible_g;
-  [apply compatible_m_nonempty | apply compatible_m_plus | ]; easy.
-Qed.
-
-Lemma compatible_g_m_equiv :
-  forall {PG : G -> Prop},
-    opp_closed PG -> compatible_g PG <-> compatible_m PG.
-Proof. intros; split; [apply compatible_g_m | apply compatible_m_g; easy]. Qed.
-
-Lemma compatible_g_zero_sub_struct : compatible_g (@zero_sub_struct G).
-Proof.
-split; try now exists 0.
-intros x y Hx Hy; rewrite Hx Hy opp_zero plus_zero_l; easy.
-Qed.
-
-Lemma compatible_g_full : forall {PG : G -> Prop}, full PG -> compatible_g PG.
-Proof. intros; split; try exists 0; easy. Qed.
-
-Lemma compatible_g_fullset : compatible_g (@fullset G).
-Proof. apply compatible_g_full; easy. Qed.
-
-Lemma compatible_g_inter :
-  forall {PG QG : G -> Prop},
-    compatible_g PG -> compatible_g QG -> compatible_g (inter PG QG).
-Proof.
-intros PG QG HPG HQG; split.
-destruct HPG as [HPG _], HQG as [HQG _];
-    intros x y [Hx1 Hx2] [Hy1 Hy2]; split; auto.
-exists 0; split; apply compatible_g_zero; easy.
-Qed.
-
-(* Gostiaux T1, Th 4.6 p. 89. *)
-Lemma compatible_g_inter_any :
-  forall {Idx : Type} {PG : Idx -> G -> Prop},
-    (forall i, compatible_g (PG i)) -> compatible_g (inter_any PG).
-Proof.
-intros Idx PG HPG; split.
-intros x y Hx Hy i; apply HPG; easy.
-exists 0; intros i; apply compatible_g_zero; easy.
-Qed.
-
-(* Gostiaux T1, Th 4.10 pp. 89-91. *)
-Definition span_g (gen : G -> Prop) := span_struct compatible_g gen.
-
-Lemma span_g_compatible_g : forall gen, compatible_g (span_g gen).
-Proof.
-apply span_struct_compatible; move=>>; apply compatible_g_inter_any.
-Qed.
-
-Lemma span_g_nonempty : forall gen, nonempty (span_g gen).
-Proof. intros; apply compatible_g_nonempty, span_g_compatible_g. Qed.
-
-Lemma span_g_zero : forall gen, zero_closed (span_g gen).
-Proof. intros; apply compatible_g_zero, span_g_compatible_g. Qed.
-
-Lemma span_g_plus : forall gen, plus_closed (span_g gen).
-Proof. intros; apply compatible_g_plus, span_g_compatible_g. Qed.
-
-Lemma span_g_opp : forall gen, opp_closed (span_g gen).
-Proof. intros; apply compatible_g_opp, span_g_compatible_g. Qed.
-
-Lemma span_g_minus : forall gen, minus_closed (span_g gen).
-Proof. intros; apply compatible_g_minus, span_g_compatible_g. Qed.
-
-Lemma span_g_incl : forall gen, incl gen (span_g gen).
-Proof. apply span_struct_incl. Qed.
-
-(* Gostiaux T1, Def 4.9 p. 90. *)
-Lemma span_g_lub :
-  forall gen PG, compatible_g PG -> incl gen PG -> incl (span_g gen) PG.
-Proof. apply span_struct_lub. Qed.
-
-Lemma span_g_full : forall PG, compatible_g PG -> span_g PG = PG.
-Proof. apply span_struct_full. Qed.
-
-End Compatible_Group_Facts.
-
-
-Section Compatible_Group_Morphism_Facts1.
-
-Context {G1 G2 : AbelianGroup}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_g f.
-
-Lemma compatible_g_image : compatible_g PG1 -> compatible_g (image f PG1).
-Proof.
-intros [HPG1 [x01 Hx01]]; split; try now exists (f x01).
-intros _ _ [x1 Hx1] [y1 Hy1]; rewrite -(morphism_g_opp f Hf) -Hf.
-apply Im; auto.
-Qed.
-
-Lemma compatible_g_preimage : compatible_g PG2 -> compatible_g (preimage f PG2).
-Proof.
-intros HPG2; split; unfold preimage.
-destruct HPG2 as [HPG2 _]; intros; rewrite Hf morphism_g_opp; auto.
-apply compatible_g_zero in HPG2; exists 0; rewrite morphism_g_zero; easy.
-Qed.
-
-Lemma compatible_g_morphism_g : compatible_g (@morphism_g G1 G2).
-Proof.
-split; [apply morphism_g_fct_minus | exists 0; apply morphism_g_fct_zero].
-Qed.
-
-End Compatible_Group_Morphism_Facts1.
-
-
-Section Compatible_Group_Morphism_Def.
-
-Context {G1 G2 : AbelianGroup}.
-
-Variable PG1 : G1 -> Prop.
-
-Variable f : G1 -> G2.
-
-Definition morphism_g_sub : Prop := f_plus_compat_sub PG1 f.
-
-End Compatible_Group_Morphism_Def.
-
-
-Section Compatible_Group_Morphism_Facts2a.
-
-Context {G1 G2 : AbelianGroup}.
-
-Context {PG1 : G1 -> Prop}.
-Hypothesis HPG1 : compatible_g PG1.
-
-Context {PG2 : G2 -> Prop}.
-Hypothesis HPG2 : compatible_g PG2.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_g f.
-
-Lemma KerS_compatible_g : compatible_g (KerS PG1 f).
-Proof.
-apply compatible_g_inter, compatible_g_preimage,
-    compatible_g_zero_sub_struct; easy.
-Qed.
-
-Lemma RgS_gen_compatible_g : compatible_g (RgS_gen PG1 PG2 f).
-Proof. apply compatible_g_inter, compatible_g_image; easy. Qed.
-
-Lemma RgS_compatible_g : compatible_g (RgS PG1 f).
-Proof. apply compatible_g_image; easy. Qed.
-
-Lemma KerS_g_zero_equiv :
-  KerS PG1 f = zero_sub_struct <-> incl (KerS PG1 f) zero_sub_struct.
-Proof.
-apply KerS_m_zero_equiv; [apply compatible_g_m | apply morphism_g_m]; easy.
-Qed.
-
-Lemma gmS_injS : incl (KerS PG1 f) zero_sub_struct -> injS PG1 f.
-Proof.
-intros Hf1 x1 y1 Hx1 Hy1; rewrite -2!minus_zero_equiv -(morphism_g_minus _ Hf).
-intros H1; apply Hf1, KerS_equiv; split; [apply HPG1 |]; easy.
-Qed.
-
-Lemma gmS_injS_rev : injS PG1 f -> KerS PG1 f = zero_sub_struct.
-Proof.
-apply mmS_injS_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
-Qed.
-
-Lemma gmS_injS_equiv : injS PG1 f <-> KerS PG1 f = zero_sub_struct.
-Proof.
-split; [apply gmS_injS_rev | rewrite KerS_g_zero_equiv; apply gmS_injS].
-Qed.
-
-Lemma gmS_bijS_gen :
-  funS PG1 PG2 f -> incl (KerS PG1 f) zero_sub_struct -> incl PG2 (RgS PG1 f) ->
-  bijS PG1 PG2 f.
-Proof.
-intros; apply bijS_equiv; [apply inhabited_g | repeat split];
-    [| apply gmS_injS | apply surjS_correct]; easy.
-Qed.
-
-Lemma gmS_bijS_gen_rev :
-  bijS PG1 PG2 f ->
-  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
-Proof.
-apply mmS_bijS_gen_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
-Qed.
-
-Lemma gmS_bijS_gen_equiv :
-  bijS PG1 PG2 f <->
-  funS PG1 PG2 f /\ KerS PG1 f = zero_sub_struct /\ RgS_gen PG1 PG2 f = PG2.
-Proof.
-split; [apply gmS_bijS_gen_rev | rewrite KerS_g_zero_equiv RgS_gen_full_equiv].
-intros; apply gmS_bijS_gen; easy.
-Qed.
-
-Lemma gmS_bijS : incl (KerS PG1 f) zero_sub_struct -> bijS PG1 (RgS PG1 f) f.
-Proof.
-rewrite bijS_RgS_equiv; [| apply inhabited_g]; apply gmS_injS; easy.
-Qed.
-
-Lemma gmS_bijS_rev : bijS PG1 (RgS PG1 f) f -> KerS PG1 f = zero_sub_struct.
-Proof.
-apply mmS_bijS_rev; [apply compatible_g_m | apply morphism_g_m]; easy.
-Qed.
-
-Lemma gmS_bijS_equiv : bijS PG1 (RgS PG1 f) f <-> KerS PG1 f = zero_sub_struct.
-Proof.
-split; [apply gmS_bijS_rev | rewrite KerS_g_zero_equiv//; apply gmS_bijS].
-Qed.
-
-Lemma gmS_bijS_injS_equiv : bijS PG1 (RgS PG1 f) f <-> injS PG1 f.
-Proof. apply mmS_bijS_injS_equiv. Qed.
-
-End Compatible_Group_Morphism_Facts2a.
-
-
-Section Compatible_Group_Morphism_Facts2b.
-
-Context {G1 G2 : AbelianGroup}.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : morphism_g f.
-
-Lemma Ker_compatible_g : compatible_g (Ker f).
-Proof.
-rewrite -KerS_full; apply (KerS_compatible_g compatible_g_fullset Hf).
-Qed.
-
-Lemma Rg_compatible_g : compatible_g (Rg f).
-Proof.
-rewrite -RgS_full; apply (RgS_compatible_g compatible_g_fullset Hf).
-Qed.
-
-Lemma Ker_g_zero_equiv :
-  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
-Proof. apply Ker_m_zero_equiv, morphism_g_m; easy. Qed.
-
-Lemma gm_inj : incl (Ker f) zero_sub_struct -> injective f.
-Proof.
-rewrite -KerS_full inj_S_equiv; apply (gmS_injS compatible_g_fullset Hf).
-Qed.
-
-Lemma gm_inj_rev : injective f -> Ker f = zero_sub_struct.
-Proof. apply mm_inj_rev, morphism_g_m; easy. Qed.
-
-Lemma gm_inj_equiv : injective f <-> Ker f = zero_sub_struct.
-Proof.
-rewrite -KerS_full inj_S_equiv; apply (gmS_injS_equiv compatible_g_fullset Hf).
-Qed.
-
-Lemma gm_bij :
-  incl (Ker f) zero_sub_struct -> incl fullset (Rg f) -> bijective f.
-Proof.
-rewrite -KerS_full -RgS_full bij_S_equiv; [| apply inhabited_g].
-apply (gmS_bijS_gen compatible_g_fullset Hf), funS_full.
-Qed.
-
-Lemma gm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
-Proof. apply mm_bij_rev, morphism_g_m; easy. Qed.
-
-Lemma gm_bij_equiv : bijective f <-> Ker f = zero_sub_struct /\ Rg f = fullset.
-Proof.
-rewrite -KerS_full -RgS_gen_full bij_S_equiv; [| apply inhabited_g].
-rewrite (gmS_bijS_gen_equiv compatible_g_fullset Hf); easy.
-Qed.
-
-End Compatible_Group_Morphism_Facts2b.
-
-
-Section Compatible_Group_Morphism_Facts3.
-
-Context {E1 E2 E3 : AbelianGroup}.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : morphism_g f.
-
-Context {g : E2 -> E3}.
-Hypothesis Hg : morphism_g g.
-
-Lemma Ker_g_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
-Proof. apply Ker_m_comp_r; apply morphism_g_m; easy. Qed.
-
-Lemma Ker_g_comp_r_alt :
-  forall h x2, injective g -> h = g \o f ->
-    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
-Proof. apply Ker_m_comp_r_alt; apply morphism_g_m; easy. Qed.
-
-End Compatible_Group_Morphism_Facts3.
-
-
-Section Compatible_Group_Morphism_Facts4a.
-
-Context {G1 G2 : AbelianGroup}.
-
-Variable PG1 : G1 -> Prop.
-
-Lemma morphism_g_sub_id : morphism_g_sub PG1 ssrfun.id.
-Proof. easy. Qed.
-
-Variable f : G1 -> G2.
-
-Lemma morphism_g_is_sub : morphism_g f -> morphism_g_sub PG1 f.
-Proof. easy. Qed.
-
-End Compatible_Group_Morphism_Facts4a.
-
-
-Section Compatible_Group_Morphism_Facts4b.
-
-Context {G1 G2 G3 : AbelianGroup}.
-
-Context {PG1 : G1 -> Prop}.
-Variable PG2 : G2 -> Prop.
-
-Variable f : G1 -> G2.
-Variable g : G2 -> G3.
-
-Hypothesis Hf : funS PG1 PG2 f.
-
-Lemma morphism_g_comp_sub :
-  morphism_g_sub PG1 f -> morphism_g_sub PG2 g ->
-  morphism_g_sub PG1 (g \o f).
-Proof. apply f_plus_compat_comp_sub; easy. Qed.
-
-End Compatible_Group_Morphism_Facts4b.
-
-
-Section Sub_Group_Def.
-
-Context {G : AbelianGroup}.
-Context {PG : G -> Prop}.
-
-Definition sub_g (HPG : compatible_g PG) : Type :=
-  sub_struct (compatible_g_m HPG).
-Definition mk_sub_g_ (HPG : compatible_g PG) {x} (Hx : PG x) : sub_g HPG :=
-  mk_sub_ (compatible_g_m HPG) x Hx.
-Definition mk_sub_g {HPG : compatible_g PG} {x} (Hx : PG x) : sub_g HPG :=
-  mk_sub_g_ HPG Hx.
-
-Hypothesis HPG : compatible_g PG.
-
-Definition sub_opp (x : sub_g HPG) : sub_g HPG :=
-  mk_sub_g (compatible_g_opp HPG _ (in_sub x)).
-
-Lemma sub_plus_opp_r : forall x, x + (sub_opp x) = 0.
-Proof. intros; apply val_inj, plus_opp_r. Qed.
-
-Definition sub_AbelianGroup_mixin := AbelianGroup.Mixin _ _ sub_plus_opp_r.
-
-Canonical Structure sub_AbelianGroup :=
-  AbelianGroup.Pack _
-    (AbelianGroup.Class _ _ sub_AbelianGroup_mixin) (sub_g HPG).
-
-Lemma val_opp : f_opp_compat val.
-Proof. easy. Qed.
-
-Lemma val_minus : f_minus_compat val.
-Proof. easy. Qed.
-
-Lemma val_morphism_g : morphism_g val.
-Proof. easy. Qed.
-
-Lemma mk_sub_g_zero : mk_sub_g (compatible_g_zero HPG) = (0 : sub_g HPG).
-Proof. apply val_inj; easy. Qed.
-
-Lemma mk_sub_g_zero_equiv :
-  forall {x} (Hx : PG x), mk_sub_g Hx = 0 :> sub_g HPG <-> x = 0.
-Proof. apply mk_sub_zero_equiv. Qed.
-
-Lemma mk_sub_opp :
-  forall (x : G) (Hx : PG x), - mk_sub Hx = mk_sub (compatible_g_opp HPG _ Hx).
-Proof. easy. Qed.
-
-Lemma mk_sub_minus :
-  forall (x y : G) (Hx : PG x) (Hy : PG y),
-    mk_sub Hx - mk_sub Hy = mk_sub (compatible_g_minus HPG _ _ Hx Hy).
-Proof. intros; apply val_inj; easy. Qed.
-
-Lemma sub_opp_eq :
-  forall (x : sub_g HPG), - x = mk_sub (compatible_g_opp HPG _ (in_sub x)).
-Proof. easy. Qed.
-
-Lemma sub_minus_eq :
-  forall (x y : sub_g HPG),
-    x - y = mk_sub (compatible_g_minus HPG _ _ (in_sub x) (in_sub y)).
-Proof. intros; apply mk_sub_minus. Qed.
-
-Lemma mk_sub_g_inj :
-  forall {x y} (Hx : PG x) (Hy : PG y),
-    mk_sub_g Hx = mk_sub_g Hy :> sub_g HPG -> x = y.
-Proof. apply mk_sub_inj. Qed.
-
-End Sub_Group_Def.
-
-
-Section Sub_Group_Facts1.
-
-Context {G : AbelianGroup}.
-Context {PGa PGb : G -> Prop}.
-Hypothesis HPG : incl PGa PGb.
-Hypothesis HPGb : compatible_g PGb.
-Let PGb_g := sub_g HPGb.
-Let PGa' : PGb_g -> Prop := preimage val PGa.
-
-Lemma RgS_val_g_eq : RgS PGa' val = PGa.
-Proof.
-apply RgS_preimage_equiv; intros x Hx; apply Rg_ex;
-    exists (mk_sub (HPG _ Hx)); easy.
-Qed.
-
-Lemma preimage_val_compatible_g : compatible_g PGa -> compatible_g PGa'.
-Proof. intros; apply compatible_g_preimage; easy. Qed.
-
-Lemma preimage_val_compatible_g_rev : compatible_g PGa' -> compatible_g PGa.
-Proof. intros; rewrite -RgS_val_g_eq; apply RgS_compatible_g; easy. Qed.
-
-Lemma preimage_val_compatible_g_equiv : compatible_g PGa' <-> compatible_g PGa.
-Proof.
-split; [apply preimage_val_compatible_g_rev | apply preimage_val_compatible_g].
-Qed.
-
-End Sub_Group_Facts1.
-
-
-Section Sub_Group_Facts2.
-
-Context {G : AbelianGroup}.
-Context {PGb : G -> Prop}.
-Hypothesis HPGb : compatible_g PGb.
-Let PGb_g := sub_g HPGb.
-Context {PGa : PGb_g -> Prop}.
-Let PGa' := RgS PGa val.
-
-Lemma preimage_val_g_eq : preimage val PGa' = PGa.
-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_g : compatible_g PGa -> compatible_g PGa'.
-Proof. intros; apply RgS_compatible_g; easy. Qed.
-
-Lemma RgS_val_compatible_g_rev : compatible_g PGa' -> compatible_g PGa.
-Proof.
-intros; rewrite -preimage_val_g_eq; apply compatible_g_preimage; easy.
-Qed.
-
-Lemma RgS_val_compatible_g_equiv : compatible_g PGa' <-> compatible_g PGa.
-Proof.
-split; [apply RgS_val_compatible_g_rev | apply RgS_val_compatible_g].
-Qed.
-
-End Sub_Group_Facts2.
-
-
-Section Sub_Group_Morphism_Facts1.
-
-Context {G1 G2 : AbelianGroup}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-Context {HPG1 : compatible_g PG1}.
-Context {HPG2 : compatible_g PG2}.
-
-Context {f : G1 -> G2}.
-Context {fS : sub_g HPG1 -> sub_g HPG2}.
-Hypothesis HfS : forall x, val (fS x) = f (val x).
-
-Lemma sub_g_fun_rev : funS PG1 PG2 f.
-Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
-
-Lemma sub_g_inj : injS PG1 f -> injective fS.
-Proof. apply sub_m_inj, HfS. Qed.
-
-Lemma sub_g_inj_rev : injective fS -> injS PG1 f.
-Proof. apply sub_m_inj_rev, HfS. Qed.
-
-Lemma sub_g_inj_equiv : injective fS <-> injS PG1 f.
-Proof. apply sub_m_inj_equiv, HfS. Qed.
-
-Lemma sub_g_surj : surjS PG1 PG2 f -> surjective fS.
-Proof. apply sub_m_surj, HfS. Qed.
-
-Lemma sub_g_surj_rev : surjective fS -> surjS PG1 PG2 f.
-Proof. apply sub_m_surj_rev, HfS. Qed.
-
-Lemma sub_g_surj_equiv : surjective fS <-> surjS PG1 PG2 f.
-Proof. apply sub_m_surj_equiv, HfS. Qed.
-
-Lemma sub_g_bij : bijS PG1 PG2 f -> bijective fS.
-Proof. apply sub_m_bij, HfS. Qed.
-
-Lemma sub_g_bij_rev : bijective fS -> bijS PG1 PG2 f.
-Proof. apply sub_m_bij_rev, HfS. Qed.
-
-Lemma sub_g_bij_equiv : bijective fS <-> bijS PG1 PG2 f.
-Proof. apply sub_m_bij_equiv, HfS. Qed.
-
-Lemma sub_g_f_opp_compat : f_opp_compat f -> f_opp_compat fS.
-Proof. intros Hf x1; apply val_inj; rewrite HfS Hf -HfS; easy. Qed.
-
-Lemma sub_g_f_minus_compat : f_minus_compat f -> f_minus_compat fS.
-Proof. intros Hf x1 y1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
-
-Lemma sub_g_morphism : morphism_g f -> morphism_g fS.
-Proof. apply sub_m_f_plus_compat, HfS. Qed.
-
-Lemma Ker_sub_g_KerS_zero_equiv :
-  morphism_g f ->
-  Ker fS = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
-Proof. move=> /morphism_g_m; apply Ker_sub_zero_equiv, HfS. Qed.
-
-Lemma KerS_g_zero_equiv_alt :
-  morphism_g f ->
-  KerS PG1 f = zero_sub_struct <->
-  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
-Proof.
-move=>>; rewrite subset_ext_equiv; split; [intros [Hf1 _] | intros Hf1; split].
-intros [x1 Hx1] Hx1a; apply val_inj, Hf1; easy.
-intros x1 [Hx1 Hx1a]; apply (mk_sub_g_inj HPG1 Hx1 (compatible_g_zero HPG1)).
-rewrite mk_sub_g_zero; apply Hf1; easy.
-move=> x1 ->; apply compatible_g_zero, KerS_compatible_g; easy.
-Qed.
-
-Lemma Ker_sub_g_zero_equiv :
-  morphism_g f ->
-  Ker fS = zero_sub_struct <->
-  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros; rewrite Ker_sub_g_KerS_zero_equiv//; apply KerS_g_zero_equiv_alt; easy.
-Qed.
-
-Lemma gmS_injS_sub_equiv :
-  morphism_g f -> injS PG1 f <-> Ker fS = zero_sub_struct.
-Proof.
-intros Hf; rewrite gmS_injS_equiv//; split;
-    intros Hf1; apply subset_ext; intros x1.
-rewrite (Ker_sub_KerS_equiv HfS) Hf1; apply val_zero_equiv.
-destruct (classic_dec (PG1 x1)) as [Hx1 | Hx1a].
-rewrite (KerS_Ker_sub_equiv HfS Hx1) Hf1; apply mk_sub_zero_equiv.
-split; intros Hx1b; contradict Hx1a;
-    [apply Hx1b | rewrite Hx1b; apply compatible_g_zero, HPG1].
-Qed.
-
-Lemma gmS_injS_sub_equiv_alt :
-  morphism_g f ->
-  injS PG1 f <-> forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros H; rewrite (gmS_injS_sub_equiv H); apply (Ker_sub_g_zero_equiv H).
-Qed.
-
-Lemma gmS_bijS_sub_equiv :
-  morphism_g f -> bijS PG1 (RgS PG1 f) f <-> Ker fS = zero_sub_struct.
-Proof. rewrite gmS_bijS_injS_equiv; apply gmS_injS_sub_equiv. Qed.
-
-Lemma gmS_bijS_sub_equiv_alt :
-  morphism_g f ->
-  bijS PG1 (RgS PG1 f) f <-> forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
-Proof. rewrite gmS_bijS_injS_equiv; apply gmS_injS_sub_equiv_alt. Qed.
-
-End Sub_Group_Morphism_Facts1.
-
-
-Section Sub_Group_Morphism_Facts2.
-
-Context {G1 G2 : AbelianGroup}.
-
-Context {PG1 : G1 -> Prop}.
-Context {PG2 : G2 -> Prop}.
-Hypothesis HPG1 : compatible_g PG1.
-Hypothesis HPG2 : compatible_g PG2.
-
-Context {f : G1 -> G2}.
-Hypothesis Hf : funS PG1 PG2 f.
-
-Definition fct_sub_g : sub_g HPG1 -> sub_g HPG2 :=
-  fct_sub (compatible_g_m HPG1) (compatible_g_m HPG2) Hf.
-
-Lemma fct_sub_g_inj : injS PG1 f -> injective fct_sub_g.
-Proof. apply fct_sub_m_inj. Qed.
-
-Lemma fct_sub_g_inj_rev : injective fct_sub_g -> injS PG1 f.
-Proof. apply fct_sub_m_inj_rev. Qed.
-
-Lemma fct_sub_g_inj_equiv : injective fct_sub_g <-> injS PG1 f.
-Proof. apply fct_sub_m_inj_equiv. Qed.
-
-Lemma fct_sub_g_surj : surjS PG1 PG2 f -> surjective fct_sub_g.
-Proof. apply fct_sub_m_surj. Qed.
-
-Lemma fct_sub_g_surj_rev : surjective fct_sub_g -> surjS PG1 PG2 f.
-Proof. apply fct_sub_m_surj_rev. Qed.
-
-Lemma fct_sub_g_surj_equiv : surjective fct_sub_g <-> surjS PG1 PG2 f.
-Proof. apply fct_sub_m_surj_equiv. Qed.
-
-Lemma fct_sub_g_bij : bijS PG1 PG2 f -> bijective fct_sub_g.
-Proof. apply fct_sub_m_bij. Qed.
-
-Lemma fct_sub_g_bij_rev : bijective fct_sub_g -> bijS PG1 PG2 f.
-Proof. apply fct_sub_m_bij_rev. Qed.
-
-Lemma fct_sub_g_bij_equiv : bijective fct_sub_g <-> bijS PG1 PG2 f.
-Proof. apply fct_sub_m_bij_equiv. Qed.
-
-Lemma fct_sub_g_f_opp_compat : f_opp_compat f -> f_opp_compat fct_sub_g.
-Proof. apply sub_g_f_opp_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_g_f_minus_compat : f_minus_compat f -> f_minus_compat fct_sub_g.
-Proof. apply sub_g_f_minus_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_g_morphism : morphism_g f -> morphism_g fct_sub_g.
-Proof. apply sub_g_morphism, fct_sub_correct. Qed.
-
-Lemma Ker_fct_sub_g_KerS_zero_equiv :
-  morphism_g f ->
-  Ker fct_sub_g = zero_sub_struct <-> KerS PG1 f = zero_sub_struct.
-Proof. apply Ker_sub_g_KerS_zero_equiv, fct_sub_correct. Qed.
-
-Lemma Ker_fct_sub_g_zero_equiv :
-  morphism_g f ->
-  Ker fct_sub_g = zero_sub_struct <->
-  forall (x1 : sub_g HPG1), f (val x1) = 0 -> x1 = 0.
-Proof. apply Ker_sub_g_zero_equiv, fct_sub_correct. Qed.
-
-Lemma gmS_injS_fct_sub_equiv :
-  morphism_g f -> injS PG1 f <-> Ker fct_sub_g = zero_sub_struct.
-Proof. apply gmS_injS_sub_equiv, fct_sub_correct. Qed.
-
-Lemma gmS_bijS_fct_sub_equiv :
-  morphism_g f -> bijS PG1 (RgS PG1 f) f <-> Ker fct_sub_g = zero_sub_struct.
-Proof. apply gmS_bijS_sub_equiv, fct_sub_correct. Qed.
-
-End Sub_Group_Morphism_Facts2.
-
-
-Local Open Scope Ring_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.
-
-
-Section Compatible_ModuleSpace_Def.
-
-Context {K : Ring}.
-Context {E : ModuleSpace K}.
-
-Definition scal_closed (PE : E -> Prop) : Prop :=
-  forall a u, PE u -> PE (scal a u).
-
-Definition scal_rev_closed (PE : E -> Prop) : Prop :=
-  forall a u, invertible a -> PE (scal a u) -> PE u.
-
-Definition comb_lin_closed (PE : E -> Prop) : Prop :=
-  forall n L (B : 'E^n), inclF B PE -> PE (comb_lin L B).
-
-End Compatible_ModuleSpace_Def.
-
-
-Section Compatible_ModuleSpace_Facts.
-
-Context {K : Ring}.
-Context {E : ModuleSpace K}.
-
-Lemma scal_scal_rev_closed :
-  forall (PE : E -> Prop), scal_closed PE -> scal_rev_closed PE.
-Proof.
-intros PE HPE a u [b [Hb _]]; rewrite -{2}(scal_one u) -Hb -scal_assoc; auto.
-Qed.
-
-Lemma scal_rev_scal_closed :
-  forall (PE : E -> Prop),
-    has_inv K -> zero_closed PE -> scal_rev_closed PE -> scal_closed PE.
-Proof.
-move=>> HK HPE0 HPE a u Hu; destruct (classic (a = 0)) as [Ha | Ha].
-rewrite Ha scal_zero_l; easy. specialize (HK a Ha).
-generalize HK; intros [b Hb]; generalize Hb; intros [Hb1 _].
-rewrite -(scal_one u) -Hb1 -scal_assoc in Hu; apply (HPE b); try easy.
-apply (is_inverse_invertible_r a _ Hb).
-Qed.
-
-Lemma scal_closed_equiv :
-  forall (PE : E -> Prop),
-    has_inv K -> zero_closed PE -> scal_closed PE <-> scal_rev_closed PE.
-Proof.
-move=>> HK HPE; split.
-apply scal_scal_rev_closed.
-apply scal_rev_scal_closed; easy.
-Qed.
-
-Lemma scal_zero_closed :
-  forall {PE : E -> Prop}, nonempty PE -> scal_closed PE -> zero_closed PE.
-Proof.
-intros PE [x Hx]; unfold zero_closed; rewrite -(scal_zero_l x); auto.
-Qed.
-
-Lemma scal_opp_closed :
-  forall {PE : E -> Prop}, scal_closed PE -> opp_closed PE.
-Proof. intros PE H u; rewrite <- scal_opp_one; apply H. Qed.
-
-Lemma comb_lin_zero_closed :
-  forall {PE : E -> Prop}, comb_lin_closed PE -> zero_closed PE.
-Proof.
-move=>> HPE.
-destruct (hat0F_is_nonempty K) as [L], (hat0F_is_nonempty E) as [B].
-specialize (HPE 0%nat L B); rewrite comb_lin_nil in HPE.
-apply HPE, inclF_nil.
-Qed.
-
-Lemma plus_scal_comb_lin_closed :
-  forall {PE : E -> Prop},
-    nonempty PE -> plus_closed PE -> scal_closed PE -> comb_lin_closed PE.
-Proof.
-intros PE HPE0 HPE1 HPE2 n L B HB; induction n as [| n Hn].
-rewrite comb_lin_nil; apply scal_zero_closed; easy.
-rewrite comb_lin_ind_l; apply HPE1. apply HPE2, HB. apply Hn; intro; apply HB.
-Qed.
-
-Lemma comb_lin_plus_scal_closed :
-  forall {PE : E -> Prop},
-    comb_lin_closed PE -> plus_closed PE /\ scal_closed PE.
-Proof.
-intros PE HPE; split.
-(* *)
-intros u v Hu Hv; rewrite -sum_coupleF -comb_lin_ones_l.
-apply HPE; intros i; destruct (ord2_dec i);
-    [rewrite coupleF_l | rewrite coupleF_r]; easy.
-(* *)
-intros a u Hu; replace (scal a u) with (comb_lin (singleF a) (singleF u)).
-apply HPE; intro; rewrite singleF_0; easy.
-rewrite comb_lin_1 2!singleF_0; easy.
-Qed.
-
-Lemma comb_lin_plus_scal_closed_equiv :
-  forall {PE : E -> Prop},
-    nonempty PE -> comb_lin_closed PE <-> plus_closed PE /\ scal_closed PE.
-Proof.
-intros; split.
-apply comb_lin_plus_scal_closed.
-intros; apply plus_scal_comb_lin_closed; easy.
-Qed.
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (09/06/2023):
-   compatible_ms should be compatible_g /\ scal_closed.
-   Then, compatible_ms <-> nonempty /\ plus_closed /\ scal_closed
-                       <-> comb_lin_closed. *)
-
-Lemma plus_scal_closed_compatible_ms :
-  forall {PE : E -> Prop},
-    nonempty PE -> plus_closed PE -> scal_closed PE -> compatible_ms PE.
-Proof.
-intros; split; auto.
-apply plus_opp_closed_compatible_g; try apply scal_opp_closed; easy.
-Qed.
-
-Lemma comb_lin_closed_compatible_ms :
-  forall {PE : E -> Prop}, comb_lin_closed PE -> compatible_ms PE.
-Proof.
-move=>> HPE; apply plus_scal_closed_compatible_ms.
-exists 0; apply comb_lin_zero_closed; easy.
-1,2: apply comb_lin_plus_scal_closed; easy.
-Qed.
-
-Lemma compatible_ms_g :
-  forall {PE : E -> Prop}, compatible_ms PE -> compatible_g PE.
-Proof. move=>> H; apply H. Qed.
-
-Lemma compatible_g_ms :
-  forall {PE : E -> Prop},
-    scal_closed PE -> compatible_g PE -> compatible_ms PE.
-Proof.
-intros; apply plus_scal_closed_compatible_ms;
-    [apply compatible_g_nonempty | apply: compatible_g_plus |]; easy.
-Qed.
-
-Lemma compatible_ms_g_equiv :
-  forall {PE : E -> Prop},
-    scal_closed PE -> compatible_ms PE <-> compatible_g PE.
-Proof.
-intros; split; [apply compatible_ms_g | apply compatible_g_ms; easy].
-Qed.
-
-Lemma compatible_ms_m :
-  forall {PE : E -> Prop}, compatible_ms PE -> compatible_m PE.
-Proof. move=>> /compatible_ms_g; apply: compatible_g_m. Qed.
-
-Lemma compatible_m_ms :
-  forall {PE : E -> Prop},
-    scal_closed PE -> compatible_m PE -> compatible_ms PE.
-Proof.
-intros; apply compatible_g_ms;
-    [| apply compatible_m_g; [apply scal_opp_closed |]]; easy.
-Qed.
-
-Lemma compatible_ms_m_equiv :
-  forall {PE : E -> Prop},
-    scal_closed PE -> compatible_ms PE <-> compatible_m PE.
-Proof.
-intros; split; [apply compatible_ms_m | apply compatible_m_ms; easy].
-Qed.
-
-Lemma compatible_ms_nonempty :
-  forall {PE : E -> Prop}, compatible_ms PE -> nonempty PE.
-Proof. move=>> /compatible_ms_g; apply compatible_g_nonempty. Qed.
-
-Lemma compatible_ms_zero :
-  forall {PE : E -> Prop}, compatible_ms PE -> zero_closed PE.
-Proof. move=>> /compatible_ms_g; apply: compatible_g_zero. Qed.
-
-Lemma compatible_ms_plus :
-  forall {PE : E -> Prop}, compatible_ms PE -> plus_closed PE.
-Proof. move=>> /compatible_ms_g; apply: compatible_g_plus. Qed.
-
-Lemma compatible_ms_opp :
-  forall {PE : E -> Prop}, compatible_ms PE -> opp_closed PE.
-Proof. move=>> /compatible_ms_g; apply compatible_g_opp. Qed.
-
-Lemma compatible_ms_minus :
-  forall {PE : E -> Prop}, compatible_ms PE -> minus_closed PE.
-Proof. move=>> /compatible_ms_g; apply compatible_g_minus. Qed.
-
-Lemma compatible_ms_scal :
-  forall {PE : E -> Prop}, compatible_ms PE -> scal_closed PE.
-Proof. move=>> [_ H]; move=>>; apply H. Qed.
-
-Lemma compatible_ms_scal_rev :
-  forall {PE : E -> Prop}, compatible_ms PE -> scal_rev_closed PE.
-Proof. move=>> /compatible_ms_scal; apply scal_scal_rev_closed. Qed.
-
-Lemma compatible_ms_comb_lin :
-  forall {PE : E -> Prop}, compatible_ms PE -> comb_lin_closed PE.
-Proof.
-intros PE HPE n L B HB; induction n as [| n Hn].
-(* *)
-rewrite comb_lin_nil.
-apply compatible_ms_zero; easy.
-(* *)
-rewrite comb_lin_ind_l.
-apply compatible_ms_plus; try easy.
-apply HPE, HB.
-apply Hn.
-intros i; apply HB.
-Qed.
-
-Lemma compatible_ms_plus_equiv :
-  forall {PE : E -> Prop},
-    compatible_ms PE -> forall u v, PE u -> PE (u + v) <-> PE v.
-Proof.
-intros PE HPE u v Hu; split; [rewrite -{2}(minus_plus_l u v) minus_sym | ].
-1,2: apply compatible_ms_plus; try easy.
-apply compatible_ms_opp; easy.
-Qed.
-
-Lemma compatible_ms_zero_sub_struct : compatible_ms (@zero_sub_struct E).
-Proof.
-split; try apply: compatible_g_zero_sub_struct.
-intros u a Hu; rewrite Hu scal_zero_r; easy.
-Qed.
-
-Lemma compatible_ms_full :
-  forall {PE : E -> Prop}, full PE -> compatible_ms PE.
-Proof. intros; split; try apply compatible_g_full; easy. Qed.
-
-Lemma compatible_ms_fullset : compatible_ms (@fullset E).
-Proof. apply compatible_ms_full; easy. Qed.
-
-Lemma compatible_ms_inter :
-  forall {PE QE : E -> Prop},
-    compatible_ms PE -> compatible_ms QE -> compatible_ms (inter PE QE).
-Proof.
-intros PE QE [HPE1 HPE2] [HQE1 HQE2]; split.
-apply compatible_g_inter; easy.
-intros u a [Hu1 Hu2]; split; auto.
-Qed.
-
-(* Gostiaux T1, Th 6.15 p. 167. *)
-Lemma compatible_ms_inter_any :
-  forall {Idx : Type} {PE : Idx -> E -> Prop},
-    (forall i, compatible_ms (PE i)) -> compatible_ms (inter_any PE).
-Proof.
-intros Idx PE HPE; split.
-apply compatible_g_inter_any; intros; apply compatible_ms_g; easy.
-intros x l Hx i; apply HPE; easy.
-Qed.
-
-(* Gostiaux T1, Th 6.17 p. 168. *)
-Definition span_ms (gen : E -> Prop) := span_struct compatible_ms gen.
-
-Lemma span_ms_compatible_ms : forall gen, compatible_ms (span_ms gen).
-Proof.
-apply span_struct_compatible; move=>>; apply compatible_ms_inter_any.
-Qed.
-
-Lemma span_ms_nonempty : forall (gen : E -> Prop), nonempty (span_ms gen).
-Proof. intros; apply compatible_ms_nonempty, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_zero : forall (gen : E -> Prop), zero_closed (span_ms gen).
-Proof. intros; apply compatible_ms_zero; apply span_ms_compatible_ms. Qed.
-
-Lemma span_ms_plus : forall (gen : E -> Prop), plus_closed (span_ms gen).
-Proof. intros; apply compatible_ms_plus, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_opp : forall (gen : E -> Prop), opp_closed (span_ms gen).
-Proof. intros; apply compatible_ms_opp, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_minus : forall (gen : E -> Prop), minus_closed (span_ms gen).
-Proof. intros; apply compatible_ms_minus, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_scal : forall (gen : E -> Prop), scal_closed (span_ms gen).
-Proof. intros; apply compatible_ms_scal, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_comb_lin :
-  forall (gen : E -> Prop), comb_lin_closed (span_ms gen).
-Proof. intros; apply compatible_ms_comb_lin, span_ms_compatible_ms. Qed.
-
-Lemma span_ms_incl : forall gen, incl gen (span_ms gen).
-Proof. apply span_struct_incl. Qed.
-
-(* Gostiaux T1, Def 6.16 p. 167. *)
-Lemma span_ms_lub :
-  forall gen PE, compatible_ms PE -> incl gen PE -> incl (span_ms gen) PE.
-Proof. apply span_struct_lub. Qed.
-
-Lemma span_ms_full : forall PE, compatible_ms PE -> span_ms PE = PE.
-Proof. apply span_struct_full. Qed.
-
-End Compatible_ModuleSpace_Facts.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts1.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Context {PE1 : E1 -> Prop}.
-Context {PE2 : E2 -> Prop}.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : is_linear_mapping f.
-
-Lemma compatible_ms_image : compatible_ms PE1 -> compatible_ms (image f PE1).
-Proof.
-destruct Hf as [_ Hf1]; intros [HPE1a HPE1b]; split.
-apply compatible_g_image; [apply is_linear_mapping_morphism_g |]; easy.
-intros _ a [u Hu]; rewrite -Hf1; apply Im; auto.
-Qed.
-
-Lemma compatible_ms_preimage :
-  compatible_ms PE2 -> compatible_ms (preimage f PE2).
-Proof.
-destruct Hf as [_ Hf1]; intros [HPE2a HPE2b]; split.
-apply compatible_g_preimage; [apply is_linear_mapping_morphism_g |]; easy.
-intros u a Hu; unfold preimage; rewrite Hf1; auto.
-Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts1.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Def.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Variable PE1 : E1 -> Prop.
-
-Variable f : E1 -> E2.
-
-Definition f_scal_compat_sub : Prop :=
-  forall a x1 , PE1 x1 -> f (scal a x1) = scal a (f x1).
-
-Definition is_linear_mapping_sub : Prop :=
-  f_plus_compat_sub PE1 f /\ f_scal_compat_sub.
-
-End Compatible_ModuleSpace_Linear_Mapping_Def.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts2a.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Context {PE1 : E1 -> Prop}.
-Hypothesis HPE1 : compatible_ms PE1.
-
-Context {PE2 : E2 -> Prop}.
-Hypothesis HPE2 : compatible_ms PE2.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : is_linear_mapping f.
-
-Lemma KerS_compatible_ms : compatible_ms (KerS PE1 f).
-Proof.
-apply compatible_ms_inter, compatible_ms_preimage,
-    compatible_ms_zero_sub_struct; easy.
-Qed.
-
-Lemma RgS_gen_compatible_ms : compatible_ms (RgS_gen PE1 PE2 f).
-Proof. apply compatible_ms_inter, compatible_ms_image; easy. Qed.
-
-Lemma RgS_compatible_ms : compatible_ms (RgS PE1 f).
-Proof. apply compatible_ms_image; easy. Qed.
-
-Lemma KerS_zero_equiv :
-  KerS PE1 f = zero_sub_struct <-> incl (KerS PE1 f) zero_sub_struct.
-Proof.
-apply: KerS_g_zero_equiv;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_injS : incl (KerS PE1 f) zero_sub_struct -> injS PE1 f.
-Proof.
-apply: gmS_injS;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_injS_rev : injS PE1 f -> KerS PE1 f = zero_sub_struct.
-Proof.
-apply: gmS_injS_rev;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_injS_equiv : injS PE1 f <-> KerS PE1 f = zero_sub_struct.
-Proof.
-apply: gmS_injS_equiv;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_gen :
-  funS PE1 PE2 f -> incl (KerS PE1 f) zero_sub_struct -> incl PE2 (RgS PE1 f) ->
-  bijS PE1 PE2 f.
-Proof.
-apply: gmS_bijS_gen;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_gen_rev :
-  bijS PE1 PE2 f ->
-  funS PE1 PE2 f /\ KerS PE1 f = zero_sub_struct /\ RgS_gen PE1 PE2 f = PE2.
-Proof.
-apply: gmS_bijS_gen_rev;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_gen_equiv :
-  bijS PE1 PE2 f <->
-  funS PE1 PE2 f /\ KerS PE1 f = zero_sub_struct /\ RgS_gen PE1 PE2 f = PE2.
-Proof.
-apply: gmS_bijS_gen_equiv;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS : incl (KerS PE1 f) zero_sub_struct -> bijS PE1 (RgS PE1 f) f.
-Proof.
-apply: gmS_bijS;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_rev : bijS PE1 (RgS PE1 f) f -> KerS PE1 f = zero_sub_struct.
-Proof.
-apply: gmS_bijS_rev;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_equiv : bijS PE1 (RgS PE1 f) f <-> KerS PE1 f = zero_sub_struct.
-Proof.
-apply: gmS_bijS_equiv;
-    [apply compatible_ms_g | apply is_linear_mapping_morphism_g]; easy.
-Qed.
-
-Lemma lmS_bijS_injS_equiv : bijS PE1 (RgS PE1 f) f <-> injS PE1 f.
-Proof. apply gmS_bijS_injS_equiv. Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts2a.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts2b.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : is_linear_mapping f.
-
-Lemma Ker_compatible_ms : compatible_ms (Ker f).
-Proof.
-rewrite -KerS_full; apply (KerS_compatible_ms compatible_ms_fullset Hf).
-Qed.
-
-Lemma Rg_compatible_ms : compatible_ms (Rg f).
-Proof.
-rewrite -RgS_full; apply (RgS_compatible_ms compatible_ms_fullset Hf).
-Qed.
-
-Lemma Ker_zero_equiv :
-  Ker f = zero_sub_struct <-> incl (Ker f) zero_sub_struct.
-Proof.
-apply: Ker_g_zero_equiv; apply is_linear_mapping_morphism_g; easy.
-Qed.
-
-Lemma lm_inj : incl (Ker f) zero_sub_struct -> injective f.
-Proof. apply: gm_inj; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma lm_inj_rev : injective f -> Ker f = zero_sub_struct.
-Proof. apply: gm_inj_rev; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma lm_inj_equiv : injective f <-> Ker f = zero_sub_struct.
-Proof. apply: gm_inj_equiv; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma lm_bij :
-  incl (Ker f) zero_sub_struct -> incl fullset (Rg f) -> bijective f.
-Proof. apply: gm_bij; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma lm_bij_rev : bijective f -> Ker f = zero_sub_struct /\ Rg f = fullset.
-Proof. apply: gm_bij_rev; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma lm_bij_equiv : bijective f <-> Ker f = zero_sub_struct /\ Rg f = fullset.
-Proof. apply: gm_bij_equiv; apply is_linear_mapping_morphism_g; easy. Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts2b.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts3.
-
-Context {K : Ring}.
-Context {E1 E2 E3 : ModuleSpace K}.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : is_linear_mapping f.
-
-Context {g : E2 -> E3}.
-Hypothesis Hg : is_linear_mapping g.
-
-Lemma Ker_comp_r : injective g -> RgS (Ker (g \o f)) f = Ker g.
-Proof. apply: Ker_g_comp_r; apply is_linear_mapping_morphism_g; easy. Qed.
-
-Lemma Ker_comp_r_alt :
-  forall h x2, injective g -> h = g \o f ->
-    (exists x1, h x1 = 0 /\ f x1 = x2) <-> g x2 = 0.
-Proof.
-apply: Ker_g_comp_r_alt; apply is_linear_mapping_morphism_g; easy.
-Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts3.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts4a.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Variable PE1 : E1 -> Prop.
-
-Lemma f_scal_compat_sub_id : f_scal_compat_sub PE1 ssrfun.id.
-Proof. easy. Qed.
-
-Lemma is_linear_mapping_sub_id : is_linear_mapping_sub PE1 ssrfun.id.
-Proof. easy. Qed.
-
-Variable f : E1 -> E2.
-
-Lemma f_scal_compat_is_sub : f_scal_compat f -> f_scal_compat_sub PE1 f.
-Proof. easy. Qed.
-
-Lemma is_linear_mapping_is_sub :
-  is_linear_mapping f -> is_linear_mapping_sub PE1 f.
-Proof. intros [Hfa Hfb]; easy. Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts4a.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_Facts4b.
-
-Context {K : Ring}.
-Context {E1 E2 E3 : ModuleSpace K}.
-
-Context {PE1 : E1 -> Prop}.
-Variable PE2 : E2 -> Prop.
-
-Variable f : E1 -> E2.
-Variable g : E2 -> E3.
-
-Hypothesis Hf : funS PE1 PE2 f.
-
-Lemma f_scal_compat_comp_sub :
-  f_scal_compat_sub PE1 f -> f_scal_compat_sub PE2 g ->
-  f_scal_compat_sub PE1 (g \o f).
-Proof.
-intros Hf1 Hg1 a x1 Hx1; rewrite comp_correct Hf1// Hg1//; apply Hf; easy.
-Qed.
-
-Lemma is_linear_mapping_comp_sub :
-  is_linear_mapping_sub PE1 f -> is_linear_mapping_sub PE2 g ->
-  is_linear_mapping_sub PE1 (g \o f).
-Proof.
-intros [Hfa Hfb] [Hga Hgb]; split;
-    [apply (f_plus_compat_comp_sub PE2) | apply f_scal_compat_comp_sub]; easy.
-Qed.
-
-End Compatible_ModuleSpace_Linear_Mapping_Facts4b.
-
-
-Section Sub_ModuleSpace_Def.
-
-Context {K : Ring}.
-Context {E : ModuleSpace K}.
-Context {PE : E -> Prop}.
-
-Definition sub_ms (HPE : compatible_ms PE) : Type :=
-  sub_g (compatible_ms_g HPE).
-Definition mk_sub_ms_ (HPE : compatible_ms PE) {x} (Hx : PE x) : sub_ms HPE :=
-  mk_sub_g_ (compatible_ms_g HPE) Hx.
-Definition mk_sub_ms {HPE : compatible_ms PE} {x} (Hx : PE x) : sub_ms HPE :=
-  mk_sub_ms_ HPE Hx.
-
-Hypothesis HPE : compatible_ms PE.
-
-Definition sub_scal a (x : sub_ms HPE) : sub_ms HPE :=
-  mk_sub_ms (compatible_ms_scal HPE a (val x) (in_sub x)).
-
-Lemma sub_scal_assoc :
-  forall a b x, sub_scal a (sub_scal b x) = sub_scal (a * b) x.
-Proof. intros; apply val_inj, scal_assoc. Qed.
-
-Lemma sub_scal_one_l : forall x, sub_scal 1 x = x.
-Proof. intros; apply val_inj, scal_one. Qed.
-
-Lemma sub_scal_distr_l :
-  forall a x y, sub_scal a (x + y) = sub_scal a x + sub_scal a y.
-Proof. intros; apply val_inj, scal_distr_l. Qed.
-
-Lemma sub_scal_distr_r  :
-  forall a b x, sub_scal (a + b) x = sub_scal a x + sub_scal b x.
-Proof. intros; apply val_inj, scal_distr_r. Qed.
-
-Definition sub_ModuleSpace_mixin :=
-  ModuleSpace.Mixin _ _ _
-    sub_scal_assoc sub_scal_one_l sub_scal_distr_l sub_scal_distr_r.
-
-Canonical Structure sub_ModuleSpace :=
-  ModuleSpace.Pack _ _
-    (ModuleSpace.Class _ _ _ sub_ModuleSpace_mixin) (sub_ms HPE).
-
-Lemma val_scal : f_scal_compat val.
-Proof. easy. Qed.
-
-Lemma val_comb_lin : f_comb_lin_compat val.
-Proof.
-intros n; induction n as [| n Hn].
-intros; rewrite 2!comb_lin_nil; easy.
-intros; rewrite 2!comb_lin_ind_l; simpl; f_equal; apply Hn.
-Qed.
-
-Lemma val_lin_map : is_linear_mapping val.
-Proof. easy. Qed.
-
-Lemma mk_sub_scal :
-  forall a (x : E) (Hx : PE x),
-    scal a (mk_sub_ms Hx) = mk_sub_ms (compatible_ms_scal HPE a _ Hx).
-Proof. easy. Qed.
-
-Lemma mk_sub_comb_lin :
-  forall {n} L (B : 'E^n) (HB : inclF B PE),
-    comb_lin L (fun i => mk_sub_ms (HB i)) =
-      mk_sub_ms (compatible_ms_comb_lin HPE _ L _ HB).
-Proof. intros; apply val_inj, val_comb_lin. Qed.
-
-Lemma sub_scal_eq :
-  forall a (x : sub_ms HPE),
-    scal a x = mk_sub_ms (compatible_ms_scal HPE a _ (in_sub x)).
-Proof. easy. Qed.
-
-Lemma sub_comb_lin_eq :
-  forall {n} L (B : '(sub_ms HPE)^n),
-    comb_lin L B =
-      mk_sub_ms (compatible_ms_comb_lin HPE _ L _ (fun i => in_sub (B i))).
-Proof. intros; apply mk_sub_comb_lin. Qed.
-
-Lemma mk_sub_ms_inj :
-  forall {x y} (Hx : PE x) (Hy : PE y),
-    mk_sub_ms Hx = mk_sub_ms Hy :> sub_ms HPE -> x = y.
-Proof. apply mk_sub_g_inj. Qed.
-
-End Sub_ModuleSpace_Def.
-
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (30/05/2023):
-  Define sum of sub_ms (Gostiaux T1, Th 6.20 p. 169),
-         direct sum, supplementary (Gostiaux T1, S 6.5 pp. 182-187). *)
-
-
-Section Sub_ModuleSpace_Facts1.
-
-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 RgS_val_ms_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_ms : compatible_ms PEa -> compatible_ms PEa'.
-Proof. intros; apply compatible_ms_preimage; easy. Qed.
-
-Lemma preimage_val_compatible_ms_rev : compatible_ms PEa' -> compatible_ms PEa.
-Proof. intros; rewrite -RgS_val_ms_eq; apply RgS_compatible_ms; easy. Qed.
-
-Lemma preimage_val_compatible_ms_equiv : compatible_ms PEa' <-> compatible_ms PEa.
-Proof.
-split; [apply preimage_val_compatible_ms_rev | apply preimage_val_compatible_ms].
-Qed.
-
-End Sub_ModuleSpace_Facts1.
-
-
-Section Sub_ModuleSpace_Facts2.
-
-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_val_ms_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_ms : compatible_ms PEa -> compatible_ms PEa'.
-Proof. intros; apply RgS_compatible_ms; easy. Qed.
-
-Lemma RgS_val_compatible_ms_rev : compatible_ms PEa' -> compatible_ms PEa.
-Proof.
-intros; rewrite -preimage_val_ms_eq; apply compatible_ms_preimage; easy.
-Qed.
-
-Lemma RgS_val_compatible_ms_equiv : compatible_ms PEa' <-> compatible_ms PEa.
-Proof.
-split; [apply RgS_val_compatible_ms_rev | apply RgS_val_compatible_ms].
-Qed.
-
-End Sub_ModuleSpace_Facts2.
-
-
-Section Sub_ModuleSpace_Linear_Mapping_Facts1.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Context {PE1 : E1 -> Prop}.
-Context {PE2 : E2 -> Prop}.
-Hypothesis HPE1 : compatible_ms PE1.
-Hypothesis HPE2 : compatible_ms PE2.
-
-Context {f : E1 -> E2}.
-Context {fS : sub_ms HPE1 -> sub_ms HPE2}.
-Hypothesis HfS : forall x, val (fS x) = f (val x).
-
-Lemma sub_ms_fun_rev : funS PE1 PE2 f.
-Proof. intros _ [x1 Hx1]; rewrite -(HfS (mk_sub Hx1)); apply in_sub. Qed.
-
-Lemma sub_ms_inj : injS PE1 f -> injective fS.
-Proof. apply sub_g_inj, HfS. Qed.
-
-Lemma sub_ms_inj_rev : injective fS -> injS PE1 f.
-Proof. apply sub_g_inj_rev, HfS. Qed.
-
-Lemma sub_ms_inj_equiv : injective fS <-> injS PE1 f.
-Proof. apply sub_g_inj_equiv, HfS. Qed.
-
-Lemma sub_ms_surj : surjS PE1 PE2 f -> surjective fS.
-Proof. apply sub_g_surj, HfS. Qed.
-
-Lemma sub_ms_surj_rev : surjective fS -> surjS PE1 PE2 f.
-Proof. apply sub_g_surj_rev, HfS. Qed.
-
-Lemma sub_ms_surj_equiv : surjective fS <-> surjS PE1 PE2 f.
-Proof. apply sub_g_surj_equiv, HfS. Qed.
-
-Lemma sub_ms_bij : bijS PE1 PE2 f -> bijective fS.
-Proof. apply sub_g_bij, HfS. Qed.
-
-Lemma sub_ms_bij_rev : bijective fS -> bijS PE1 PE2 f.
-Proof. apply sub_g_bij_rev, HfS. Qed.
-
-Lemma sub_ms_bij_equiv : bijective fS <-> bijS PE1 PE2 f.
-Proof. apply sub_g_bij_equiv, HfS. Qed.
-
-Lemma sub_ms_f_scal_compat : f_scal_compat f -> f_scal_compat fS.
-Proof. intros Hf a x1; apply val_inj; rewrite HfS Hf -!HfS; easy. Qed.
-
-Lemma sub_ms_f_comb_lin_compat : f_comb_lin_compat f -> f_comb_lin_compat fS.
-Proof.
-intros Hf n L B; apply val_inj; rewrite HfS !val_comb_lin Hf; f_equal.
-apply extF; intros i; rewrite !mapF_correct; easy.
-Qed.
-
-Lemma sub_ms_linear_mapping : is_linear_mapping f -> is_linear_mapping fS.
-Proof.
-intros [Hf1 Hf2]; split; [apply (sub_g_morphism HfS Hf1) |
-    move=>>; apply sub_ms_f_scal_compat; move=>>; easy].
-Qed.
-
-Lemma Ker_sub_ms_KerS_zero_equiv :
-  is_linear_mapping f ->
-  Ker fS = zero_sub_struct <-> KerS PE1 f = zero_sub_struct.
-Proof.
-intros Hf; apply (Ker_sub_g_KerS_zero_equiv HfS),
-    is_linear_mapping_morphism_g, Hf.
-Qed.
-
-Lemma KerS_ms_zero_equiv_alt :
-  is_linear_mapping f ->
-  KerS PE1 f = zero_sub_struct <->
-  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros H; apply: KerS_g_zero_equiv_alt; apply is_linear_mapping_morphism_g, H.
-Qed.
-
-Lemma Ker_sub_ms_zero_equiv :
-  is_linear_mapping f ->
-  Ker fS = zero_sub_struct <->
-  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros Hf; apply (Ker_sub_g_zero_equiv HfS), is_linear_mapping_morphism_g, Hf.
-Qed.
-
-Lemma lmS_injS_sub_equiv :
-  is_linear_mapping f -> injS PE1 f <-> Ker fS = zero_sub_struct.
-Proof.
-intros Hf; apply (gmS_injS_sub_equiv HfS), is_linear_mapping_morphism_g, Hf.
-Qed.
-
-Lemma lmS_injS_sub_equiv_alt :
-  is_linear_mapping f ->
-  injS PE1 f <-> forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros Hf; apply: (gmS_injS_sub_equiv_alt HfS);
-    apply is_linear_mapping_morphism_g, Hf.
-Qed.
-
-Lemma lmS_bijS_sub_equiv :
-  is_linear_mapping f -> bijS PE1 (RgS PE1 f) f <-> Ker fS = zero_sub_struct.
-Proof.
-intros Hf; apply (gmS_bijS_sub_equiv HfS), is_linear_mapping_morphism_g, Hf.
-Qed.
-
-Lemma lmS_bijS_sub_equiv_alt :
-  is_linear_mapping f ->
-  bijS PE1 (RgS PE1 f) f <-> forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
-Proof.
-intros Hf; apply: (gmS_bijS_sub_equiv_alt HfS);
-    apply is_linear_mapping_morphism_g, Hf.
-Qed.
-
-End Sub_ModuleSpace_Linear_Mapping_Facts1.
-
-
-Section Sub_ModuleSpace_Linear_Mapping_Facts2.
-
-Context {K : Ring}.
-Context {E1 E2 : ModuleSpace K}.
-
-Context {PE1 : E1 -> Prop}.
-Context {PE2 : E2 -> Prop}.
-Hypothesis HPE1 : compatible_ms PE1.
-Hypothesis HPE2 : compatible_ms PE2.
-
-Context {f : E1 -> E2}.
-Hypothesis Hf : funS PE1 PE2 f.
-
-Definition fct_sub_ms : sub_ms HPE1 -> sub_ms HPE2 :=
-  fct_sub_g (compatible_ms_g HPE1) (compatible_ms_g HPE2) Hf.
-
-Lemma fct_sub_ms_inj : injS PE1 f -> injective fct_sub_ms.
-Proof. apply fct_sub_g_inj. Qed.
-
-Lemma fct_sub_ms_inj_rev : injective fct_sub_ms -> injS PE1 f.
-Proof. apply fct_sub_g_inj_rev. Qed.
-
-Lemma fct_sub_ms_inj_equiv : injective fct_sub_ms <-> injS PE1 f.
-Proof. apply fct_sub_g_inj_equiv. Qed.
-
-Lemma fct_sub_ms_surj : surjS PE1 PE2 f -> surjective fct_sub_ms.
-Proof. apply fct_sub_g_surj. Qed.
-
-Lemma fct_sub_ms_surj_rev : surjective fct_sub_ms -> surjS PE1 PE2 f.
-Proof. apply fct_sub_g_surj_rev. Qed.
-
-Lemma fct_sub_ms_surj_equiv : surjective fct_sub_ms <-> surjS PE1 PE2 f.
-Proof. apply fct_sub_g_surj_equiv. Qed.
-
-Lemma fct_sub_ms_bij : bijS PE1 PE2 f -> bijective fct_sub_ms.
-Proof. apply fct_sub_g_bij. Qed.
-
-Lemma fct_sub_ms_bij_rev : bijective fct_sub_ms -> bijS PE1 PE2 f.
-Proof. apply fct_sub_g_bij_rev. Qed.
-
-Lemma fct_sub_ms_bij_equiv : bijective fct_sub_ms <-> bijS PE1 PE2 f.
-Proof. apply fct_sub_g_bij_equiv. Qed.
-
-Lemma fct_sub_ms_f_scal_compat : f_scal_compat f -> f_scal_compat fct_sub_ms.
-Proof. apply sub_ms_f_scal_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_ms_f_comb_lin_compat :
-  f_comb_lin_compat f -> f_comb_lin_compat fct_sub_ms.
-Proof. apply sub_ms_f_comb_lin_compat, fct_sub_correct. Qed.
-
-Lemma fct_sub_ms_linear_mapping :
-  is_linear_mapping f -> is_linear_mapping fct_sub_ms.
-Proof. apply sub_ms_linear_mapping, fct_sub_correct. Qed.
-
-Lemma Ker_fct_sub_ms_KerS_zero_equiv :
-  is_linear_mapping f ->
-  Ker fct_sub_ms = zero_sub_struct <-> KerS PE1 f = zero_sub_struct.
-Proof. apply Ker_sub_ms_KerS_zero_equiv, fct_sub_correct. Qed.
-
-Lemma Ker_fct_sub_ms_zero_equiv :
-  is_linear_mapping f ->
-  Ker fct_sub_ms = zero_sub_struct <->
-  forall (x1 : sub_ms HPE1), f (val x1) = 0 -> x1 = 0.
-Proof. apply Ker_sub_ms_zero_equiv, fct_sub_correct. Qed.
-
-Lemma lmS_injS_fct_sub_equiv :
-  is_linear_mapping f -> injS PE1 f <-> Ker fct_sub_ms = zero_sub_struct.
-Proof. apply lmS_injS_sub_equiv, fct_sub_correct. Qed.
-
-Lemma lmS_bijS_fct_sub_equiv :
-  is_linear_mapping f ->
-  bijS PE1 (RgS PE1 f) f <-> Ker fct_sub_ms = zero_sub_struct.
-Proof. apply lmS_bijS_sub_equiv, fct_sub_correct. Qed.
-
-End Sub_ModuleSpace_Linear_Mapping_Facts2.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_R_Facts1.
-
-Variable E1 E2 : ModuleSpace R_Ring.
-
-Lemma compatible_ms_linear_mapping : compatible_ms (@is_linear_mapping _ E1 E2).
-Proof.
-split; [apply compatible_g_is_linear_mapping |].
-intros; apply is_linear_mapping_f_scal; easy.
-Qed.
-
-(* Lm is the vector subspace of linear mappings. *)
-Definition Lm := sub_ms compatible_ms_linear_mapping.
-
-End Compatible_ModuleSpace_Linear_Mapping_R_Facts1.
-
-
-Section Compatible_ModuleSpace_Linear_Mapping_R_Facts2.
-
-Variable E : ModuleSpace R_Ring.
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (25/11/2023): define ring of endomorphisms (with + and \o). *)
-
-
-
-Definition Endom : (E -> E) -> Prop := fun f => is_linear_mapping f.
-
-(* TO BE REMOVED FOR PUBLICATION!
- TODO FC (25/11/2023): define non-Abelian group of automorphisms (with \o). *)
-
-Definition Autom : (E -> E) -> Prop :=
-  fun f => is_linear_mapping f /\ bijective f.
-
-End Compatible_ModuleSpace_Linear_Mapping_R_Facts2.
-
-
-Local Open Scope AffineSpace_scope.
-
-
-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.
-
-
-(* 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.
-
-Lemma PE_as_inhabited : inhabited (sub_struct HPE).
-Proof. apply sub_struct_inhabited; exists O; easy. Qed.
-
-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 _ _ _ PE_as_inhabited _ 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/_CoqProject b/_CoqProject
index 7421132aeede6ae44b0a2d41d6eb53cb44fb95b5..4306e85658e527e64d4981b562cd3791982842c0 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -103,6 +103,7 @@
 ./FEM/Algebra/ord_compl.v
 ./FEM/Algebra/Finite_family.v
 ./FEM/Algebra/Finite_table.v
+./FEM/Algebra/Sub_struct.v
 ./FEM/Algebra/Monoid_compl.v
 ./FEM/Algebra/Group_compl.v
 ./FEM/Algebra/GroupMult.v
@@ -111,9 +112,8 @@
 ./FEM/Algebra/ModuleSpace_compl.v
 ./FEM/Algebra/ModuleSpace_R_compl.v
 ./FEM/Algebra/matrix.v
-./FEM/Algebra/AffineSpace.v
-./FEM/Algebra/Sub_struct.v
 ./FEM/Algebra/MonoidComp.v
+./FEM/Algebra/AffineSpace.v
 ./FEM/Algebra/Finite_dim.v
 ./FEM/Algebra/Finite_dim_R.v
 ./FEM/Algebra/Algebra.v