Theory Algebra7

theory Algebra7
imports Algebra6
(**        Algebra7  
                            author Hidetsune Kobayashi
                            Group You Santo
                            Department of Mathematics
                            Nihon University
                            h_koba@math.cst.nihon-u.ac.jp
                            May 3, 2004.
                            April 6, 2007 (revised)

   chapter 5. Modules
    section 3.   a module over two rings 
    section 4.   eSum and Generators
     subsection 4-1. sum up coefficients
     subsection 4-2. free generators 
   **)

theory Algebra7 imports Algebra6 begin

chapter "Modules"

section "Basic properties of Modules"

record ('a, 'b) Module = "'a aGroup" +
  sprod  :: "'b ⇒ 'a ⇒ 'a" (infixl "⋅sı" 76)

locale Module = aGroup M for M (structure) +
  fixes R (structure)
  assumes  sc_Ring: "Ring R" 
  and  sprod_closed :
      "⟦ a ∈ carrier R; m ∈ carrier M⟧ ⟹ a ⋅s m ∈ carrier M" 
    and sprod_l_distr:
      "⟦a ∈ carrier R; b ∈ carrier R; m ∈ carrier M⟧ ⟹
       (a ±R b) ⋅s m = a ⋅s m ±M b ⋅s m" 
    and sprod_r_distr:
      "⟦ a ∈ carrier R; m ∈ carrier M; n ∈ carrier M ⟧ ⟹
      a ⋅s (m ±M n) = a ⋅s m ±M a ⋅s n"
    and sprod_assoc:
      "⟦ a ∈ carrier R; b ∈ carrier R; m ∈ carrier M ⟧ ⟹
      (a ⋅rR b) ⋅s m = a ⋅s (b ⋅s m)"  
    and sprod_one:
      "m ∈ carrier M ⟹ (1rR) ⋅s m = m" 

definition 
  submodule :: "[('b, 'm) Ring_scheme, ('a, 'b, 'c) Module_scheme, 'a set] ⇒
            bool" where
  "submodule R A H ⟷ H ⊆ carrier A ∧ A +> H ∧ (∀a. ∀m. 
                     (a ∈ carrier R ∧ m ∈ H) ⟶ (sprod A a m) ∈ H)"

definition
  mdl :: "[('a, 'b, 'm) Module_scheme, 'a set] ⇒ ('a, 'b) Module" where
  "mdl M H = ⦇carrier = H, pop = pop M, mop = mop M, zero = zero M,
    sprod = λa. λx∈H. sprod M a x⦈" 

abbreviation
  MODULE  (infixl "module" 58) where
 "R module M == Module M R"
 

lemma (in Module) module_is_ag: "aGroup M" ..

lemma (in Module) module_inc_zero:" 𝟬M ∈ carrier M"
apply (simp add:ag_inc_zero) (** type of M is ('c, 'a, 'd) Module_scheme **)
done                         (** type of M is (?'b, ?'b, ?'z) Module_scheme **)

lemma (in Module) submodule_subset:"submodule R M H ⟹ H ⊆ carrier M"
apply (simp add:submodule_def)
done

lemma (in Module) submodule_asubg:"submodule R M H ⟹ M +> H"
by (simp add:submodule_def)

lemma (in Module) submodule_subset1:"⟦submodule R M H; h ∈ H⟧ ⟹
                            h ∈ carrier M"
apply (simp add:submodule_def)
apply (erule conjE)+
apply (simp add:subsetD)
done

lemma (in Module) submodule_inc_0:"submodule R M H ⟹
                                           𝟬M ∈ H" 
apply (simp add:submodule_def, (erule conjE)+)
apply (rule asubg_inc_zero, assumption+)
done

lemma (in Module) sc_un:" m ∈ carrier M ⟹ 1rRs m = m"
apply (simp add:sprod_one)
done

lemma (in Module) sc_mem:"⟦a ∈ carrier R; m ∈ carrier M⟧ ⟹
           a ⋅s m ∈ carrier M"
apply (simp add:sprod_closed)
done

lemma (in Module) submodule_sc_closed:"⟦submodule R M H; 
 a ∈ carrier R; h ∈ H⟧ ⟹  a ⋅s h ∈ H"
apply (simp add:submodule_def)
done

lemma (in Module) sc_assoc:"⟦a ∈ carrier R; b ∈ carrier R; 
 m ∈ carrier M⟧ ⟹ (a ⋅rR b) ⋅s m =  a ⋅s ( b ⋅s m)"
apply (simp add:sprod_assoc)
done

lemma (in Module) sc_l_distr:"⟦a ∈ carrier R; b ∈ carrier R; 
 m ∈ carrier M⟧ ⟹ (a ±R b)⋅s m = a ⋅s m ±  b ⋅s m"
apply (simp add:sprod_l_distr)
done

lemma (in Module) sc_r_distr:"⟦a ∈ carrier R; m ∈ carrier M; n ∈ carrier M⟧ ⟹
                 a ⋅s (m ± n) = a ⋅s m ±  a ⋅s n"
apply (simp add:sprod_r_distr)
done

lemma (in Module) sc_0_m:"m ∈ carrier M ⟹ 𝟬Rs m = 𝟬M"
apply (cut_tac sc_Ring,
       frule Ring.ring_is_ag,
       frule aGroup.ag_inc_zero [of "R"],
       frule sc_l_distr[of "𝟬R" "𝟬R" "m"], assumption+,
       frule sc_mem [of "𝟬R" m], assumption+)
apply (simp add:aGroup.ag_l_zero, frule sym,
       thin_tac "𝟬Rs m = 𝟬Rs m ± 𝟬Rs m")
apply (frule ag_eq_sol1 [of "𝟬Rs m" "𝟬Rs m" "𝟬Rs m"], assumption+,   
       simp add:ag_l_inv1)
done

lemma (in Module) sc_a_0:"a ∈ carrier R ⟹ a ⋅s 𝟬  = 𝟬"
apply (cut_tac ag_inc_zero,
       frule sc_r_distr[of a 𝟬 𝟬], assumption+,
       frule sc_mem [of a 𝟬], assumption+)
apply (simp add:ag_l_zero, frule sym,
       thin_tac "a ⋅s 𝟬 = a ⋅s 𝟬 ± a ⋅s 𝟬")
apply (frule ag_eq_sol1 [of "a ⋅s 𝟬" "a ⋅s 𝟬" "a ⋅s 𝟬"], assumption+,   
       simp add:ag_l_inv1)
done

lemma (in Module) sc_minus_am:"⟦a ∈ carrier R; m ∈ carrier M⟧
                     ⟹ -a (a ⋅s m) = a ⋅s (-a m)"
apply (frule ag_mOp_closed [of m],
       frule sc_r_distr[of a m "-a m"], assumption+,
       simp add:ag_r_inv1,
       simp add:sc_a_0, frule sym,
       thin_tac "𝟬 = a ⋅s m ± a ⋅s (-a m)")
 apply (frule sc_mem [of a m], assumption+,
        frule sc_mem [of a "-a m"], assumption+,
        frule ag_eq_sol1 [of "a ⋅s m" "a ⋅s (-a m)" "𝟬"], assumption+,
        simp add:ag_inc_zero, assumption)
 apply (frule ag_mOp_closed [of "a ⋅s m"],
        simp add:ag_r_zero)
done

lemma (in Module) sc_minus_am1:"⟦a ∈ carrier R; m ∈ carrier M⟧
            ⟹ -a (a ⋅s m) = (-aR a) ⋅s m"
apply (cut_tac sc_Ring, frule Ring.ring_is_ag,
       frule aGroup.ag_mOp_closed [of R a], assumption+,
       frule sc_l_distr[of a "-aR a" m], assumption+,
       simp add:aGroup.ag_r_inv1 [of "R"],
       simp add:sc_0_m, frule sym) apply (
       thin_tac "𝟬 = a ⋅s m ± (-aR a) ⋅s m")
 apply (frule sc_mem [of a m], assumption+,
        frule sc_mem [of "-aR a" m], assumption+)
 apply (frule ag_eq_sol1 [of "a ⋅s m" "(-aR a) ⋅s m" 𝟬], assumption+,
        simp add:ag_inc_zero, assumption)
 apply (frule ag_mOp_closed [of "a ⋅s m"])
 apply (thin_tac "a ⋅s m ± (-aR a) ⋅s m = 𝟬",
        simp add:ag_r_zero)
done

lemma (in Module) submodule_0:"submodule R M {𝟬}" 
apply (simp add:submodule_def)
apply (simp add:ag_inc_zero)
apply (simp add:asubg_zero)
apply (rule allI, rule impI)
apply (simp add:sc_a_0)
done   

lemma (in Module) submodule_whole:"submodule R M (carrier M)" 
apply (simp add:submodule_def)
apply (simp add:asubg_whole)
apply ((rule allI)+, rule impI, erule conjE)
apply (simp add:sc_mem)
done

lemma (in Module) submodule_pOp_closed:"⟦submodule R M H; h ∈ H; k ∈ H⟧ ⟹ 
                  h ± k ∈ H"
apply (simp add:submodule_def)
apply (erule conjE)+
apply (thin_tac "∀a m. a ∈ carrier R ∧ m ∈ H ⟶ a ⋅s m ∈ H")
apply (simp add:asubg_pOp_closed)
done

lemma (in Module) submodule_mOp_closed:"⟦submodule R M H; h ∈ H⟧
                 ⟹ -a h ∈ H"
apply (simp add:submodule_def,
       (erule conjE)+,
       thin_tac "∀a m. a ∈ carrier R ∧ m ∈ H ⟶ a ⋅s m ∈ H")
apply (rule asubg_mOp_closed, assumption+)
done 

definition
  mHom :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
                    ('c, 'b, 'm2) Module_scheme] ⇒  ('a ⇒ 'c) set"
        (*  ("(3HOM_/ _/ _)" [90, 90, 91]90 ) *) where
  "mHom R M N = {f. f ∈ aHom M N ∧ 
             (∀a∈carrier R. ∀m∈carrier M. f (a ⋅sM m) = a ⋅sN (f m))}"

definition
  mimg :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
           ('c, 'b, 'm2) Module_scheme, 'a ⇒ 'c] ⇒  ('c, 'b) Module" 
                 ("(4mimg_ _,_/ _)" [88,88,88,89]88) where
  "mimgR M,N f = mdl N (f ` (carrier M))"

definition
  mzeromap :: "[('a, 'b, 'm1) Module_scheme, ('c, 'b, 'm2) Module_scheme]
                              ⇒ ('a ⇒ 'c)" where
  "mzeromap M N = (λx∈carrier M. 𝟬N)"

lemma (in Ring) mHom_func:"f ∈ mHom R M N ⟹ f ∈ carrier M → carrier N"
by (simp add:mHom_def aHom_def)

lemma (in Module) mHom_test:"⟦R module N; f ∈ carrier M → carrier N ∧ 
      f ∈ extensional (carrier M) ∧ 
     (∀m∈carrier M. ∀n∈carrier M. f (m ±M n) = f m ±N (f n)) ∧ 
     (∀a∈carrier R. ∀m∈carrier M. f (a ⋅sM m) = a ⋅sN (f m))⟧ ⟹
     f ∈ mHom R M N"  
apply (simp add:mHom_def)
apply (simp add:aHom_def)
done

lemma (in Module) mHom_mem:"⟦R module N; f ∈ mHom R M N; m ∈ carrier M⟧
 ⟹ f m ∈ carrier N"
apply (simp add:mHom_def aHom_def) apply (erule conjE)+
apply (simp add:Pi_def)
done

lemma (in Module) mHom_add:"⟦R module N; f ∈ mHom R M N; m ∈ carrier M; 
             n ∈ carrier M⟧ ⟹ f (m ± n) = f m ±N (f n)"
apply (simp add:mHom_def) apply (erule conjE)+
apply (frule Module.module_is_ag [of N R],
       cut_tac module_is_ag)
apply (simp add:aHom_add)
done 
 
lemma (in Module) mHom_0:"⟦R module N; f ∈ mHom R M N⟧ ⟹ f (𝟬) = 𝟬N"
apply (simp add:mHom_def, (erule conjE)+,
       frule Module.module_is_ag [of N],
       cut_tac module_is_ag)
apply (simp add:aHom_0_0)
done

lemma (in Module) mHom_inv:"⟦R module N; m ∈ carrier M; f ∈ mHom R M N⟧ ⟹ 
                 f (-a m) = -aN (f m)"
apply (cut_tac module_is_ag,
       frule Module.module_is_ag [of N])
apply (simp add:mHom_def, (erule conjE)+)
apply (rule aHom_inv_inv, assumption+)
done

lemma (in Module) mHom_lin:"⟦R module N; m ∈ carrier M; f ∈ mHom R M N;
                    a ∈ carrier R⟧ ⟹ f (a ⋅s m) = a ⋅sN (f m)"
apply (simp add:mHom_def)
done

lemma (in Module) mker_inc_zero:
           "⟦R module N; f ∈ mHom R M N ⟧ ⟹ 𝟬 ∈ (kerM,N f)" 
apply (simp add:ker_def) 
apply (simp add:module_inc_zero)
apply (simp add:mHom_0)
done

lemma (in Module) mHom_eq_ker:"⟦R module N; f ∈ mHom R M N; a ∈ carrier M; 
      b∈ carrier M; a ± (-a b) ∈ kerM,N f⟧ ⟹ f a = f b"
apply (simp add:ker_def, erule conjE)
apply (cut_tac module_is_ag,
       frule aGroup.ag_mOp_closed [of "M" "b"], assumption+,
       simp add:mHom_add, simp add:mHom_inv,
       thin_tac "aGroup M")
apply (frule mHom_mem [of N f a], assumption+,
       frule mHom_mem [of N f b], assumption+,
       frule Module.module_is_ag[of N]) 
apply (subst aGroup.ag_eq_diffzero[of N], assumption+)
done  

lemma (in Module) mHom_ker_eq:"⟦R module N; f ∈ mHom R M N; a ∈ carrier M; 
      b∈ carrier M; f a = f b⟧ ⟹ a ± (-a b) ∈ kerM,N f"
apply (simp add:ker_def)
 apply (frule ag_mOp_closed[of b])
 apply (simp add:ag_pOp_closed)
 apply (simp add:mHom_add mHom_inv)
 apply (frule mHom_mem [of N f b], assumption+)
 apply (frule_tac R = R and M = N in Module.module_is_ag,
        simp add:aGroup.ag_r_inv1)
done
 
lemma (in Module) mker_submodule:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                                    submodule R M (kerM,N f)"
apply (cut_tac module_is_ag,
       frule Module.module_is_ag [of N])
apply (simp add:submodule_def)
apply (rule conjI)
 apply (rule subsetI, simp add:ker_def)

apply (rule conjI)
 apply (simp add:mHom_def, (erule conjE)+, simp add:ker_subg)

apply ((rule allI)+, rule impI, erule conjE)
 apply (simp add:ker_def, erule conjE)
 apply (simp add:sc_mem)
 apply (subst mHom_lin [of N _ f], assumption+, simp) (* key *)
apply (simp add:Module.sc_a_0[of N])
done

lemma (in Module) mker_mzeromap:"R module N ⟹
                         kerM,N (mzeromap M N) = carrier M"
apply (simp add:ker_def mzeromap_def)
done

lemma (in Module) mdl_carrier:"submodule R M H ⟹ carrier (mdl M H) = H"
apply (simp add:mdl_def)
done 

lemma (in Module) mdl_is_ag:"submodule R M H ⟹ aGroup (mdl M H)"
apply (cut_tac module_is_ag)
apply (rule aGroup.intro)
 apply (simp add:mdl_def)
 apply (clarsimp simp: submodule_def asubg_pOp_closed)

 apply (simp add:mdl_def)
 apply (simp add:submodule_def, (erule conjE)+,
        frule_tac c = a in subsetD[of H "carrier M"], assumption+,
        frule_tac c = b in subsetD[of H "carrier M"], assumption+,
        frule_tac c = c in subsetD[of H "carrier M"], assumption+,
        simp add:aGroup.ag_pOp_assoc)

 apply (simp add:submodule_def, (erule conjE)+,
        simp add:mdl_def,
        frule_tac c = a in subsetD[of H "carrier M"], assumption+,
        frule_tac c = b in subsetD[of H "carrier M"], assumption+,
        simp add:aGroup.ag_pOp_commute)

 apply (simp add:mdl_def)
 apply (simp add:submodule_def aGroup.asubg_mOp_closed)

 apply (simp add:mdl_def,
        simp add:submodule_def, (erule conjE)+,
        frule_tac c = a in subsetD[of H "carrier M"], assumption+,
        rule aGroup.ag_l_inv1, assumption+)         

 apply (simp add:mdl_def,
        simp add:submodule_def, (erule conjE)+,
        simp add:asubg_inc_zero)

 apply (simp add:mdl_def, simp add:submodule_def, (erule conjE)+,
        frule_tac c = a in subsetD[of H "carrier M"], assumption+)
 apply (simp add:ag_l_zero)
done

lemma (in Module) mdl_is_module:"submodule R M H ⟹ R module (mdl M H)" 
apply (rule Module.intro)
apply (simp add:mdl_is_ag)

apply (rule Module_axioms.intro)
apply (simp add:sc_Ring)

apply (simp add:mdl_def)
 apply (simp add:submodule_def) 

apply (simp add:mdl_def)
 apply (simp add:submodule_def, (erule conjE)+,
        frule_tac c = m in subsetD[of H "carrier M"], assumption+,
        simp add:sc_l_distr)

apply (simp add:mdl_def submodule_def, (erule conjE)+,
       simp add:asubg_pOp_closed,
       frule_tac c = m in subsetD[of H "carrier M"], assumption+,
       frule_tac c = n in subsetD[of H "carrier M"], assumption+,
       simp add:sc_r_distr)
apply (simp add:mdl_def submodule_def, (erule conjE)+,
       frule_tac c = m in subsetD[of H "carrier M"], assumption+,
       simp add:sc_assoc)
apply (simp add:mdl_def submodule_def, (erule conjE)+,
       frule_tac c = m in subsetD[of H "carrier M"], assumption+,
       simp add:sprod_one)
done   

lemma (in Module) submodule_of_mdl:"⟦submodule R M H; submodule R M N; H ⊆ N⟧
                   ⟹ submodule R (mdl M N) H"
apply (subst submodule_def)
 apply (rule conjI, simp add:mdl_def)
 apply (rule conjI)
 apply (rule aGroup.asubg_test[of "mdl M N" H])
 apply (frule mdl_is_module[of N],
        simp add:Module.module_is_ag, simp add:mdl_def)
 apply (simp add:submodule_def[of R M H], (erule conjE)+)
 apply (frule asubg_inc_zero[of H], simp add:nonempty)

 apply ((rule ballI)+, simp add:mdl_def)
 apply (simp add:submodule_def[of R M H], (erule conjE)+)
 apply (frule_tac x = b in asubg_mOp_closed[of H], assumption+)
 apply (rule asubg_pOp_closed[of H], assumption+)

apply ((rule allI)+, rule impI, erule conjE)
 apply (simp add:mdl_def subsetD)
 apply (simp add:submodule_def[of R M H])
done

lemma (in Module) img_set_submodule:"⟦R module N; f ∈ mHom R M N⟧ ⟹
         submodule R N (f ` (carrier M))"
apply (simp add:submodule_def)
apply (rule conjI)
 apply (rule subsetI)
 apply (simp add:image_def)
 apply (erule bexE, simp, thin_tac "x = f xa")
  apply (simp add:mHom_mem)
apply (rule conjI)
 apply (frule Module.module_is_ag [of N])
 apply (rule aGroup.asubg_test, assumption+)
 apply (rule subsetI) apply (simp add:image_def)
 apply (erule bexE) apply (simp add:mHom_mem)
 apply (cut_tac ag_inc_zero,
        simp add:mHom_mem,  simp add:nonempty)
 apply ((rule ballI)+, simp add:image_def)
 apply ((erule bexE)+, simp)
 apply (simp add:mHom_inv[THEN sym],
        frule_tac x = xa in ag_mOp_closed,
        simp add:mHom_add[THEN sym, of N f],
        frule_tac x = "x" and y = "-a xa" in ag_pOp_closed, assumption+)
 apply blast

apply ((rule allI)+, rule impI, erule conjE)
 apply (simp add:image_def, erule bexE, simp)
 apply (simp add:mHom_lin[THEN sym, of N _ f])
 apply (frule_tac a = a and m = x in sc_mem, assumption) 
 apply blast 
done

lemma (in Module) mimg_module:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                                              R module (mimg R M N f)"
apply (simp add:mimg_def)
apply (rule Module.mdl_is_module[of N R "f ` (carrier M)"], assumption)
apply (simp add:img_set_submodule)
done
   
lemma (in Module) surjec_to_mimg:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                                       surjecM, (mimg R M N f) f"
apply (simp add:surjec_def)
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:mimg_def mdl_def)
 apply (rule conjI)
 apply (simp add:mHom_def aHom_def restrict_def extensional_def)
 apply ((rule ballI)+, simp add:mimg_def mdl_def, simp add:mHom_add)
apply (simp add:mimg_def mdl_def)
 apply (simp add:surj_to_def image_def)
done
 
definition
  tOp_mHom :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
    ('c, 'b, 'm2) Module_scheme] ⇒  ('a ⇒ 'c) ⇒ ('a ⇒ 'c) ⇒ ('a ⇒ 'c)" where
  "tOp_mHom R M N f g = (λx ∈ carrier M. (f x ±N (g x)))"

definition
  iOp_mHom :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
    ('c, 'b, 'm2) Module_scheme] ⇒  ('a ⇒ 'c) ⇒ ('a ⇒ 'c)" where
  "iOp_mHom R M N f = (λx ∈ carrier M. (-aN (f x)))" 

definition
  sprod_mHom ::"[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
    ('c, 'b, 'm2) Module_scheme] ⇒ 'b ⇒ ('a ⇒ 'c) ⇒ ('a ⇒ 'c)" where
  "sprod_mHom R M N a f = (λx ∈ carrier M. a ⋅sN (f x))"

definition
  HOM :: "[('b, 'more) Ring_scheme, ('a, 'b, 'more1) Module_scheme, 
    ('c, 'b, 'more2) Module_scheme] ⇒ ('a ⇒ 'c, 'b) Module"   
    ("(3HOM_ _/ _)" [90, 90, 91] 90) where
 "HOMR M N = ⦇carrier = mHom R M N, pop = tOp_mHom R M N, 
  mop = iOp_mHom R M N, zero = mzeromap M N,  sprod =sprod_mHom R M N ⦈"

lemma (in Module) zero_HOM:"R module N ⟹
         mzeromap M N = 𝟬HOMR M N"
apply (simp add:HOM_def)
done

lemma (in Module) tOp_mHom_closed:"⟦R module N; f ∈ mHom R M N; g ∈ mHom R M N⟧
      ⟹ tOp_mHom R M N f g ∈ mHom R M N"
apply (rule mHom_test, assumption+)
apply (rule conjI)
 apply (rule Pi_I)
 apply (simp add:tOp_mHom_def)
 apply (frule_tac f = f and m = x in mHom_mem [of N], assumption+,
        frule_tac f = g and m = x in mHom_mem [of N], assumption+,
        frule Module.module_is_ag [of N], 
        simp add:aGroup.ag_pOp_closed[of N])
apply (rule conjI)
 apply (simp add:tOp_mHom_def restrict_def extensional_def)
apply (rule conjI)
 apply (rule ballI)+
 apply (simp add:tOp_mHom_def)
 apply (simp add:ag_pOp_closed)
            
apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+,
       frule_tac f = f and m = n in mHom_mem [of N], assumption+,
       frule_tac f = g and m = m in mHom_mem [of N], assumption+,
       frule_tac f = g and m = n in mHom_mem [of N], assumption+,
       simp add:mHom_add,
       frule Module.module_is_ag [of N],
       subst aGroup.pOp_assocTr43[of "N"], assumption+,
       frule_tac x = "f n" and y = "g m" in aGroup.ag_pOp_commute [of "N"],
                                                              assumption+)
apply simp
apply (subst aGroup.pOp_assocTr43[of "N"], assumption+, simp) 

apply (rule ballI)+
apply (simp add:tOp_mHom_def) 
apply (frule_tac a = a and m = m in sc_mem, assumption, simp) 
apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+,
       frule_tac f = g and m = m in mHom_mem [of N], assumption+,
       frule_tac a = a and m = "f m" and n = "g m" in 
                                  Module.sc_r_distr[of N R], assumption+,
      simp)
apply (simp add:mHom_lin)
done

lemma (in Module) iOp_mHom_closed:"⟦R module N; f ∈ mHom R M N⟧
                                     ⟹ iOp_mHom R M N f ∈ mHom R M N"
apply (rule mHom_test, assumption+)
apply (rule conjI)
 apply (rule Pi_I)
 apply (simp add:iOp_mHom_def)
 apply (frule_tac f = f and m = x in mHom_mem [of N], assumption+)
 apply (frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_mOp_closed)
apply (rule conjI)
 apply (simp add:iOp_mHom_def restrict_def extensional_def)
apply (rule conjI) apply (rule ballI)+
 apply (simp add:iOp_mHom_def)
 apply (simp add:ag_pOp_closed)
 apply (simp add:mHom_add)
  apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+,
         frule_tac f = f and m = n in mHom_mem [of N], assumption+)
 apply (frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_p_inv)

apply (rule ballI)+
apply (simp add:iOp_mHom_def)
apply (simp add:sc_mem)
 apply (simp add:mHom_lin)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+)
 apply (simp add:Module.sc_minus_am[of N])
done

lemma (in Module) mHom_ex_zero:"R module N ⟹  mzeromap M N ∈ mHom R M N"
apply (simp add:mHom_def)
apply (rule conjI)
 apply (simp add:aHom_def,
        rule conjI,
        simp add:mzeromap_def, simp add:Module.module_inc_zero)

 apply (simp add:mzeromap_def extensional_def)

 apply ((rule ballI)+,
         simp add:ag_pOp_closed,
         frule Module.module_is_ag [of N],
         frule aGroup.ag_inc_zero [of "N"],
         simp add:aGroup.ag_l_zero)
apply ((rule ballI)+,
       simp add:mzeromap_def,
       simp add:sc_mem)
 apply (simp add:Module.sc_a_0)
done

lemma (in Module) mHom_eq:"⟦R module N; f ∈ mHom R M N; g ∈ mHom R M N; 
                            ∀m∈carrier M. f m = g m⟧ ⟹ f = g"  
apply (simp add:mHom_def aHom_def)
 apply (erule conjE)+
 apply (rule funcset_eq, assumption+)
done

lemma (in Module) mHom_l_zero:"⟦R module N; f ∈ mHom R M N⟧
              ⟹ tOp_mHom R M N (mzeromap M N) f = f"
apply (frule mHom_ex_zero [of N])
apply (frule tOp_mHom_closed [of N "mzeromap M N" f], assumption+)
apply (rule mHom_eq, assumption+)
 apply (rule ballI)
 apply (simp add:tOp_mHom_def, simp add:mzeromap_def)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+)
 apply (frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_l_zero[of N])
done

lemma  (in Module) mHom_l_inv:"⟦R module N; f ∈ mHom R M N⟧
       ⟹ tOp_mHom R M N (iOp_mHom R M N f) f = mzeromap M N"
apply (frule mHom_ex_zero [of N])
apply (frule_tac f = f in iOp_mHom_closed [of N], assumption,
       frule_tac f = "iOp_mHom R M N f" and g = f in tOp_mHom_closed [of N],
        assumption+,
       frule mHom_ex_zero [of N])
apply (rule mHom_eq, assumption+, rule ballI)
 apply (simp add:tOp_mHom_def iOp_mHom_def, simp add:mzeromap_def)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+)
 apply (frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_l_inv1)
done

lemma  (in Module) mHom_tOp_assoc:"⟦R module N; f ∈ mHom R M N; g ∈ mHom R M N;
        h ∈ mHom R M N⟧ ⟹ tOp_mHom R M N (tOp_mHom R M N f g) h =
          tOp_mHom R M N f (tOp_mHom R M N g h)"
apply (frule_tac f = f and g = g in tOp_mHom_closed [of N], assumption+,
       frule_tac f = "tOp_mHom R M N f g" and g = h in 
                      tOp_mHom_closed [of N], assumption+,
       frule_tac f = g and g = h in tOp_mHom_closed [of N], assumption+,
       frule_tac f = f and g = "tOp_mHom R M N g h" in 
                      tOp_mHom_closed [of N], assumption+) 
 apply (rule mHom_eq, assumption+, rule ballI,
        thin_tac "tOp_mHom R M N f g ∈ mHom R M N",
        thin_tac "tOp_mHom R M N (tOp_mHom R M N f g) h ∈ mHom R M N",
        thin_tac "tOp_mHom R M N g h ∈ mHom R M N",
        thin_tac "tOp_mHom R M N f (tOp_mHom R M N g h) ∈ mHom R M N")
 apply (simp add:tOp_mHom_def)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+,
        frule_tac f = g and m = m in mHom_mem [of N], assumption+,
        frule_tac f = h and m = m in mHom_mem [of N], assumption+)
apply (frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_pOp_assoc)
done

lemma (in Module) mHom_tOp_commute:"⟦R module N; f ∈ mHom R M N; 
        g ∈ mHom R M N⟧ ⟹ tOp_mHom R M N f g = tOp_mHom R M N g f"
apply (frule_tac f = f and g = g in tOp_mHom_closed [of N], assumption+,
       frule_tac f = g and g = f in tOp_mHom_closed [of N], assumption+)
apply (rule mHom_eq, assumption+)
 apply (rule ballI)
 apply (thin_tac "tOp_mHom R M N f g ∈ mHom R M N",
        thin_tac "tOp_mHom R M N g f ∈ mHom R M N")
 apply (simp add:tOp_mHom_def)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+,
        frule_tac f = g and m = m in mHom_mem [of N], assumption+,
        frule Module.module_is_ag [of N])
 apply (simp add:aGroup.ag_pOp_commute)
done

lemma  (in Module) HOM_is_ag:"R module N ⟹ aGroup (HOMR M N)"
apply (rule aGroup.intro)
 apply (simp add:HOM_def)
 apply (simp add:tOp_mHom_closed)

apply (simp add:HOM_def)
 apply (simp add:mHom_tOp_assoc)

apply (simp add:HOM_def)
 apply (simp add:mHom_tOp_commute)

apply (simp add:HOM_def)
 apply (simp add:iOp_mHom_closed)

apply (simp add:HOM_def,
       simp add:mHom_l_inv)

apply (simp add:HOM_def)
 apply (simp add:mHom_ex_zero)

apply (simp add:HOM_def,
       simp add:mHom_l_zero)
done

lemma (in Module) sprod_mHom_closed:"⟦R module N; a ∈ carrier R; 
       f ∈ mHom R M N⟧ ⟹ sprod_mHom R M N a f ∈ mHom R M N"
apply (rule mHom_test, assumption+)
apply (rule conjI)
 apply (simp add:Pi_def)
 apply (rule allI, rule impI, simp add:sprod_mHom_def,
        frule_tac f = f and m = x in mHom_mem [of N], assumption+,
        simp add:Module.sc_mem [of N R a])
apply (rule conjI)
 apply (simp add:sprod_mHom_def restrict_def extensional_def)
apply (rule conjI)
 apply (rule ballI)+
 apply (frule_tac x = m and y = n in ag_pOp_closed, assumption+)
 apply (simp add:sprod_mHom_def)
apply (subst mHom_add [of N f], assumption+)
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+, 
        frule_tac f = f and m = n in mHom_mem [of N], assumption+)
 apply (simp add:Module.sc_r_distr)

apply (rule ballI)+
 apply (simp add:sprod_mHom_def)
 apply (frule_tac a = aa and m = m in sc_mem, assumption+, simp)
 apply (simp add:mHom_lin) 
 apply (frule_tac f = f and m = m in mHom_mem [of N], assumption+)
apply (simp add:Module.sc_assoc[THEN sym, of N R]) 
apply (cut_tac sc_Ring, simp add:Ring.ring_tOp_commute)
done

lemma (in Module) HOM_is_module:"R module N ⟹ R module (HOMR M N)"
apply (rule Module.intro)
apply (simp add:HOM_is_ag)
apply (rule Module_axioms.intro)
 apply (simp add:sc_Ring)

 apply (simp add:HOM_def)
 apply (simp add:sprod_mHom_closed)

 apply (simp add:HOM_def)
 apply (cut_tac sc_Ring,
        frule Ring.ring_is_ag[of R],
        frule_tac x = a and y = b in aGroup.ag_pOp_closed[of R], assumption+,
        frule_tac a = "a ±R b" and f = m in sprod_mHom_closed[of N], 
        assumption+)
  apply(frule_tac a = a and f = m in sprod_mHom_closed[of N], assumption+,
        frule_tac a = b and f = m in sprod_mHom_closed[of N], assumption+,
        frule_tac f = "sprod_mHom R M N a m" and g = "sprod_mHom R M N b m" in
        tOp_mHom_closed[of N], assumption+)
  apply (rule mHom_eq[of N], assumption+, rule ballI,
         simp add:sprod_mHom_def tOp_mHom_def)
  apply (rename_tac a b f m)
  apply (frule_tac f = f and m = m in mHom_mem[of N], assumption+)
  apply (simp add:Module.sc_l_distr[of N])

apply (simp add:HOM_def)
 apply (rename_tac a f g,
        frule_tac f = f and g = g in tOp_mHom_closed[of N], assumption+,
        frule_tac a = a and f = "tOp_mHom R M N f g" in 
                                     sprod_mHom_closed[of N], assumption+,
        frule_tac a = a and f = f in sprod_mHom_closed[of N], assumption+,
        frule_tac a = a and f = g in sprod_mHom_closed[of N], assumption+,
        frule_tac f = "sprod_mHom R M N a f" and g = "sprod_mHom R M N a g" 
        in tOp_mHom_closed[of N], assumption+)   
 apply (rule mHom_eq[of N], assumption+, rule ballI,
        simp add:sprod_mHom_def tOp_mHom_def,
        frule_tac f = f and m = m in mHom_mem[of N], assumption+,
        frule_tac f = g and m = m in mHom_mem[of N], assumption+)
 apply (simp add:Module.sc_r_distr)

apply (simp add:HOM_def)
 apply (rename_tac a b f)
 apply (cut_tac sc_Ring,
        frule_tac x = a and y = b in Ring.ring_tOp_closed, assumption+,
        frule_tac a = "a ⋅rR b" and f = f in sprod_mHom_closed[of N], 
                                                            assumption+,
        frule_tac a = b and f = f in sprod_mHom_closed[of N], assumption+,
        frule_tac a = a and f = "sprod_mHom R M N b f" in 
                                     sprod_mHom_closed[of N], assumption+) 
 apply (rule mHom_eq[of N], assumption+, rule ballI,
        simp add:sprod_mHom_def,
        frule_tac f = f and m = m in mHom_mem[of N], assumption+,
        simp add:Module.sc_assoc)

apply (simp add:HOM_def)
 apply (cut_tac sc_Ring,
        frule Ring.ring_one,
        frule_tac a = "1rR" and f = m in sprod_mHom_closed[of N], assumption+)
 apply (rule mHom_eq, assumption+, rule ballI, rename_tac f m,
        simp add:sprod_mHom_def,
        frule_tac f = f and m = m in mHom_mem[of N], assumption+,
        simp add:Module.sprod_one)
done

section "Injective hom, surjective hom, bijective hom and inverse hom"

definition
  invmfun :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
              ('c, 'b, 'm2) Module_scheme, 'a ⇒ 'c] ⇒ 'c ⇒ 'a" where
  "invmfun R M N (f :: 'a ⇒ 'c) =
                    (λy∈(carrier N). SOME x. (x ∈ (carrier M) ∧ f x = y))"

definition
  misomorphic :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
              ('c, 'b, 'm2) Module_scheme] ⇒ bool" where
  "misomorphic R M N ⟷ (∃f. f ∈ mHom R M N ∧ bijecM,N f)"

definition
  mId :: "('a, 'b, 'm1) Module_scheme ⇒ 'a ⇒ 'a"   ("(mId_/ )" [89]88) where
  "mIdM = (λm∈carrier M. m)"

definition
  mcompose :: "[('a, 'r, 'm1) Module_scheme, 'b ⇒ 'c, 'a ⇒ 'b] ⇒ 'a ⇒ 'c" where
  "mcompose M g f = compose (carrier M) g f"

abbreviation
  MISOM  ("(3_ ≅_ _)" [82,82,83]82) where
  "M ≅R N == misomorphic R M N"

lemma (in Module) minjec_inj:"⟦R module N; injecM,N f⟧ ⟹
                            inj_on f (carrier M)" 
apply (simp add:inj_on_def, (rule ballI)+, rule impI)
 apply (simp add:injec_def, erule conjE)
 apply (frule Module.module_is_ag[of N])
 apply (cut_tac module_is_ag) 
 apply (frule_tac a = x in aHom_mem[of M N f], assumption+,
        frule_tac a = y in aHom_mem[of M N f], assumption+)
 apply (simp add:aGroup.ag_eq_diffzero[of N])
 apply (simp add:aHom_inv_inv[THEN sym, of M N f],
       frule_tac x = y in aGroup.ag_mOp_closed, assumption+,
       simp add:aHom_add[THEN sym, of M N f])
 apply (simp add:ker_def)
 apply (frule_tac x = x and y = "-a y" in ag_pOp_closed, assumption+)
 apply (subgoal_tac "(x ± -a y) ∈ {a ∈ carrier M. f a = 𝟬N}", simp)
 apply (simp add:ag_eq_diffzero)
 apply blast
done 

lemma (in Module) invmfun_l_inv:"⟦R module N; bijecM,N f; m ∈ carrier M⟧ ⟹
                            (invmfun R M N f) (f m) = m"
apply (simp add:bijec_def, erule conjE)
apply (frule minjec_inj [of N f], assumption+)
apply (simp add:surjec_def, erule conjE, simp add:aHom_def)
apply (frule conjunct1) 
apply (thin_tac "f ∈ carrier M → carrier N ∧
     f ∈ extensional (carrier M) ∧
     (∀a∈carrier M. ∀b∈carrier M. f (a ± b) = f a ±N f b)")
apply (frule invfun_l [of "f" "carrier M" "carrier N" "m"], assumption+)
 apply (simp add:surj_to_def) 
apply (simp add:invfun_def invmfun_def)
done
 
lemma (in Module) invmfun_mHom:"⟦R module N; bijecM,N f; f ∈ mHom R M N ⟧ ⟹
                 invmfun R M N f ∈ mHom R N M"
apply (frule minjec_inj [of N f])
 apply (simp add:bijec_def)
 apply (subgoal_tac "surjecM,N f") prefer 2 apply (simp add:bijec_def)
 apply (rule Module.mHom_test) apply assumption apply (rule Module_axioms)

apply (rule conjI) 
 apply (simp add:surjec_def, erule conjE)
 apply (simp add:aHom_def, frule conjunct1)
 apply (thin_tac "f ∈ carrier M → carrier N ∧
     f ∈ extensional (carrier M) ∧
     (∀a∈carrier M. ∀b∈carrier M. f (a ± b) = f a ±N f b)")
 apply (frule inv_func [of "f" "carrier M" "carrier N"], assumption+)
 apply (simp add:invmfun_def invfun_def)

apply (rule conjI)
 apply (simp add:invmfun_def restrict_def extensional_def)

apply (rule conjI)
 apply (rule ballI)+
 apply (simp add:surjec_def)
 apply (erule conjE, simp add:surj_to_def)
 apply (frule sym, thin_tac "f ` carrier M = carrier N", simp,
        thin_tac "carrier N = f ` carrier M")
 apply (simp add:image_def, (erule bexE)+, simp)
 apply (simp add:mHom_add[THEN sym])
 apply (frule_tac x = x and y = xa in ag_pOp_closed, assumption+)
 apply (simp add:invmfun_l_inv)

apply (rule ballI)+
 apply (simp add:surjec_def, erule conjE)
 apply (simp add:surj_to_def, frule sym, thin_tac "f ` carrier M = carrier N") 
 apply (simp add:image_def, (erule bexE)+, simp)
 apply (simp add:mHom_lin[THEN sym])
 apply (frule_tac a = a and m = x in sc_mem, assumption+)
 apply (simp add:invmfun_l_inv)
done

lemma (in Module) invmfun_r_inv:"⟦R module N; bijecM,N f; n ∈ carrier N⟧ ⟹
                           f ((invmfun R M N f) n) = n"
apply (frule minjec_inj[of N f])
 apply (simp add:bijec_def)
 apply (unfold bijec_def, frule conjunct2, fold bijec_def)
 apply (simp add:surjec_def, erule conjE, simp add:surj_to_def)
 apply (frule sym, thin_tac "f ` carrier M = carrier N", simp,
        thin_tac "carrier N = f ` carrier M")
 apply (simp add:image_def, erule bexE, simp)
 apply (simp add:invmfun_l_inv)
done

lemma (in Module) mHom_compos:"⟦R module L; R module N; f ∈ mHom R L M; 
       g ∈ mHom R M N ⟧ ⟹ compos L g f ∈ mHom R L N" 
apply (simp add:mHom_def [of "R" "L" "N"])
 apply (frule Module.module_is_ag [of L],
        frule Module.module_is_ag [of N])

apply (rule conjI) 
 apply (simp add:mHom_def, (erule conjE)+)
   apply (rule aHom_compos[of L M N f], assumption+)
   apply (cut_tac module_is_ag, assumption+)

apply (rule ballI)+
apply (simp add:compos_def compose_def)
 apply (simp add:Module.sc_mem)
 apply (subst Module.mHom_lin[of L R M _ f], assumption, rule Module_axioms, assumption+) (*apply (
        simp add:Module_def, rule conjI, assumption+) *)
 apply (subst Module.mHom_lin[of M R N _ g], rule Module_axioms, assumption) (*apply (
        simp add:Module_def, rule conjI)*)  (** ordering **)
 apply (rule Module.mHom_mem[of L R M f], assumption, rule Module_axioms, assumption+) 
 apply simp
done

lemma (in Module) mcompos_inj_inj:"⟦R module L; R module N; f ∈ mHom R L M; 
       g ∈ mHom R M N; injecL,M f; injecM,N g ⟧ ⟹ injecL,N (compos L g f)"
apply (frule Module.module_is_ag [of L],
       frule Module.module_is_ag [of N])
apply (simp add:injec_def [of "L" "N"])
apply (rule conjI)
 apply (simp add:injec_def, (erule conjE)+,
        rule_tac aHom_compos[of L M N], assumption+,
        rule module_is_ag)
 apply assumption+
 apply (simp add:compos_def compose_def)
 apply (rule equalityI)
 apply (rule subsetI, simp) 
 apply (simp add:injec_def [of _ _ "g"], erule conjE, simp add:ker_def)
 apply (subgoal_tac "f x ∈ {a. a ∈ carrier M ∧ g a = 𝟬N}")
 apply simp
 apply (simp add:injec_def [of _ _ "f"], erule conjE)
 apply (subgoal_tac "x ∈ kerL,M f", simp, thin_tac "kerL,M f = {𝟬L}")
 apply (simp add:ker_def)
 apply (thin_tac "{a ∈ carrier M. g a = 𝟬N} = {𝟬}")
 apply (simp, erule conjE, simp)
 apply (rule Module.mHom_mem[of L R M f], assumption, rule Module_axioms, assumption+) 

apply (rule subsetI, simp)
 apply (frule Module.module_inc_zero [of L R])
 apply (frule Module.mHom_0[of L R M f], rule Module_axioms, assumption+) 
 apply (simp add:ker_def)
 apply (subst mHom_0[of N], assumption+, simp)
done

lemma (in Module) mcompos_surj_surj:"⟦R module L; R module N; surjecL,M f;
        surjecM,N g; f ∈ mHom R L M; g ∈ mHom R M N ⟧ ⟹ 
                                        surjecL,N (compos L g f)"
apply (frule Module.module_is_ag [of L],
       frule Module.module_is_ag [of N],
       cut_tac module_is_ag)
apply (simp add:surjec_def [of "L" "N"])
apply (rule conjI)
 apply (simp add:mHom_def, (erule conjE)+)
 apply (rule aHom_compos[of L M N f g], assumption+)

apply (rule surj_to_test)
 apply (cut_tac Module.mHom_compos [of M R L N f g]) 
 apply (simp add:mHom_def aHom_def) 
 apply (rule Module_axioms, assumption+)

apply (rule ballI)
 apply (simp add: compos_def compose_def)
 apply (simp add:surjec_def [of _ _ "g"])
 apply (erule conjE) apply (simp add:surj_to_def)
 apply (frule sym, thin_tac "g ` carrier M = carrier N", simp add:image_def,
        thin_tac "carrier N = {y. ∃x∈carrier M. y = g x}",
        erule bexE, simp)
  apply (simp add:surjec_def [of _ _ "f"], erule conjE, simp add:surj_to_def,
         rotate_tac -1, frule sym, thin_tac "f ` carrier L = carrier M",
          simp add:image_def, erule bexE, simp)
 apply blast
done

lemma (in Module) mId_mHom:"mIdM ∈ mHom R M M"
apply (simp add:mHom_def)
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:mId_def)
apply (simp add:mId_def extensional_def)
apply (rule ballI)+
 apply (simp add:ag_pOp_closed)
apply (rule ballI)+
 apply (simp add:mId_def)
 apply (simp add:sc_mem)
done

lemma (in Module) mHom_mId_bijec:"⟦R module N; f ∈ mHom R M N; g ∈ mHom R N M;
      compose (carrier M) g f = mIdM; compose (carrier N) f g = mIdN⟧ ⟹
      bijecM,N f"
apply (simp add:bijec_def)
apply (rule conjI)
apply (simp add:injec_def)
 apply (rule conjI)
 apply (simp add:mHom_def)
 apply (simp add:ker_def)
 apply (rule equalityI)
 apply (rule subsetI, simp, erule conjE)
 apply (frule_tac x = "f x" and y = "𝟬N" and f = g in eq_elems_eq_val)
 apply (frule_tac f = "compose (carrier M) g f" and g = "mIdM" and x = x in
        eq_fun_eq_val, thin_tac "compose (carrier M) g f = mIdM", 
        simp add:compose_def)
 apply (cut_tac Module.mHom_0[of N R M g], simp add:mId_def, assumption,
   rule Module_axioms, assumption) 
apply (rule subsetI, simp,
       simp add:ag_inc_zero, simp add:mHom_0)

apply (simp add:surjec_def)
 apply (rule conjI, simp add:mHom_def)
 apply (rule surj_to_test)
 apply (simp add:mHom_def aHom_def)
 apply (rule ballI)
  apply (frule_tac f = "compose (carrier N) f g" and g = "mIdN" and x = b in
        eq_fun_eq_val, thin_tac "compose (carrier M) g f = mIdM",
        thin_tac "compose (carrier N) f g = mIdN", 
        simp add:compose_def)
 apply (simp add:mId_def)
 apply (frule_tac m = b in Module.mHom_mem [of N R M g], rule Module_axioms, assumption+)
 apply blast
done

definition
  sup_sharp :: "[('r, 'n) Ring_scheme, ('b, 'r, 'm1) Module_scheme, 
    ('c, 'r, 'm2) Module_scheme, ('a, 'r, 'm) Module_scheme, 'b ⇒ 'c] 
     ⇒ ('c ⇒ 'a) ⇒ ('b ⇒ 'a)" where
  "sup_sharp R M N L u = (λf∈mHom R N L. compos M f u)"

definition
  sub_sharp :: "[('r, 'n) Ring_scheme, ('a, 'r, 'm) Module_scheme, 
    ('b, 'r, 'm1) Module_scheme, ('c, 'r, 'm2) Module_scheme, 'b ⇒ 'c] 
     ⇒ ('a ⇒ 'b) ⇒ ('a ⇒ 'c)" where
  "sub_sharp R L M N u = (λf∈mHom R L M. compos L u f)"

       (*  L
          f| u
           M → N,  f → u o f   *)

lemma (in Module) sup_sharp_homTr:"⟦R module N; R module L; u ∈ mHom R M N; 
      f ∈ mHom R N L ⟧ ⟹ sup_sharp R M N L u f ∈ mHom R M L"
apply (simp add:sup_sharp_def)
apply (rule Module.mHom_compos, assumption, rule Module_axioms, assumption+) 
done

lemma (in Module) sup_sharp_hom:"⟦R module N; R module L; u ∈ mHom R M N⟧ ⟹ 
           sup_sharp R M N L u ∈ mHom R (HOMR N L) (HOMR M L)"
apply (simp add:mHom_def [of "R" "HOMR N L"])
apply (rule conjI) 
 apply (simp add:aHom_def) 
 apply (rule conjI)
 apply (simp add:HOM_def sup_sharp_homTr)

 apply (rule conjI)
 apply (simp add:sup_sharp_def extensional_def,
        rule allI, rule impI, simp add:HOM_def)

 apply (rule ballI)+
 apply (simp add:HOM_def)
 apply (frule_tac f = a and g = b in Module.tOp_mHom_closed, assumption+)
 apply (subgoal_tac "R module M")        
 apply (frule_tac f = a in Module.sup_sharp_homTr [of M R N L u], assumption+)
 apply (frule_tac f = b in Module.sup_sharp_homTr [of M R N L u], assumption+)
 apply (frule_tac f = "tOp_mHom R N L a b" in 
                            Module.sup_sharp_homTr[of M R N L u], assumption+) 
 apply (rule Module.mHom_eq, assumption+)
 apply (rule Module.tOp_mHom_closed, assumption+)

 apply (rule ballI)
 apply (simp add:sup_sharp_def tOp_mHom_def compose_def compos_def)
 apply (simp add:mHom_mem, rule Module_axioms)

apply (rule ballI)+
 apply (simp add:HOM_def)
 apply (frule_tac a = a and f = m in Module.sprod_mHom_closed [of N R L],
                                                                assumption+)
 apply (subgoal_tac "R module M",
        frule_tac f = "sprod_mHom R N L a m" in 
                 Module.sup_sharp_homTr [of M R N L u], assumption+)
 apply (frule_tac f = m in Module.sup_sharp_homTr [of M R N L u], assumption+)
 apply (frule_tac a = a and f = "sup_sharp R M N L u m" in 
           Module.sprod_mHom_closed [of M R L], assumption+)
 apply (rule mHom_eq, assumption+)
 apply (rule ballI)
 apply (simp add:sprod_mHom_def sup_sharp_def compose_def compos_def)
apply (simp add:Module.mHom_mem, rule Module_axioms)
done

lemma (in Module) sub_sharp_homTr:"⟦R module N; R module L; u ∈ mHom R M N; 
       f ∈ mHom R L M⟧ ⟹ sub_sharp R L M N u f ∈ mHom R L N"
apply (simp add:sub_sharp_def)
apply (simp add:mHom_compos)
done

lemma (in Module) sub_sharp_hom:"⟦R module N; R module L; u ∈ mHom R M N⟧ ⟹ 
          sub_sharp R L M N u ∈ mHom R (HOMR L M) (HOMR L N)"
apply (simp add:mHom_def [of _ "HOMR L M"])
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:HOM_def)
 apply (simp add:sub_sharp_homTr)

apply (rule conjI)
 apply (simp add:sub_sharp_def extensional_def)
 apply (simp add:HOM_def)

apply (rule ballI)+
 apply (simp add:HOM_def)
 apply (frule_tac f = a and g = b in Module.tOp_mHom_closed [of L R M],
   rule Module_axioms, assumption+)
 apply (subgoal_tac "R module M")
 apply (frule_tac f = "tOp_mHom R L M a b" in Module.sub_sharp_homTr 
                                 [of M R N L u], assumption+)
 apply (frule_tac f = b in Module.sub_sharp_homTr[of M R N L u],
                                                  assumption+,
        frule_tac f = a in Module.sub_sharp_homTr[of M R N L u], assumption+) 
 apply (frule_tac f = "sub_sharp R L M N u a" and 
  g = "sub_sharp R L M N u b" in Module.tOp_mHom_closed [of L R N],assumption+)
apply (rule Module.mHom_eq, assumption+)
 apply (rule ballI)
 apply (simp add:tOp_mHom_def sub_sharp_def mcompose_def compose_def,
        simp add:compos_def compose_def)
 apply (rule Module.mHom_add [of M R], assumption+)
 apply (simp add:Module.mHom_mem, simp add:Module.mHom_mem)
 apply (rule Module_axioms)

apply (rule ballI)+
 apply (simp add:HOM_def)
 apply (subgoal_tac "R module M")
 apply (frule_tac a = a and f = m in Module.sprod_mHom_closed [of L R M],
                                          assumption+)
 apply (frule_tac f = "sprod_mHom R L M a m" in Module.sub_sharp_homTr 
                                 [of M R N L u], assumption+) 
 apply (frule_tac f = m in Module.sub_sharp_homTr 
                                 [of M R N L u], assumption+) 
 apply (frule_tac a = a and f = "sub_sharp R L M N u m" in 
                       Module.sprod_mHom_closed [of L R N], assumption+)
apply (rule Module.mHom_eq, assumption+)
 apply (rule ballI)
 apply (simp add:sprod_mHom_def sub_sharp_def mcompose_def compose_def)
 apply (frule_tac  f = m and m = ma in Module.mHom_mem [of L R M], assumption+)
apply (simp add:compos_def compose_def) 
apply (simp add:mHom_lin)
apply (rule Module_axioms)
done   

lemma (in Module) mId_bijec:"bijecM,M (mIdM)" 
apply (simp add:bijec_def)
apply (cut_tac mId_mHom)
apply (rule conjI)
 apply (simp add:injec_def)
 apply (rule conjI) apply (simp add:mHom_def)
 apply (simp add:ker_def) apply (simp add:mId_def)
 apply (rule equalityI) apply (rule subsetI, simp) 
 apply (rule subsetI, simp, simp add:ag_inc_zero) 

apply (simp add:surjec_def)
 apply (rule conjI, simp add:mHom_def)
 apply (rule surj_to_test)
 apply (simp add:mHom_def aHom_def)
 apply (rule ballI)
 apply (simp add:mId_def)
done

lemma (in Module) invmfun_bijec:"⟦R module N; f ∈ mHom R M N; bijecM,N f⟧ ⟹
                  bijecN,M (invmfun R M N f)"
apply (frule invmfun_mHom [of N f], assumption+)
apply (simp add:bijec_def [of N M])
apply (rule conjI)
apply (simp add:injec_def)
 apply (simp add:mHom_def [of "R" "N" "M"]) apply (erule conjE)+
 apply (thin_tac "∀a∈carrier R.
        ∀m∈carrier N. invmfun R M N f (a ⋅sN m) = a ⋅s invmfun R M N f m")
 apply (rule equalityI) apply (rule subsetI) apply (simp add:ker_def CollectI)
 apply (erule conjE)
 apply (frule_tac x = "invmfun R M N f x" and y = "𝟬" and f = f in 
       eq_elems_eq_val,
       thin_tac "invmfun R M N f x = 𝟬")
 apply (simp add:invmfun_r_inv)
  apply (simp add:mHom_0)

apply (rule subsetI, simp)
 apply (simp add:ker_def)
 apply (simp add:Module.module_inc_zero)
 apply (cut_tac ag_inc_zero,
        frule invmfun_l_inv[of N f 𝟬], assumption+)
 apply (simp add:mHom_0)

apply (simp add:surjec_def,
       frule invmfun_mHom[of N f], assumption+)
 apply (rule conjI, simp add:mHom_def)
 apply (simp add:surj_to_def)
 apply (rule equalityI, rule subsetI, simp add:image_def, erule bexE,
        simp) thm Module.mHom_mem[of N R M "invmfun R M N f"]
 apply (rule Module.mHom_mem[of N R M "invmfun R M N f"], assumption,
   rule Module_axioms, assumption+) 
 apply (rule subsetI, simp add:image_def)
 apply (frule_tac m = x in invmfun_l_inv[of N f], assumption+)
 apply (frule_tac m = x in mHom_mem[of N f], assumption+)
 apply (frule sym, thin_tac "invmfun R M N f (f x) = x", blast)
done
  
lemma (in Module) misom_self:"M ≅R M"
apply (cut_tac mId_bijec)
apply (cut_tac mId_mHom)
apply (simp add:misomorphic_def)
apply blast
done

lemma (in Module) misom_sym:"⟦R module N; M ≅R N⟧ ⟹ N ≅R M"
apply (simp add:misomorphic_def [of "R" "M" "N"])
apply (erule exE, erule conjE)
apply (frule_tac f = f in invmfun_mHom [of N], assumption+)
apply (frule_tac f = f in invmfun_bijec [of N], assumption+)
apply (simp add:misomorphic_def)
apply blast
done

lemma (in Module) misom_trans:"⟦R module L; R module N; L ≅R M; M ≅R N⟧ ⟹ 
                               L ≅R N"
apply (simp add:misomorphic_def)
 apply ((erule exE)+, (erule conjE)+)
 apply (subgoal_tac  "bijecL,N (compos L fa f)")
 apply (subgoal_tac "(compos L fa f) ∈ mHom R L N")
 apply blast
 apply (rule Module.mHom_compos[of M R L N], rule Module_axioms, assumption+) 

apply (simp add:bijec_def) apply (erule conjE)+
apply (simp add:mcompos_inj_inj)                                
apply (simp add:mcompos_surj_surj)
done

definition
  mr_coset :: "['a, ('a, 'b, 'more) Module_scheme, 'a set] ⇒ 'a set" where
  "mr_coset a M H = a ⊎M H"

definition
  set_mr_cos :: "[('a, 'b, 'more) Module_scheme, 'a set] ⇒ 'a set set" where
  "set_mr_cos M H = {X. ∃a∈carrier M. X = a ⊎M H}"

definition
  mr_cos_sprod :: "[('a, 'b, 'more) Module_scheme, 'a set] ⇒ 
                                              'b ⇒ 'a set ⇒ 'a set" where
  "mr_cos_sprod M H a X = {z. ∃x∈X. ∃h∈H. z = h ±M (a ⋅sM x)}"

definition
  mr_cospOp :: "[('a, 'b, 'more) Module_scheme, 'a set] ⇒ 
                                               'a set ⇒ 'a set ⇒ 'a set" where
  "mr_cospOp M H = (λX. λY. c_top (b_ag M) H X Y)"  

definition
  mr_cosmOp :: "[('a, 'b, 'more) Module_scheme, 'a set] ⇒ 
                                                  'a set ⇒ 'a set" where
  "mr_cosmOp M H = (λX. c_iop (b_ag M) H X)"

definition
  qmodule :: "[('a, 'r, 'more) Module_scheme, 'a set] ⇒
                 ('a set, 'r) Module" where
  "qmodule M H = ⦇ carrier = set_mr_cos M H, pop = mr_cospOp M H, 
    mop = mr_cosmOp M H, zero = H, sprod = mr_cos_sprod M H⦈"

definition
  sub_mr_set_cos :: "[('a, 'r, 'more) Module_scheme, 'a set, 'a set] ⇒
                            'a set set" where
 "sub_mr_set_cos M H N = {X. ∃n∈N. X = n ⊎M H}" 
 (* N/H, where N is a submodule *)

abbreviation
  QMODULE  (infixl "'/'m" 200) where
  "M /m H == qmodule M H"

abbreviation
  SUBMRSET  ("(3_/ s'/'_/ _)" [82,82,83]82) where
  "N s/M H == sub_mr_set_cos M H N"

lemma (in Module) qmodule_carr:"submodule R M H ⟹
            carrier (qmodule M H) = set_mr_cos M H"
apply (simp add:qmodule_def)
done

lemma (in Module) set_mr_cos_mem:"⟦submodule R M H; m ∈ carrier M⟧ ⟹
                        m ⊎M H ∈ set_mr_cos M H"
apply (simp add:set_mr_cos_def) 
apply blast
done

lemma (in Module) mem_set_mr_cos:"⟦submodule R M N; x ∈ set_mr_cos M N⟧ ⟹
                          ∃m ∈ carrier M. x = m  ⊎M N"
by (simp add:set_mr_cos_def)

lemma (in Module) m_in_mr_coset:"⟦submodule R M H; m ∈ carrier M⟧ ⟹
                                   m ∈ m ⊎M H"
apply (cut_tac module_is_ag)
apply (frule aGroup.b_ag_group)
apply (simp add:ar_coset_def)
apply (simp add:aGroup.ag_carrier_carrier [THEN sym])
apply (simp add:submodule_def) apply (erule conjE)+ 
apply (simp add:asubGroup_def)
apply (rule Group.a_in_rcs [of "b_ag M" "H" "m"], assumption+)
done

lemma (in Module) mr_cos_h_stable:"⟦submodule R M H; h ∈ H⟧ ⟹
                                                       H = h ⊎M H"
apply (cut_tac module_is_ag)
apply (frule aGroup.b_ag_group [of "M"])
apply (simp add:ar_coset_def) 
apply (rule Group.rcs_Unit2[THEN sym], assumption+,
        simp add:submodule_def, (erule conjE)+, 
        simp add:asubGroup_def) 
apply assumption
done

lemma (in Module) mr_cos_h_stable1:"⟦submodule R M H; m ∈ carrier M; h ∈ H⟧
             ⟹ (m ± h) ⊎M H = m ⊎M H"
apply (cut_tac module_is_ag)
apply (subst aGroup.ag_pOp_commute, assumption+)
 apply (simp add:submodule_def, (erule conjE)+, simp add:subsetD)
apply (frule aGroup.b_ag_group [of "M"])
apply (simp add:ar_coset_def)
apply (simp add:aGroup.agop_gop [THEN sym])
apply (simp add:aGroup.ag_carrier_carrier [THEN sym])
apply (simp add:submodule_def, (erule conjE)+, simp add:asubGroup_def)
apply (rule Group.rcs_fixed1 [THEN sym, of "b_ag M" "H" "m" "h"], assumption+)
done

lemma (in Module) x_in_mr_coset:"⟦submodule R M H; m ∈ carrier M; x ∈ m ⊎M H⟧
                 ⟹ ∃h∈H. m ± h = x"
apply (cut_tac module_is_ag)
 apply (frule aGroup.b_ag_group [of "M"])
 apply (simp add:submodule_def, (erule conjE)+,
        simp add:asubGroup_def)
 apply (simp add:aGroup.ag_carrier_carrier [THEN sym])
 apply (simp add:aGroup.agop_gop [THEN sym])
 apply (simp add:ar_coset_def)
 apply (frule Group.rcs_tool2[of "b_ag M" H m x], assumption+,
        erule bexE)
 apply (frule sym, thin_tac "h ⋅b_ag M m = x", simp)
 apply (simp add:aGroup.agop_gop)
 apply (simp add:aGroup.ag_carrier_carrier)
 apply (frule_tac c = h in subsetD[of H "carrier M"], assumption+)
 apply (subst ag_pOp_commute[of _ m], assumption+)
 apply blast
done

lemma (in Module) mr_cos_sprodTr:"⟦submodule R M H; a ∈ carrier R; 
       m ∈ carrier M⟧ ⟹ mr_cos_sprod M H a (m ⊎M H) = (a ⋅s m) ⊎M H"
apply (cut_tac module_is_ag,
       frule aGroup.b_ag_group,
       frule sc_mem[of a m], assumption)
 apply (simp add:ar_coset_def,
        simp add:mr_cos_sprod_def)
 apply (simp add:submodule_def, (erule conjE)+)
 apply (simp add:aGroup.ag_carrier_carrier [THEN sym],
        simp add:aGroup.agop_gop [THEN sym])
 apply (simp add:asubGroup_def)
apply (rule equalityI)
 apply (rule subsetI, simp) 
 apply (erule bexE)+
 apply (frule_tac x = xa in Group.rcs_tool2[of "b_ag M" H m], assumption+)
 apply (erule bexE, rotate_tac -1, frule sym, thin_tac "ha ⋅b_ag M m = xa",
        simp)
 apply (simp add:aGroup.agop_gop, simp add:aGroup.ag_carrier_carrier)
 apply (frule_tac c = ha in subsetD[of H "carrier M"], assumption+,
        simp add:sc_r_distr,
        drule_tac x = a in spec,
        drule_tac a = ha in forall_spec, simp,
        frule_tac c = "a ⋅s ha" in subsetD[of H "carrier M"], assumption+,
        frule_tac c = h in subsetD[of H "carrier M"], assumption+,
        subst ag_pOp_assoc[THEN sym], assumption+)
 apply (simp add:aGroup.agop_gop[THEN sym], 
        simp add:aGroup.ag_carrier_carrier[THEN sym]) 
 apply (frule_tac x = h and y = "a ⋅s ha" in 
                  Group.sg_mult_closed[of "b_ag M" H], assumption+)
 apply (frule_tac a = "a ⋅s m" and h = "h ⋅b_ag M (a ⋅s ha)" in 
                  Group.rcs_fixed1[of "b_ag M" H], assumption+)
 apply simp
 apply (rule Group.a_in_rcs [of "b_ag M" "H"], assumption+)
 apply (simp add:aGroup.agop_gop, simp add:aGroup.ag_carrier_carrier)  
 apply (rule ag_pOp_closed, simp add:subsetD, assumption)

apply (rule subsetI, simp,
       frule_tac x = x in Group.rcs_tool2[of "b_ag M" H "a ⋅s m"], assumption+,
       erule bexE,
       rotate_tac -1, frule sym, thin_tac "h ⋅b_ag M (a ⋅s m) = x",
       frule Group.a_in_rcs[of "b_ag M" H m], assumption+)
 apply blast
done

lemma (in Module) mr_cos_sprod_mem:"⟦submodule R M H; a ∈ carrier R; 
       X ∈ set_mr_cos M H⟧ ⟹ mr_cos_sprod M H a X ∈ set_mr_cos M H"
apply (simp add:set_mr_cos_def)
 apply (erule bexE, rename_tac m, simp) 
 apply (subst mr_cos_sprodTr, assumption+)
 apply (frule_tac m = m in sc_mem [of a], assumption)
apply blast
done  

lemma (in Module) mr_cos_sprod_assoc:"⟦submodule R M H; a ∈ carrier R;
 b ∈ carrier R; X ∈ set_mr_cos M H⟧ ⟹ mr_cos_sprod  M H (a ⋅rR b) X = 
                           mr_cos_sprod M H a (mr_cos_sprod M H b X)"
apply (simp add:set_mr_cos_def, erule bexE, simp)
 apply (frule_tac m = aa in sc_mem [of b], assumption)
 apply (cut_tac sc_Ring,
        frule Ring.ring_tOp_closed [of "R" "a" "b"], assumption+)
 apply (subst mr_cos_sprodTr, assumption+)+
 apply (simp add: sc_assoc)
done

lemma (in Module) mr_cos_sprod_one:"⟦submodule R M H; X ∈ set_mr_cos M H⟧ ⟹
                   mr_cos_sprod M H (1rR) X = X"
apply (simp add:set_mr_cos_def, erule bexE, simp,
       thin_tac "X = a ⊎M H")
 apply (cut_tac sc_Ring,
        frule Ring.ring_one[of "R"])
 apply (subst mr_cos_sprodTr, assumption+) 
 apply (simp add:sprod_one)
done

lemma (in Module) mr_cospOpTr:"⟦submodule R M H; m ∈ carrier M; n ∈ carrier M⟧
        ⟹ mr_cospOp M H (m ⊎M H) (n ⊎M H) = (m ± n) ⊎M H" 
apply (cut_tac module_is_ag, frule aGroup.b_ag_group) 
apply (simp add:mr_cospOp_def mr_coset_def agop_gop [THEN sym])
apply (simp add:ag_carrier_carrier [THEN sym])

apply (simp add:submodule_def, (erule conjE)+,
       frule aGroup.asubg_nsubg, assumption+, simp add:ar_coset_def)
apply (simp add:Group.c_top_welldef[THEN sym, of "b_ag M" H m n])
done

lemma(in Module) mr_cos_sprod_distrib1:"⟦submodule R M H; a ∈ carrier R; 
                b ∈ carrier R;  X ∈ set_mr_cos M H⟧ ⟹ 
                mr_cos_sprod M H (a ±R b) X =  
                 mr_cospOp M H (mr_cos_sprod M H a X) (mr_cos_sprod M H b X)"
apply (simp add:set_mr_cos_def, erule bexE, rename_tac m)
 apply simp
 apply (cut_tac sc_Ring,
        frule Ring.ring_is_ag[of R])
 apply (frule aGroup.ag_pOp_closed [of R a b], assumption+)
apply (subst mr_cos_sprodTr [of H], assumption+)+
apply (subst mr_cospOpTr, assumption+)
 apply (simp add:sc_mem, simp add:sc_mem)
 apply (simp add:sc_l_distr)
done

lemma (in Module) mr_cos_sprod_distrib2:"⟦submodule R M H; 
 a ∈ carrier R; X ∈ set_mr_cos M H; Y ∈ set_mr_cos M H⟧ ⟹ 
 mr_cos_sprod M H a (mr_cospOp M H X Y) =  
           mr_cospOp M H (mr_cos_sprod M H a X) (mr_cos_sprod M H a Y)"
apply (simp add:set_mr_cos_def, (erule bexE)+, rename_tac m n, simp,
       thin_tac "X = m ⊎M H", thin_tac "Y = n ⊎M H")
apply (subst mr_cos_sprodTr [of H], assumption+)+
 apply (subst mr_cospOpTr, assumption+)
 apply (subst mr_cospOpTr, assumption+)
 apply (simp add:sc_mem)+
apply (subst mr_cos_sprodTr [of H], assumption+)
 apply (rule ag_pOp_closed, assumption+)
apply (simp add:sc_r_distr)
done

lemma (in Module) mr_cosmOpTr:"⟦submodule R M H; m ∈ carrier M⟧ ⟹ 
                mr_cosmOp M H (m ⊎M H) = (-a m) ⊎M H"
apply (simp add:ar_coset_def) 
apply (cut_tac module_is_ag)
apply (frule aGroup.b_ag_group)
apply (simp add:ag_carrier_carrier [THEN sym])
apply (simp add:agiop_giop [THEN sym])
apply (simp add:mr_cosmOp_def)
 apply (simp add:submodule_def, (erule conjE)+,
        frule aGroup.asubg_nsubg[of M H], assumption)
 apply (simp add:Group.c_iop_welldef[of "b_ag M" H m])
done

lemma (in Module) mr_cos_oneTr:"submodule R M H ⟹ H =  𝟬 ⊎M H"
apply (cut_tac module_is_ag,
       cut_tac ag_inc_zero)
 apply (simp add:ar_coset_def)
 apply (frule aGroup.b_ag_group)
 apply (simp add:ag_carrier_carrier [THEN sym])
 apply (subst aGroup.agunit_gone[THEN sym, of M], assumption)
 apply (subst Group.rcs_Unit1, assumption)
 apply (simp add:submodule_def, (erule conjE)+, simp add:asubGroup_def)
 apply simp
done

lemma (in Module) mr_cos_oneTr1:"⟦submodule R M H; m ∈ carrier M⟧ ⟹
                            mr_cospOp M H H (m ⊎M H) = m ⊎M H"
apply (subgoal_tac "mr_cospOp M H (𝟬 ⊎M H) (m ⊎M H) = m ⊎M H")
apply (simp add:mr_cos_oneTr [THEN sym, of H])
apply (subst mr_cospOpTr, assumption+)
 apply (simp add:ag_inc_zero)
 apply assumption
 apply (simp add:ag_l_zero)
done

lemma (in Module) qmodule_is_ag:"submodule R M H ⟹ aGroup (M /m H)"
apply (cut_tac sc_Ring)
apply (rule aGroup.intro) 
 apply (simp add:qmodule_def)
 apply (rule Pi_I)+
 apply (rename_tac X Y)
 apply (simp add:set_mr_cos_def, (erule bexE)+, rename_tac n m, simp)
 apply (subst mr_cospOpTr, assumption+,
        frule_tac x = n and y = m in ag_pOp_closed, assumption+, blast)

 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def, (erule bexE)+, rename_tac a b c m n n')
 apply (simp add:mr_cospOpTr,
        frule_tac x = m and y = n in ag_pOp_closed, assumption+,
        frule_tac x = n and y = n' in ag_pOp_closed, assumption+,
       simp add:mr_cospOpTr, simp add:ag_pOp_assoc)

 apply (simp add:qmodule_def) 
  apply (simp add:set_mr_cos_def, (erule bexE)+, rename_tac a b m n, simp)
  apply (simp add:mr_cospOpTr,
         simp add:ag_pOp_commute)

 apply (simp add:qmodule_def,
        rule Pi_I,
        simp add:set_mr_cos_def, erule bexE, simp)
 apply (subst mr_cosmOpTr, assumption+,
         frule_tac x = a in ag_mOp_closed, blast)

 apply (simp add:qmodule_def,
        simp add:set_mr_cos_def, erule bexE, simp,
        simp add:mr_cosmOpTr,
        frule_tac x = aa in ag_mOp_closed)  
 apply (simp add:mr_cospOpTr,
        frule_tac x = "-a aa" and y = aa in ag_pOp_closed, assumption+,
        simp add:ag_l_inv1, simp add:mr_cos_oneTr[THEN sym])
 apply (simp add:qmodule_def,
        simp add:set_mr_cos_def,
        cut_tac mr_cos_oneTr[of H],
        cut_tac ag_inc_zero, blast, assumption)

 apply (simp add:qmodule_def)
  apply (simp add:set_mr_cos_def, erule bexE, simp)
 apply (subgoal_tac "mr_cospOp M H (𝟬 ⊎M H) (aa ⊎M H) = aa ⊎M H")
  apply (simp add:mr_cos_oneTr[THEN sym, of H])
 apply (subst mr_cospOpTr, assumption+,
        simp add:ag_inc_zero, assumption, simp add:ag_l_zero)
done

lemma (in Module) qmodule_module:"submodule R M H ⟹ R module (M /m H)"
apply (rule Module.intro)
apply (simp add:qmodule_is_ag)
apply (rule Module_axioms.intro)
 apply (cut_tac sc_Ring, simp)

apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprod_mem)

apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprod_distrib1[of H])

apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprod_distrib2[of H])

apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprod_assoc)

apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprod_one)
done

definition
  indmhom :: "[('b, 'm) Ring_scheme, ('a, 'b, 'm1) Module_scheme, 
    ('c, 'b, 'm2) Module_scheme, 'a ⇒ 'c] ⇒  'a set ⇒ 'c" where
  "indmhom R M N f = (λX∈ (set_mr_cos M (kerM,N f)). f ( SOME x. x ∈ X))"

abbreviation
  INDMHOM  ("(4__ _, _)" [92,92,92,93]92) where
  "fR M,N == indmhom R M N f"


lemma (in Module) indmhom_someTr:"⟦R module N; f ∈ mHom R M N; 
      X ∈ set_mr_cos M (kerM,N f)⟧ ⟹ f (SOME xa. xa ∈ X) ∈ f `(carrier M)"
apply (simp add:set_mr_cos_def)
 apply (erule bexE, simp) 
apply (frule mker_submodule [of N f], assumption+)
apply (simp add:submodule_def) apply (erule conjE)+
apply (simp add:asubGroup_def)
 apply (thin_tac "∀a m. a ∈ carrier R ∧ m ∈ kerM,N f ⟶ a ⋅s m ∈ kerM,N f")
 apply (cut_tac module_is_ag)
 apply (frule aGroup.b_ag_group)
apply (rule someI2_ex)
 apply (simp add:ar_coset_def)
 apply (frule_tac a = a in Group.a_in_rcs[of "b_ag M" "kerM,N f"], 
        assumption+, simp add:ag_carrier_carrier [THEN sym], blast)
apply (simp add:ar_coset_def)
 apply (frule_tac a = a and x = x in 
                  Group.rcs_subset_elem[of "b_ag M" "kerM,N f"], assumption+)
 apply (simp add:ag_carrier_carrier, assumption+)

apply (simp add:image_def,
       simp add:ag_carrier_carrier, blast)
done

lemma (in Module) indmhom_someTr1:"⟦R module N; f ∈ mHom R M N; m ∈ carrier M⟧
        ⟹  f (SOME xa. xa ∈ (ar_coset m M (kerM,N f))) = f m"
apply (rule someI2_ex)
 apply (frule mker_submodule[of N f], assumption)
 apply (frule_tac m_in_mr_coset[of "kerM,N f" m], assumption+,
        blast)

 apply (frule mker_submodule [of N f], assumption+) 
 apply (frule_tac x = x in x_in_mr_coset [of  "kerM,N f" "m"], 
                                         assumption+, erule bexE,
        frule sym , thin_tac "m ± h = x", simp)
 apply (simp add:ker_def, erule conjE)
 apply (subst mHom_add[of N f ], assumption+, simp)
apply (frule Module.module_is_ag [of N R])
 apply (frule mHom_mem [of "N" "f" "m"], assumption+)
apply (simp add:aGroup.ag_r_zero)
done

lemma (in Module) indmhom_someTr2:"⟦R module N; f ∈ mHom R M N; 
       submodule R M H; m ∈ carrier M; H ⊆ kerM,N f⟧ ⟹ 
                       f (SOME xa. xa ∈ m ⊎M H) = f m"
apply (rule someI2_ex)
  apply (frule_tac m_in_mr_coset[of "H" m], assumption+, blast) 
   apply (frule_tac x = x in x_in_mr_coset [of  H m], 
                                         assumption+, erule bexE,
        frule sym , thin_tac "m ± h = x", simp)
 apply (frule_tac c = h in subsetD[of H "kerM,N f"], assumption+)
 apply (frule mker_submodule [of N f], assumption+, 
         simp add:submodule_def[of R M "kerM,N f"], (erule conjE)+,
        frule_tac c = h in subsetD[of "kerM,N f" "carrier M"], assumption+)
 apply (simp add:ker_def mHom_add,
        frule_tac m = m in mHom_mem[of "N" "f"], assumption+)
 apply (frule Module.module_is_ag[of N R])
 apply (simp add:aGroup.ag_r_zero)
done

lemma (in Module) indmhomTr1:"⟦R module N; f ∈ mHom R M N; m ∈ carrier M⟧ ⟹
               (fR M,N) (m ⊎M (kerM,N f)) = f m" 
apply (simp add:indmhom_def)
apply (subgoal_tac "m ⊎M kerM,N f ∈ set_mr_cos M (kerM,N f)", simp)
 apply (rule indmhom_someTr1, assumption+)
 apply (rule set_mr_cos_mem)
apply (rule mker_submodule, assumption+)
done

lemma (in Module) indmhomTr2:"⟦R module N; f ∈ mHom R M N⟧ 
      ⟹ (fR M,N) ∈ set_mr_cos M (kerM,N f) → carrier N" 
 apply (rule Pi_I)
 apply (simp add:set_mr_cos_def)
 apply (erule bexE)
 apply (frule_tac m = a in indmhomTr1 [of N f], assumption+)
 apply (simp add:mHom_mem)
done

lemma (in Module) indmhom:"⟦R module N; f ∈ mHom R M N⟧ 
                           ⟹ (fR M,N) ∈ mHom R (M /m (kerM,N f)) N"
apply (simp add:mHom_def [of R "M /m (kerM,N f)" N])
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:qmodule_def)
 apply (simp add:indmhomTr2)

apply (rule conjI)
 apply (simp add:qmodule_def indmhom_def extensional_def) 

apply (rule ballI)+
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def, (erule bexE)+, simp, rename_tac  m n)
 apply (frule mker_submodule [of N f], assumption+,
        simp add:mr_cospOpTr,
        frule_tac x = m and y = n in ag_pOp_closed, assumption+)
 apply (simp add:indmhomTr1, simp add:mHom_add)

 apply (rule ballI)+ 
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def, (erule bexE)+, simp)
 apply (frule mker_submodule [of N f], assumption+,
        subst mr_cos_sprodTr [of "kerM,N f"], assumption+,
        frule_tac a = a and m = aa in sc_mem, assumption)
 apply (simp add:indmhomTr1)
 apply (simp add:mHom_lin)
done

lemma (in Module) indmhom_injec:"⟦R module N; f ∈ mHom R M N⟧ ⟹
       injec(M /m (kerM,N f)),N (fR M,N)"
apply (simp add:injec_def)
apply (frule indmhom [of N f], assumption+)
apply (rule conjI)
apply (simp add:mHom_def)
apply (simp add:ker_def [of  _ _ "fR M, N"])
apply (simp add:qmodule_def) apply (fold qmodule_def)
apply (rule equalityI)
 apply (rule subsetI) apply (simp add:CollectI) apply (erule conjE)
 apply (simp add:set_mr_cos_def, erule bexE, simp)
 apply (simp add:indmhomTr1)
apply (frule mker_submodule [of N f], assumption+)
 apply (rule_tac h1 = a in mr_cos_h_stable [THEN sym, of "kerM,N f"], 
         assumption+)
 apply (simp add:ker_def)

apply (rule subsetI) apply (simp add:CollectI)
 apply (rule conjI)
 apply (simp add:set_mr_cos_def)
 apply (frule mker_submodule [of N f], assumption+)
 apply (frule mr_cos_oneTr [of "kerM,N f"])
 apply (cut_tac  ag_inc_zero)
 apply blast
 apply (frule mker_submodule [of N f], assumption+) 
apply (subst mr_cos_oneTr [of "kerM,N f"], assumption)
 apply (cut_tac  ag_inc_zero)        
 apply (subst indmhomTr1, assumption+)
 apply (simp add:mHom_0)
done

lemma (in Module) indmhom_surjec1:"⟦R module N; surjecM,N f;
 f ∈ mHom R M N⟧ ⟹ surjec(M /m (kerM,N f)),N (fR M,N)"
apply (simp add:surjec_def)
 apply (frule indmhom [of N f], assumption+)
 apply (rule conjI)
 apply (simp add:mHom_def)
apply (rule surj_to_test)
 apply (simp add:mHom_def aHom_def)
apply (rule ballI)
 apply (erule conjE) 
 apply (simp add:surj_to_def, frule sym , thin_tac "f ` carrier M = carrier N",
        simp,
        thin_tac "carrier N = f ` carrier M")
 apply (simp add:image_def, erule bexE, simp)
 apply (frule_tac m = x in indmhomTr1 [of N f], assumption+)
 apply (frule mker_submodule [of N f], assumption+)
 apply (simp add:qmodule_carr)
 apply (frule_tac m = x in set_mr_cos_mem [of "kerM,N f"], assumption+)
apply blast
done

lemma (in Module) module_homTr:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                           f ∈ mHom R M (mimgR M,N f)"
apply (subst mHom_def, simp add:CollectI)
 apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:mimg_def mdl_def)
apply (rule conjI)
 apply (simp add:mHom_def aHom_def extensional_def)
apply (rule ballI)+
 apply (simp add:mimg_def mdl_def)
 apply (simp add:mHom_add)
apply (rule ballI)+
 apply (simp add:mimg_def mdl_def)
 apply (simp add:mHom_lin)
done

lemma (in Module) ker_to_mimg:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                kerM,mimgR M,N f f = kerM,N f"
apply (rule equalityI)
 apply (rule subsetI)
 apply (simp add:ker_def mimg_def mdl_def)
 apply (rule subsetI)
 apply (simp add:ker_def mimg_def mdl_def) 
done

lemma (in Module) module_homTr1:"⟦R module N; f ∈ mHom R M N⟧ ⟹
   (mimgR (M /m (kerM,N f)),N (fR M,N)) = mimgR M,N f"    apply (simp add:mimg_def)
apply (subgoal_tac "fR M, N ` carrier (M /m (kerM,N f))  = f ` carrier M ",
       simp)
apply (simp add:qmodule_def)
apply (rule equalityI)
 apply (rule subsetI)
 apply (simp add:image_def set_mr_cos_def)
 apply (erule exE, erule conjE, erule bexE, simp)
 apply (simp add:indmhomTr1, blast)
apply (rule subsetI,
       simp add:image_def set_mr_cos_def, erule bexE, simp)
 apply (frule_tac m1 = xa in indmhomTr1 [THEN sym, of N f], 
                                                     assumption+)
 apply blast
done

lemma (in Module) module_Homth_1:"⟦R module N; f ∈ mHom R M N⟧ ⟹
                     M /m (kerM,N f) ≅R mimgR M,N f"
apply (frule surjec_to_mimg[of N f], assumption,
       frule module_homTr[of N f], assumption,
       frule mimg_module[of N f], assumption,
       frule indmhom_surjec1[of "mimgR M,N f" f], assumption+,
       frule indmhom_injec[of "mimgR M,N f" f], assumption+,
       frule indmhom[of "mimgR M,N f" f], assumption+)
apply (simp add:misomorphic_def,
       simp add:bijec_def)
apply (simp add:ker_to_mimg)
apply blast
done

definition
  mpj :: "[('a, 'r, 'm) Module_scheme, 'a set] ⇒  ('a => 'a set)" where
  "mpj M H = (λx∈carrier M. x ⊎M H)" 

lemma (in Module) elem_mpj:"⟦m ∈ carrier M; submodule R M H⟧ ⟹
                                                 mpj M H m = m ⊎M H"
by (simp add:mpj_def)

lemma (in Module) mpj_mHom:"submodule R M H ⟹ mpj M H ∈ mHom R M (M /m H)"
apply (simp add:mHom_def)
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (rule conjI)
 apply (simp add:mpj_def qmodule_carr set_mr_cos_mem)
apply (rule conjI)
 apply (simp add:mpj_def extensional_def)
apply (rule ballI)+
 apply (simp add:qmodule_def)
 apply (simp add:mpj_def, simp add:ag_pOp_closed)
 apply (simp add:mr_cospOpTr)
apply (rule ballI)+
 apply (simp add:mpj_def sc_mem)
 apply (simp add:qmodule_def)
 apply (simp add:mr_cos_sprodTr)
done
 
lemma (in Module) mpj_mem:"⟦submodule R M H; m ∈ carrier M⟧ ⟹
                                mpj M H m ∈ carrier (M /m H)"
apply (frule mpj_mHom[of H])
apply (rule mHom_mem [of "M /m H" "mpj M H" "m"])
 apply (simp add:qmodule_module) apply assumption+
done

lemma (in Module) mpj_surjec:"submodule R M H ⟹
                             surjecM,(M /m H) (mpj M H)" 
apply (simp add:surjec_def)
apply (frule mpj_mHom [of H])
apply (rule conjI, simp add:mHom_def)
apply (rule surj_to_test,
       simp add:mHom_def aHom_def)
apply (rule ballI)
 apply (thin_tac "mpj M H ∈ mHom R M (M /m H)")

 apply (simp add:qmodule_def)
apply (simp add:set_mr_cos_def, erule bexE, simp)
 apply (frule_tac m = a in elem_mpj[of _ H], assumption, blast)
done

lemma (in Module) mpj_0:"⟦submodule R M H; h ∈ H⟧ ⟹
                                 mpj M H h  = 𝟬(M /m H)"
apply (simp add:submodule_def, (erule conjE)+)
 apply (frule_tac c = h in subsetD[of H "carrier M"], assumption+)
 apply (subst elem_mpj[of _ H], assumption+,
        simp add:submodule_def)
 apply (simp add:qmodule_def)
 apply (rule mr_cos_h_stable[THEN sym],
        simp add:submodule_def, assumption)
done

lemma (in Module) mker_of_mpj:"submodule R M H ⟹
                                 kerM,(M /m H) (mpj M H) = H"
apply (simp add:ker_def)
apply (rule equalityI)
apply (rule subsetI, simp, erule conjE)
 apply (simp add:elem_mpj, simp add:qmodule_def)
 apply (frule_tac m = x in m_in_mr_coset [of H], assumption+)
 apply simp
apply (rule subsetI)
 apply simp
 apply (simp add:submodule_def, (erule conjE)+)
 apply (simp add:subsetD)
 apply (subst elem_mpj,
        simp add:subsetD, simp add:submodule_def) 
 apply (simp add:qmodule_def)
 apply (rule mr_cos_h_stable[THEN sym],
        simp add:submodule_def, assumption)
done

lemma (in Module) indmhom1:"⟦submodule R M H; R module N; f ∈ mHom R M N;  H ⊆ kerM,N f⟧ ⟹ ∃!g. g ∈ (mHom R (M /m H) N) ∧ (compos M g (mpj M H)) = f" 
apply (rule ex_ex1I)
apply (subgoal_tac "(λX∈set_mr_cos M H. f (SOME x. x ∈ X)) ∈ mHom R  (M /m H) N ∧ compos M (λX∈set_mr_cos M H. f (SOME x. x ∈ X)) (mpj M H) = f")
apply blast
 apply (rule conjI)
 apply (rule Module.mHom_test)
 apply (simp add:qmodule_module, assumption+)
 apply (rule conjI)
 apply (rule Pi_I)
 apply (simp add:qmodule_def, simp add:set_mr_cos_def, erule bexE, simp)
 apply (simp add:indmhom_someTr2, simp add:mHom_mem)

 apply (rule conjI)
 apply (simp add:qmodule_def)

 apply (rule conjI, (rule ballI)+)
 apply (simp add:qmodule_def, simp add:set_mr_cos_def, (erule bexE)+, simp)
 apply (simp add:mr_cospOpTr,
        frule_tac x = a and y = aa in ag_pOp_closed, assumption+)
  apply (simp add:indmhom_someTr2, simp add:mHom_add)
  apply (rule impI) 
  apply (frule_tac x = "a ± aa" in bspec, assumption+, simp)

 apply ((rule ballI)+,
        simp add:qmodule_def, simp add:set_mr_cos_def, erule bexE, simp,
        simp add:mr_cos_sprodTr,
        frule_tac a = a and m = aa in sc_mem, assumption)
 apply (simp add:indmhom_someTr2, simp add:mHom_lin,
        rule impI,
        frule_tac x = "a ⋅s aa" in bspec, assumption, simp)
 apply (rule mHom_eq[of N _ f], assumption)
 apply (rule Module.mHom_compos[of "M /m H" R M N "mpj M H" 
         "λX∈set_mr_cos M H. f (SOME x. x ∈ X)"]) apply (
        simp add:qmodule_module, rule Module_axioms, assumption,
        simp add:mpj_mHom)
 apply (rule Module.mHom_test,
        simp add:qmodule_module, assumption)
 apply (rule conjI,
        rule Pi_I,
        clarsimp simp: qmodule_def set_mr_cos_def indmhom_someTr2 mHom_mem)
 apply (rule conjI,
       simp add:qmodule_def)
 apply (rule conjI,
        (rule ballI)+, simp add:qmodule_def, simp add:set_mr_cos_def,
        (erule bexE)+, simp add:mr_cospOpTr,
        frule_tac x = a and y = aa in ag_pOp_closed, assumption+,
        simp add:indmhom_someTr2 mHom_add,
        rule impI, 
        frule_tac x = "a ± aa" in bspec, assumption, simp) 
 apply ((rule ballI)+, simp add:qmodule_def set_mr_cos_def, erule bexE, simp,
        simp add:mr_cos_sprodTr,
        frule_tac a = a and m = aa in sc_mem, assumption,
        simp add:indmhom_someTr2 mHom_lin,
        rule impI,
        frule_tac x = "a ⋅s aa" in bspec, assumption, simp, 
        assumption+) 
 apply (rule ballI, simp add:compos_def compose_def elem_mpj,
        simp add:indmhom_someTr2,
        rule impI, simp add:set_mr_cos_def,
        frule_tac x = m in bspec, assumption, simp)
 
 apply (erule conjE)+ 
 apply (rule_tac f = g and g = y in Module.mHom_eq[of "M /m H" R N],
        simp add:qmodule_module, assumption+) 
 apply (rule ballI, simp add:qmodule_def, fold qmodule_def,
        simp add:set_mr_cos_def, erule bexE, simp)
 apply (rotate_tac -3, frule sym, thin_tac "compos M y (mpj M H) = f", 
        simp)
 apply (frule_tac f = "compos M g (mpj M H)" and g = "compos M y (mpj M H)"
        and x = a in eq_fun_eq_val,
        thin_tac "compos M g (mpj M H) = compos M y (mpj M H)")
 apply (simp add:compos_def compose_def elem_mpj)
done

definition
  mQmp :: "[('a, 'r, 'm) Module_scheme, 'a set, 'a set] ⇒ 
                                                   ('a set ⇒ 'a set)" where
  "mQmp M H N = (λX∈ set_mr_cos M H. {z. ∃ x ∈ X. ∃ y ∈ N. (y ±M x = z)})"
             (* H ⊆ N *)

abbreviation
  MQP  ("(3Mp_  _,_)" [82,82,83]82) where
  "MpM H,N == mQmp M H N"

 (* "⟦ R Module M; H ⊆ N ⟧ ⟹ MpM H,N ∈ rHom (M / m H) (M /m N)"  *)

lemma (in Module) mQmpTr0:"⟦submodule R M H; submodule R M N; H ⊆ N;
 m ∈ carrier M⟧ ⟹  mQmp M H N (m ⊎M H) = m ⊎M N"
apply (frule set_mr_cos_mem [of H m], assumption+)
apply (simp add:mQmp_def)
apply (rule equalityI)
 apply (rule subsetI, simp, (erule bexE)+, rotate_tac -1, frule sym,
        thin_tac "y ± xa = x", simp)
 apply (frule_tac x = xa in x_in_mr_coset[of H m], assumption+, erule bexE,
        rotate_tac -1, frule sym, thin_tac "m ± h = xa", simp)
 apply (unfold submodule_def, frule conjunct1, rotate_tac 1, frule conjunct1,
        fold submodule_def,
        frule_tac c = y in subsetD[of N "carrier M"], assumption+,
        frule_tac c = h in subsetD[of H "carrier M"], assumption+,
        simp add:ag_pOp_assoc[THEN sym],
        simp add:ag_pOp_commute[of _ m], simp add:ag_pOp_assoc,
        frule_tac c = h in subsetD[of H N], assumption+,
        frule_tac h = y and k = h in submodule_pOp_closed[of N], assumption+,
        frule_tac h1 = "y ± h" in mr_cos_h_stable1[THEN sym, of N m], 
        assumption+, simp)
 apply (rule m_in_mr_coset, assumption+,
        rule ag_pOp_closed, assumption+, simp add:subsetD)

 apply (rule subsetI, simp,
        frule_tac x = x in x_in_mr_coset[of N m], assumption+,
        erule bexE, frule sym, thin_tac "m ± h = x", simp,
        simp add:submodule_def[of R M N], frule conjunct1, fold submodule_def,
        frule_tac c = h in subsetD[of N "carrier M"], assumption+)
apply (frule_tac m_in_mr_coset[of H m], assumption+,
        subst ag_pOp_commute[of m], assumption+)
 apply blast
done

  (* show mQmp M H N is a welldefined map from M/H to M/N. step2 *)
lemma (in Module) mQmpTr1:"⟦submodule R M H; submodule R M N; H ⊆ N;
 m ∈ carrier M; n ∈ carrier M; m ⊎M H = n ⊎M H⟧ ⟹  m ⊎M N = n ⊎M N"
apply (frule_tac m_in_mr_coset [of H m], assumption+)
apply simp
apply (frule_tac x_in_mr_coset [of H n m], assumption+) 
apply (erule bexE, rotate_tac -1, frule sym, thin_tac "n ± h = m", simp)
apply (frule_tac c = h in subsetD [of "H" "N"], assumption+)
apply (rule mr_cos_h_stable1[of N n], assumption+)
done
   
lemma (in Module) mQmpTr2:"⟦submodule R M H; submodule R M N; H ⊆ N ; 
        X ∈ carrier (M /m H)⟧ ⟹ (mQmp M H N) X ∈ carrier (M /m N)" 
apply (simp add:qmodule_def)
apply (simp add:set_mr_cos_def)
apply (erule bexE, simp)
 apply (frule_tac m = a in mQmpTr0 [of H N], assumption+)
apply blast
done

lemma (in Module) mQmpTr2_1:"⟦submodule R M H; submodule R M N; H ⊆ N ⟧
 ⟹ mQmp M H N ∈ carrier (M /m H) → carrier (M /m N)"
by (simp add:mQmpTr2)

lemma (in Module) mQmpTr3:"⟦submodule R M H; submodule R M N; H ⊆ N ; 
X ∈ carrier (M /m H); Y ∈ carrier (M /m H)⟧ ⟹ (mQmp M H N) (mr_cospOp M H X Y) = mr_cospOp M N ((mQmp M H N) X) ((mQmp M H N) Y)" 
apply (simp add:qmodule_def)
apply (simp add:set_mr_cos_def)
apply ((erule bexE)+, simp)
apply (simp add:mr_cospOpTr)
apply (frule_tac x = a and y = aa in ag_pOp_closed, assumption+)
apply (subst mQmpTr0, assumption+)+
apply (subst mr_cospOpTr, assumption+) 
apply simp
done
     
lemma (in Module) mQmpTr4:"⟦submodule R M H; submodule R M N; H ⊆ N;
                            a ∈ N⟧ ⟹ mr_coset a (mdl M N) H = mr_coset a M H"
apply (simp add:mr_coset_def)
 apply (unfold submodule_def[of R M N], frule conjunct1, fold submodule_def,
        frule subsetD[of N "carrier M" a], assumption+)
apply (rule equalityI)
 apply (rule subsetI)
 apply (frule mdl_is_module[of N])
 apply (frule_tac x = x in Module.x_in_mr_coset[of "mdl M N" R H a])
 apply (simp add:submodule_of_mdl)
 apply (simp add:mdl_carrier)
 apply assumption+
 apply (erule bexE)
 apply (unfold submodule_def[of R M H], frule conjunct1, fold submodule_def)
 apply (frule_tac c = h in subsetD[of H "carrier M"], assumption+)
 apply (thin_tac "x ∈ a ⊎mdl M N H", thin_tac "R module mdl M N",
        simp add:mdl_def)
 apply (frule sym, thin_tac "a ± h = x", simp)
 apply (subst mr_cos_h_stable1[THEN sym, of H a], assumption+)
 apply (frule_tac x = a and y = h in ag_pOp_closed, assumption+)
 apply (rule m_in_mr_coset, assumption+)

apply (rule subsetI)
 apply (frule_tac x = x in x_in_mr_coset[of H a], assumption+)
 apply (erule bexE, frule sym, thin_tac "a ± h = x", simp)
 apply (frule mdl_is_module[of N])
 apply (frule submodule_of_mdl[of H N], assumption+)
 apply (subst Module.mr_cos_h_stable1[THEN sym, of "mdl M N" R H a],
         assumption+, simp add:mdl_carrier, simp)
 apply (subgoal_tac "a ± h = a ±mdl M N h", simp)
 apply (rule Module.m_in_mr_coset[of "mdl M N" R H], assumption+)
 apply (frule Module.module_is_ag[of "mdl M N" R])
 apply (rule aGroup.ag_pOp_closed, assumption,
        simp add:mdl_carrier, simp add:mdl_carrier subsetD)
 apply (subst mdl_def, simp)
done

lemma (in Module) mQmp_mHom:"⟦submodule R M H; submodule R M N; H ⊆ N⟧ ⟹
                  (MpM H,N) ∈ mHom R (M /m H) (M /m N)"
apply (simp add:mHom_def)
apply (rule conjI)  
 apply (simp add:aHom_def)
 apply (simp add:mQmpTr2_1)
apply (rule conjI)
 apply (simp add:mQmp_def extensional_def qmodule_def)
 apply (rule ballI)+
 apply (frule_tac X1 = a and Y1 = b in mQmpTr3 [THEN sym, of H N],
                                               assumption+) 
 apply (simp add:qmodule_def)

apply (rule ballI)+
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def)
 apply (erule bexE, simp)
 apply (subst mr_cos_sprodTr, assumption+)
 apply (frule_tac a = a and m = aa in sc_mem, assumption)
 apply (simp add:mQmpTr0)
 apply (subst mr_cos_sprodTr, assumption+)
apply simp
done
    
lemma (in Module) Mp_surjec:"⟦submodule R M H; submodule R M N; H ⊆ N⟧ ⟹ 
                surjec(M /m H),(M /m N) (MpM H,N)" 
apply (simp add:surjec_def)
 apply (frule mQmp_mHom [of H N], assumption+)
 apply (rule conjI)
 apply (simp add:mHom_def)
apply (rule surj_to_test)
 apply (simp add:mHom_def aHom_def)
 apply (rule ballI)
 apply (thin_tac "MpM  H,N ∈ mHom R (M /m H) (M /m N)")
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def, erule bexE, simp)
 apply (frule_tac m = a in mQmpTr0 [of H N], assumption+)
 apply blast
done

lemma (in Module) kerQmp:"⟦submodule R M H; submodule R M N; H ⊆ N⟧ 
 ⟹ ker(M /m H),(M /m N) (MpM H,N) = carrier ((mdl M N) /m H)"   
apply (simp add:ker_def)
apply (rule equalityI)
 apply (rule subsetI)
 apply (simp add:CollectI, erule conjE)
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def [of "mdl M N" "H"])
 apply (simp add:set_mr_cos_def)
 apply (erule bexE, simp)
 apply (simp add:mQmpTr0)
 apply (frule_tac m = a in m_in_mr_coset[of N], assumption+, simp)
 apply (frule_tac a = a in mQmpTr4[of H N], assumption+,
        simp add:mr_coset_def,
        rotate_tac -1, frule sym,thin_tac "a ⊎mdl M N H = a ⊎M H",
        simp only:mdl_carrier, blast)

 apply (rule subsetI)
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def [of "mdl M N" "H"])
 apply (erule bexE, simp)
 apply (simp add:mdl_carrier)
  apply (frule_tac a = a in mQmpTr4[of H N], assumption+,
         simp add:mr_coset_def)
 apply (thin_tac "a ⊎mdl M N H = a ⊎M H")
 apply (unfold submodule_def[of R M N], frule conjunct1, fold submodule_def,
        frule_tac c = a in subsetD[of N "carrier M"], assumption+)
 apply (rule conjI) 
 apply (simp add:set_mr_cos_def, blast)
 apply (simp add:mQmpTr0)
  apply (simp add:mr_cos_h_stable [THEN sym])
done

lemma (in Module) misom2Tr:"⟦submodule R M H; submodule R M N; H ⊆ N⟧ ⟹ 
            (M /m H) /m (carrier ((mdl M N) /m H)) ≅R (M /m N)"
apply (frule mQmp_mHom [of H N], assumption+)
apply (frule qmodule_module [of H])
apply (frule qmodule_module [of N]) thm Module.indmhom
apply (frule Module.indmhom [of "M /m H" R "M /m N" "MpM H,N"], assumption+)
apply (simp add:kerQmp)
apply (subgoal_tac "bijec((M /m H) /m (carrier((mdl M N) /m H))),(M /m N)
 (indmhom R (M /m H) (M /m N) (mQmp M H N))")
apply (simp add:misomorphic_def) apply blast
apply (simp add:bijec_def)
apply (rule conjI)
 apply (simp add:kerQmp [THEN sym])
 apply (rule Module.indmhom_injec [of "M /m H" R "M /m N" "MpM H,N"], assumption+)
apply (frule Mp_surjec [of H N], assumption+)
 apply (simp add:kerQmp [THEN sym])
 apply (rule Module.indmhom_surjec1, assumption+)
done

lemma (in Module) eq_class_of_Submodule:"⟦submodule R M H; submodule R M N; 
         H ⊆ N⟧ ⟹ carrier ((mdl M N) /m H) = N s/M H"
apply (rule equalityI)
 apply (rule subsetI) apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def) apply (erule bexE, simp)
 apply (frule_tac a = a in mQmpTr4 [of H N], assumption+)
 apply (simp add:mdl_def) apply (simp add:mr_coset_def)
 apply (simp add:sub_mr_set_cos_def)
 apply (simp add:mdl_carrier, blast)

apply (rule subsetI)
apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def)
 apply (simp add:sub_mr_set_cos_def)
 apply (erule bexE, simp add:mdl_carrier)
 apply (frule_tac a1 = n in mQmpTr4[THEN sym, of H N], assumption+)
 apply (simp add:mr_coset_def)
 apply blast
done

theorem (in Module) misom2:"⟦submodule R M H; submodule R M N; H ⊆ N⟧ ⟹ 
                           (M /m H) /m (N s/M H) ≅R (M /m N)"
apply (frule misom2Tr [of H N], assumption+)
apply (simp add:eq_class_of_Submodule)
done

primrec natm :: "('a, 'm) aGroup_scheme  => nat ⇒ 'a  => 'a"
where
  natm_0:  "natm M 0 x = 𝟬M"
| natm_Suc:  "natm M (Suc n) x = (natm M n x) ±M x"

definition
  finitesum_base :: "[('a, 'r, 'm) Module_scheme, 'b set, 'b ⇒ 'a set]
                      ⇒ 'a set " where
  "finitesum_base M I f = ⋃{f i | i. i ∈ I}" 

definition
  finitesum :: "[('a, 'r, 'm) Module_scheme, 'b set, 'b ⇒ 'a set]
                      ⇒ 'a set " where
  "finitesum M I f = {x. ∃n. ∃g. g ∈ {j. j ≤ (n::nat)} → finitesum_base M I f
                                           ∧ x =  nsum M g n}"


lemma (in Module) finitesumbase_sub_carrier:"f ∈ I → {X. submodule R M X} ⟹
             finitesum_base M I f ⊆ carrier M"
apply (simp add:finitesum_base_def)
apply (rule subsetI)
 apply (simp add:CollectI)
 apply (erule exE, erule conjE, erule exE, erule conjE)
 apply (frule_tac x = i in funcset_mem[of f I "{X. submodule R M X}"], 
         assumption+, simp)
 apply (thin_tac "f ∈ I → {X. submodule R M X}", unfold submodule_def,
        frule conjunct1, fold submodule_def, simp add:subsetD)
done

lemma (in Module) finitesum_sub_carrier:"f ∈ I → {X. submodule R M X} ⟹
                       finitesum M I f ⊆ carrier M"
apply (rule subsetI, simp add:finitesum_def)
apply ((erule exE)+, erule conjE, simp)
apply (frule finitesumbase_sub_carrier)
apply (rule nsum_mem, rule allI, rule impI)
apply (frule_tac x = j and f = g and A = "{j. j ≤ n}" and
        B = "finitesum_base M I f" in funcset_mem, simp)
apply (simp add:subsetD)
done

lemma (in Module) finitesum_inc_zero:"⟦f ∈ I → {X. submodule R M X}; I ≠ {}⟧
      ⟹   𝟬 ∈ finitesum M I f"
apply (simp add:finitesum_def)
apply (frule nonempty_ex)
apply (subgoal_tac "∀i. i∈I ⟶ (∃n g. g ∈ {j. j ≤ (n::nat)} → 
                    finitesum_base M I f ∧ 𝟬M = Σe M g n)")
apply blast 
apply (rule allI, rule impI)
apply (subgoal_tac "(λx∈{j. j ≤ (0::nat)}. 𝟬) ∈ 
                    {j. j ≤ (0::nat)} → finitesum_base M I f ∧
                    𝟬M = Σe M (λx∈{j. j ≤ (0::nat)}. 𝟬) 0")
apply blast
apply (rule conjI)
apply (rule Pi_I) 
 apply (simp add:finitesum_base_def, thin_tac "∃x. x ∈ I")
 apply (frule_tac x = i in funcset_mem[of f I "{X. submodule R M X}"], 
        assumption+)
 apply (frule_tac x = i in funcset_mem [of "f" "I" "{X. submodule R M X}"],
                                              assumption+, simp)
 apply (frule_tac H = "f i" in submodule_inc_0)
 apply blast

 apply simp
done

lemma (in Module) finitesum_mOp_closed:
     "⟦f ∈ I → {X. submodule R M X}; I ≠ {}; a ∈ finitesum M I f⟧ ⟹
                  -a a ∈ finitesum M I f"
apply (simp add:finitesum_def)
apply ((erule exE)+, erule conjE)
  apply (frule finitesumbase_sub_carrier [of f I])
  apply (frule_tac f = g and A = "{j. j ≤ n}" and B = "finitesum_base M I f"
          and ?B1.0 = "carrier M" in extend_fun, assumption+)
  apply (frule sym, thin_tac "a = Σe M g n")
  apply (cut_tac n = n and f = g in nsum_minus,
         rule allI, simp add:Pi_def, simp)

 apply (subgoal_tac "(λx∈{j. j ≤ n}. -a (g x)) ∈ {j. j ≤ n} → 
                                                 finitesum_base M I f")
 apply blast
 apply (rule Pi_I, simp)
 apply (frule_tac f = g and A = "{j. j ≤ n}" and B = "finitesum_base M I f" 
        and  x = x in funcset_mem, simp)
 apply (simp add:finitesum_base_def)
 apply (erule exE, erule conjE, erule exE, erule conjE)
 apply (frule_tac f = f and A = I and B = "{X. submodule R M X}" and
  x = i in funcset_mem, assumption+, simp add:CollectI)
 apply (thin_tac "f ∈ I → {X. submodule R M X}")
 apply (simp add:submodule_def, (erule conjE)+,
        frule_tac H = "f i" and x = "g x" in asubg_mOp_closed, assumption+) 
 apply blast
done

lemma (in Module) finitesum_pOp_closed:
 "⟦f ∈ I → {X. submodule R M X}; a ∈ finitesum M I f;  b ∈ finitesum M I f⟧
           ⟹  a ± b ∈ finitesum M I f"
apply (simp add:finitesum_def) 
apply ((erule exE)+, (erule conjE)+)
apply (frule_tac f = g and n = n and A = "finitesum_base M I f" and
       g = ga and m = na and B = "finitesum_base M I f" in jointfun_hom0,
       assumption+, simp)
apply (cut_tac finitesumbase_sub_carrier[of f I],
       cut_tac n1 = n and f1 = g and m1 = na and g1 = ga in 
                 nsum_add_nm[THEN sym], rule allI, rule impI,
       frule_tac x = j and f = g and A = "{j. j ≤ n}" and
        B = "finitesum_base M I f" in funcset_mem, simp,
       simp add:subsetD,
       rule allI, rule impI,
       frule_tac x = j and f = ga and A = "{j. j ≤ na}" and
        B = "finitesum_base M I f" in funcset_mem, simp,
       simp add:subsetD)
apply blast
apply assumption
done

lemma (in Module) finitesum_sprodTr:"⟦f ∈ I → {X. submodule R M X}; I ≠ {};
       r ∈ carrier R⟧  ⟹ g ∈{j. j ≤ (n::nat)} → (finitesum_base M I f)
              ⟶ r ⋅s (nsum M g n) =  nsum M (λx. r ⋅s (g x)) n"
apply (induct_tac n)
 apply (rule impI)
 apply simp
apply (rule impI)
apply (frule func_pre) apply simp
apply (frule finitesumbase_sub_carrier [of f I])
 apply (frule_tac f = g and A = "{j. j ≤ Suc n}" in extend_fun [of _ _ "finitesum_base M I f" "carrier M"], assumption+)
 apply (thin_tac "g ∈ {j. j ≤ Suc n} → finitesum_base M I f",
        thin_tac "g ∈ {j. j ≤ n} → finitesum_base M I f",
        frule func_pre)
 apply (cut_tac n = n in nsum_mem [of _ g])
 apply (rule allI, simp add:Pi_def)
 apply (frule_tac x = "Suc n" in funcset_mem [of "g" _ "carrier M"], simp)
 apply (subst sc_r_distr, assumption+)
 apply simp
done

lemma (in Module) finitesum_sprod:"⟦f ∈ I → {X. submodule R M X}; I ≠ {}; 
      r ∈ carrier R; g ∈{j. j ≤ (n::nat)} → (finitesum_base M I f) ⟧ ⟹
                       r ⋅s (nsum M g n) =  nsum M (λx. r ⋅s (g x)) n"
apply (simp add:finitesum_sprodTr)
done

lemma (in Module) finitesum_subModule:"⟦f ∈ I → {X. submodule R M X}; I ≠ {}⟧
                   ⟹ submodule R M (finitesum M I f)"
apply (simp add:submodule_def [of _ _ "(finitesum M I f)"])
apply (simp add:finitesum_sub_carrier)
apply (rule conjI)
 apply (rule asubg_test)
 apply (simp add:finitesum_sub_carrier)
 apply (frule finitesum_inc_zero, assumption, blast) 

 apply (rule ballI)+
 apply (rule finitesum_pOp_closed, assumption+,
        rule finitesum_mOp_closed, assumption+)

 apply ((rule allI)+, rule impI, erule conjE)
 apply (simp add:finitesum_def, (erule exE)+, erule conjE, simp)
 apply (simp add:finitesum_sprod)
 apply (subgoal_tac "(λx. a ⋅s g x) ∈ {j. j ≤ n} → finitesum_base M I f",
        blast)
 apply (rule Pi_I)
 apply (frule_tac x = x and f = g and A = "{j. j ≤ n}" in 
                  funcset_mem[of _ _ "finitesum_base M I f"], assumption+,
        thin_tac "g ∈ {j. j ≤ n} → finitesum_base M I f",
        simp add:finitesum_base_def, erule exE, erule conjE, erule exE,
        erule conjE, simp)
 apply (frule_tac x = i and f = f and A = I in 
        funcset_mem[of _ _ "{X. submodule R M X}"], assumption+, simp,
        frule_tac H = "f i" and a = a and h = "g x" in submodule_sc_closed,
        assumption+)
apply blast
done

lemma (in Module) sSum_cont_H:"⟦submodule R M H; submodule R M K⟧ ⟹
                     H ⊆  H ∓ K"
apply (rule subsetI)
apply (unfold submodule_def[of R M H], frule conjunct1, fold submodule_def,
       unfold submodule_def[of R M K], frule conjunct1, fold submodule_def)
apply (simp add:set_sum) 
apply (frule submodule_inc_0 [of K])
apply (cut_tac t = x in ag_r_zero [THEN sym],
       rule submodule_subset1, assumption+)
apply blast
done

lemma (in Module) sSum_commute:"⟦submodule R M H; submodule R M K⟧ ⟹
                       H ∓ K =  K ∓ H"
apply (unfold submodule_def[of R M H], frule conjunct1, fold submodule_def,
       unfold submodule_def[of R M K], frule conjunct1, fold submodule_def)   
apply (rule equalityI)
apply (rule subsetI) 
apply (simp add:set_sum)
apply ((erule bexE)+, simp)
apply (frule_tac c = h in subsetD[of H "carrier M"], assumption+,
       frule_tac c = k in subsetD[of K "carrier M"], assumption+)
apply (subst ag_pOp_commute, assumption+)
apply blast

apply (rule subsetI)
apply (simp add:set_sum)
apply ((erule bexE)+, simp)
apply (frule_tac h = h in submodule_subset1[of K ], assumption+,
       frule_tac h = k in submodule_subset1[of H ], assumption+)
apply (subst ag_pOp_commute, assumption+)
apply blast
done

lemma (in Module) Sum_of_SubmodulesTr:"⟦submodule R M H; submodule R M K⟧ ⟹
      g ∈ {j. j ≤ (n::nat)} → H ∪ K ⟶ Σe M g n ∈ H ∓ K"
apply (induct_tac n)
 apply (rule impI)
 apply simp
 apply (frule submodule_subset[of H],
        frule submodule_subset[of K])
 apply (simp add:set_sum)
 apply (erule disjE)
 apply (frule_tac c = "g 0" in subsetD[of H "carrier M"], assumption+,
        frule_tac t = "g 0" in ag_r_zero[THEN sym]) apply (
        frule submodule_inc_0[of K], blast)
 apply (frule_tac c = "g 0" in subsetD[of K "carrier M"], assumption+,
        frule_tac t = "g 0" in ag_l_zero[THEN sym]) apply (
        frule submodule_inc_0[of H], blast)
apply simp

apply (rule impI, frule func_pre, simp)
 apply (frule submodule_subset[of H],
        frule submodule_subset[of K])
 apply (simp add:set_sum[of H K], (erule bexE)+, simp)
 apply (frule_tac x = "Suc n" and f = g and A = "{j. j ≤ Suc n}" and
        B = "H ∪ K" in funcset_mem, simp,
        thin_tac "g ∈ {j. j ≤ n} → H ∪ K",
        thin_tac "g ∈ {j. j ≤ Suc n} → H ∪ K",
        thin_tac e M g n = h ± k", simp)
 apply (erule disjE)
 apply (frule_tac h = h in submodule_subset1[of H], assumption,
        frule_tac h = "g (Suc n)" in submodule_subset1[of H], assumption,
        frule_tac h = k in submodule_subset1[of K], assumption) 
 apply (subst ag_pOp_assoc, assumption+)
  apply (frule_tac x = k and y = "g (Suc n)" in ag_pOp_commute, assumption+,
         simp, subst ag_pOp_assoc[THEN sym], assumption+)
  apply (frule_tac h = h and k = "g (Suc n)" in submodule_pOp_closed[of H],
         assumption+, blast)
 apply (frule_tac h = h in submodule_subset1[of H], assumption,
        frule_tac h = "g (Suc n)" in submodule_subset1[of K], assumption,
        frule_tac h = k in submodule_subset1[of K], assumption) 
 apply (subst ag_pOp_assoc, assumption+,
        frule_tac h = k and k = "g (Suc n)" in submodule_pOp_closed[of K],
         assumption+, blast)
done

lemma (in Module) sSum_two_Submodules:"⟦submodule R M H; submodule R M K⟧ ⟹
                       submodule R M (H ∓ K)"
apply (subst submodule_def) 
 apply (frule submodule_asubg[of H],
        frule submodule_asubg[of K])
 apply (frule plus_subgs[of H K], assumption, simp add:asubg_subset)

apply (rule allI)+
apply (rule impI, erule conjE, frule asubg_subset[of H], 
       frule asubg_subset[of K])
 apply (simp add:set_sum[of H K], (erule bexE)+, simp)
 apply (frule_tac H = H and a = a and h = h in submodule_sc_closed, 
                  assumption+,
        frule_tac H = K and a = a and h = k in submodule_sc_closed, 
                  assumption+)
 apply (frule_tac c = h in subsetD[of H "carrier M"], assumption+,
        frule_tac c = k in subsetD[of K "carrier M"], assumption+,
        simp add:sc_r_distr)
 apply blast
done

definition
  iotam :: "[('a, 'r, 'm) Module_scheme, 'a set, 'a set] ⇒ ('a ⇒ 'a)"
      ("(3ιm_ _,_)" [82, 82, 83]82) where
  "ιmM H,K = (λx∈H. (x ±M 𝟬M))"  (** later we define miota. This is not 
 equal to iotam **) 

lemma (in Module) iotam_mHom:"⟦submodule R M H; submodule R M K⟧
                           ⟹ ιmM H,K ∈ mHom R (mdl M H) (mdl M (H ∓ K))"
apply (simp add:mHom_def)
apply (rule conjI)
 apply (simp add:aHom_def)
 apply (simp add:mdl_def)
 apply (rule conjI)
 apply (rule Pi_I)
 apply (simp add:iotam_def)
 apply (frule submodule_subset[of H], frule submodule_subset[of K],
        simp add:set_sum)
 apply (frule submodule_inc_0 [of K])
 apply blast
apply (rule conjI)
 apply (simp add:iotam_def extensional_def mdl_def)
apply (rule ballI)+
 apply (simp add:mdl_def iotam_def)
 apply (frule_tac h = a and k = b in submodule_pOp_closed [of H],
                                     assumption+, simp)
 apply (frule submodule_subset[of H], 
        frule_tac c = a in subsetD[of H "carrier M"], assumption) apply (
        simp add:ag_r_zero) 
 apply ( frule_tac c = b in subsetD[of H "carrier M"], assumption,
        subst ag_pOp_assoc, assumption+,
        simp add:ag_inc_zero, simp)

apply (rule ballI)+
 apply (simp add:iotam_def mdl_def)
 apply (simp add:submodule_sc_closed)
 apply (frule submodule_inc_0[of K]) 
 apply (frule submodule_asubg[of H], frule submodule_asubg[of K],
        simp add:mem_sum_subgs)

 apply (frule_tac a = a and h = m in submodule_sc_closed, assumption+,
        frule submodule_subset[of H],
        frule_tac c = m in subsetD[of H "carrier M"], assumption+,
        frule_tac c = "a ⋅s m" in subsetD[of H "carrier M"], assumption+)
 apply (simp add:ag_r_zero)
done

lemma (in Module) mhomom3Tr:"⟦submodule R M H; submodule R M K⟧ ⟹
                         submodule R (mdl M (H ∓ K)) K"
apply (subst submodule_def) 
apply (rule conjI)
 apply (simp add:mdl_def)
 apply (subst sSum_commute, assumption+) 
 apply (simp add:sSum_cont_H)
apply (rule conjI)
 apply (rule aGroup.asubg_test)
 apply (frule sSum_two_Submodules [of H K], assumption+)
 apply (frule mdl_is_module [of  "(H ∓ K)"])
 apply (rule Module.module_is_ag, assumption+)
apply (simp add:mdl_def)
 apply (subst sSum_commute, assumption+)   
  apply (simp add:sSum_cont_H)
 apply (frule submodule_inc_0 [of K])
 apply (simp add:nonempty)
apply (rule ballI)+
 apply (simp add:mdl_def)
 apply (rule submodule_pOp_closed, assumption+)
 apply (rule submodule_mOp_closed, assumption+)
apply ((rule allI)+, rule impI)
 apply (simp add:mdl_def, erule conjE)
 apply (frule sSum_cont_H[of K H], assumption,
        simp add:sSum_commute[of K H])
 apply (simp add:subsetD submodule_sc_closed)
done

lemma (in Module) mhomom3Tr0:"⟦submodule R M H; submodule R M K⟧
     ⟹ compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K)
        ∈ mHom R (mdl M H) (mdl M (H ∓ K) /m K)"
apply (frule mdl_is_module [of H])
apply (frule mhomom3Tr[of H K], assumption+)
apply (frule sSum_two_Submodules [of H K], assumption+)
apply (frule mdl_is_module [of  "H ∓ K"])
apply (frule iotam_mHom [of H K], assumption+) thm Module.mpj_mHom
apply (frule Module.mpj_mHom [of "mdl M (H ∓ K)" R "K"], assumption+)
apply (rule  Module.mHom_compos[of "mdl M (H ∓ K)" R "mdl M H"], assumption+)
apply (simp add:Module.qmodule_module, assumption)
apply (simp add:mpj_mHom)
done

lemma (in Module) mhomom3Tr1:"⟦submodule R M H; submodule R M K⟧ ⟹
  surjec(mdl M H),((mdl M (H ∓ K))/m K) 
    (compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K))"
apply (simp add:surjec_def)
apply (frule mhomom3Tr0 [of H K], assumption+)
apply (rule conjI)
apply (simp add:mHom_def)
apply (rule surj_to_test)
 apply (simp add:mHom_def aHom_def)
apply (rule ballI)
 apply (simp add:compos_def compose_def)
 apply (thin_tac "(λx∈carrier (mdl M H). mpj (mdl M (H ∓ K)) K ((ιmM H,K) x))
         ∈ mHom R (mdl M H) (mdl M (H ∓ K) /m K)")
 apply (simp add:qmodule_def)
 apply (simp add:set_mr_cos_def)
 apply (erule bexE, simp)
 apply (simp add:mdl_carrier)
 apply (simp add:iotam_def)
 apply (simp add:mpj_def)
 apply (frule sSum_two_Submodules[of H K], assumption+)
 apply (simp add:mdl_carrier)
 apply (subgoal_tac "∀aa∈H. aa ± 𝟬 ∈ H ∓ K", simp)
 apply (frule submodule_subset[of H], frule submodule_subset[of K],
        thin_tac "∀aa∈H. aa ± 𝟬 ∈ H ∓ K",
        simp add:set_sum, (erule bexE)+) 
        apply (simp add:set_sum[THEN sym])
 apply (frule mdl_is_module[of "H ∓ K"],
        frule mhomom3Tr[of H K], assumption+)
 apply (frule_tac m = h and h = k in Module.mr_cos_h_stable1[of "mdl M (H ∓ K)"
        R K], assumption+)
 apply (simp add:mdl_carrier)
 apply (frule sSum_cont_H[of H K], assumption+, simp add:subsetD, assumption)
 apply (simp add:mdl_def, fold mdl_def)
 apply (subgoal_tac "∀a∈H. a ± 𝟬 = a", simp, blast)
 apply (rule ballI)
 apply (frule_tac c = aa in subsetD[of H "carrier M"], assumption+,
        simp add:ag_r_zero)
 apply (rule ballI)
 apply (frule submodule_inc_0[of K])
 apply (rule mem_sum_subgs,
       simp add:submodule_def, simp add:submodule_def, assumption+)
done
 
lemma (in Module) mhomom3Tr2:"⟦submodule R M H; submodule R M K⟧ ⟹
  ker(mdl M H),((mdl M (H ∓ K)) /m K) 
    (compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K)) = H ∩ K"
apply (rule equalityI)
 apply (rule subsetI)
 apply (simp add:ker_def, erule conjE)
 apply (simp add:qmodule_def)
 apply (simp add:mdl_carrier) 
 apply (simp add:compos_def compose_def mdl_def iotam_def)
 apply (fold mdl_def)
apply (simp add:iotam_def mpj_def) 
 apply (frule  sSum_two_Submodules[of H K], assumption+, simp add:mdl_carrier)
 apply (frule submodule_asubg[of H], frule submodule_asubg[of K])
 apply (frule_tac h = x and k = 𝟬 in mem_sum_subgs[of H K], assumption+)
 apply (simp add:submodule_inc_0)
 apply simp apply (frule mhomom3Tr[of H K], assumption+)
 (*thm Module.m_in_mr_coset[of "mdl M (H ∓ K)" R K]
 apply (frule_tac m = "x ± 𝟬" in Module.m_in_mr_coset[of "mdl M (H ∓ K)" R K])*)
apply (frule sSum_two_Submodules[of H K], assumption,
       frule mdl_is_module [of  "H ∓ K"])
apply (frule_tac m = "x ± 𝟬" in Module.m_in_mr_coset[of "mdl M (H ∓ K)" R K],
                          assumption+)
 apply (simp add:mdl_carrier, simp)
 apply (frule submodule_subset[of H], 
        frule_tac c = x in subsetD[of H "carrier M"], assumption+) 
 apply (simp add:ag_r_zero)

apply (rule subsetI)
 apply (simp add:ker_def)
 apply (simp add:mdl_carrier)
 apply (simp add:qmodule_def)
 apply (simp add:compos_def compose_def)
 apply (simp add:mdl_carrier)
 apply (simp add:iotam_def mpj_def)
 apply (frule sSum_two_Submodules[of H K], assumption+)
 apply (simp add:mdl_carrier)
 apply (erule conjE,
        frule submodule_inc_0[of K],
        frule submodule_asubg[of H], frule submodule_asubg[of K],
       simp add:mem_sum_subgs)
 apply (frule submodule_subset[of K]) apply (
        frule_tac c = x in subsetD[of K "carrier M"], assumption+)
 apply (simp add:ag_r_zero,
        frule mdl_is_module [of  "H ∓ K"],
        frule mhomom3Tr[of H K], assumption+)
 apply (frule_tac h1 = x in Module.mr_cos_h_stable[THEN sym, of "mdl M (H ∓ K)"
         R K], assumption+)
done

lemma (in Module) mhomom_3:"⟦submodule R M H; submodule R M K⟧ ⟹
                 (mdl M H) /m (H ∩ K) ≅R (mdl M (H ∓ K)) /m K" 
apply (frule sSum_two_Submodules [of H K], assumption+)
 apply (frule mdl_is_module [of H])
 apply (frule mdl_is_module [of K])
 apply (frule mdl_is_module [of "H ∓ K"])
 apply (frule mhomom3Tr [of H K], assumption+)
 apply (frule Module.qmodule_module [of "mdl M (H ∓ K)" R K], assumption+)
apply (simp add:misomorphic_def)
apply (frule mhomom3Tr0[of H K], assumption+)
apply (frule mhomom3Tr1[of H K], assumption+)
apply (frule Module.indmhom [of "mdl M H" R "mdl M (H ∓ K) /m K" "compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K)"], assumption+)
apply (frule Module.indmhom_injec[of "mdl M H" R "mdl M (H ∓ K) /m K"
     "compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K)"], assumption+)
apply (frule Module.indmhom_surjec1[of  "mdl M H" R "mdl M (H ∓ K) /m K" "compos (mdl M H) (mpj (mdl M (H ∓ K)) K) (ιmM H,K)"], assumption+)
apply (simp add:bijec_def)
apply (simp add:mhomom3Tr2[of H K])
apply blast
done

definition
  l_comb :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, nat] ⇒
    (nat ⇒ 'r) ⇒ (nat ⇒ 'a) ⇒ 'a" where
  "l_comb R M n s m = nsum M (λj. (s j) ⋅sM (m j)) n" 

definition
  linear_span :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, 'r set,
              'a set] ⇒ 'a set" where
  "linear_span R M A H = (if H = {} then {𝟬M} else 
                           {x. ∃n. ∃f ∈ {j. j ≤ (n::nat)} → H.
         ∃s∈{j. j ≤ (n::nat)} → A.  x = l_comb R M n s f})"

definition
  coefficient :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
               nat, nat ⇒ 'r, nat ⇒ 'a] ⇒ nat ⇒ 'r" where
  "coefficient R M n s m j = s j"

definition
  body :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, nat, nat ⇒ 'r, 
         nat ⇒ 'a] ⇒ nat ⇒ 'a" where
  "body R M n s m j = m j"

lemma (in Module) l_comb_mem_linear_span:"⟦ideal R A; H ⊆ carrier M; 
       s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ n} → H⟧ ⟹
                    l_comb R M n s f ∈ linear_span R M A H"
apply (frule_tac x = 0 in funcset_mem[of f "{j. j ≤ n}" H], simp)
 apply (frule nonempty[of "f 0" H])
 apply (simp add:linear_span_def)
 apply blast
done

lemma (in Module) linear_comb_eqTr:"H ⊆ carrier M ⟹ 
      s ∈ {j. j ≤ (n::nat)} → carrier R ∧ 
      f ∈ {j. j ≤ n} → H ∧ 
      g ∈ {j. j ≤ n} → H ∧ 
      (∀j∈{j. j ≤ n}. f j = g j) ⟶ 
      l_comb R M n s f = l_comb R M n s g" 
apply (induct_tac n)
 apply (rule impI) apply (erule conjE)+ apply (simp add:l_comb_def)
 
apply (rule impI) apply (erule conjE)+ 
 apply (frule_tac f = s in func_pre)
 apply (frule_tac f = f in func_pre)
 apply (frule_tac f = g in func_pre)
 apply (cut_tac n = n in Nsetn_sub_mem1, simp)
 apply (thin_tac "s ∈ {j. j ≤ n} → carrier R",
        thin_tac "f ∈ {j. j ≤ n} → H",
        thin_tac "g ∈ {j. j ≤ n} → H")
 apply (simp add:l_comb_def)
done
           
lemma (in Module) linear_comb_eq:"⟦H ⊆ carrier M; 
       s ∈ {j. j ≤ (n::nat)} → carrier R; f ∈ {j. j ≤ n} → H; 
       g ∈ {j. j ≤ n} → H; ∀j∈{j. j ≤ n}. f j = g j⟧  ⟹
  l_comb R M n s f = l_comb R M n s g" 
apply (simp add:linear_comb_eqTr)
done

lemma (in Module) l_comb_Suc:"⟦H ⊆ carrier M; ideal R A; 
       s ∈ {j. j ≤ (Suc n)} → carrier R; f ∈ {j. j ≤ (Suc n)} → H⟧  ⟹
       l_comb R M (Suc n) s f = l_comb R M n s f ± s (Suc n) ⋅s f (Suc n)" 
apply (simp add:l_comb_def)
done

lemma (in Module) l_comb_jointfun_jj:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ (n::nat)} → H;
        t ∈ {j. j ≤ (m::nat)} → A; g ∈ {j. j ≤ (m::nat)} → H⟧ ⟹
        nsum M (λj. (jointfun n s m t) j ⋅s (jointfun n f m g) j) n =
        nsum M (λj. s j ⋅s f j) n"
apply (cut_tac sc_Ring)
apply (rule nsum_eq)
 apply (rule allI, rule impI, simp add:jointfun_def,
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI, 
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, simp add:jointfun_def)
done

lemma (in Module) l_comb_jointfun_jj1:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ (n::nat)} → H;
        t ∈ {j. j ≤ (m::nat)} → A; g ∈ {j. j ≤ (m::nat)} → H⟧ ⟹
        l_comb R M n (jointfun n s m t) (jointfun n f m g) =
        l_comb R M n s f"
by (simp add:l_comb_def, simp add:l_comb_jointfun_jj)

lemma (in Module) l_comb_jointfun_jf:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ Suc (n + m)} → H;
        t ∈ {j. j ≤ (m::nat)} → A⟧ ⟹
        nsum M (λj. (jointfun n s m t) j ⋅s f j) n =
        nsum M (λj. s j ⋅s f j) n"
apply (cut_tac sc_Ring)
apply (rule nsum_eq)
 apply (rule allI, rule impI, simp add:jointfun_def,
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
  apply (rule allI, rule impI, 
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
  apply (rule allI, simp add:jointfun_def)
done

lemma (in Module) l_comb_jointfun_jf1:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ Suc (n + m)} → H;
        t ∈ {j. j ≤ (m::nat)} → A⟧ ⟹
        l_comb R M n (jointfun n s m t) f = l_comb R M n s f"
by (simp add:l_comb_def l_comb_jointfun_jf)

lemma (in Module) l_comb_jointfun_fj:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ Suc (n + m)} → A; f ∈ {j. j ≤ (n::nat)} → H;
        g ∈ {j. j ≤ (m::nat)} → H⟧ ⟹
        nsum M (λj. s j ⋅s (jointfun n f m g) j) n =
        nsum M (λj. s j ⋅s f j) n"
apply (cut_tac sc_Ring)
apply (rule nsum_eq)
 apply (rule allI, rule impI, simp add:jointfun_def,
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
  apply (rule allI, rule impI, 
        rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
    apply (rule allI, simp add:jointfun_def)
done

lemma (in Module) l_comb_jointfun_fj1:"⟦H ⊆ carrier M; ideal R A;
        s ∈ {j. j ≤ Suc (n + m)} → A; f ∈ {j. j ≤ (n::nat)} → H;
        g ∈ {j. j ≤ (m::nat)} → H⟧ ⟹
        l_comb R M n s (jointfun n f m g) = l_comb R M n s f"
by (simp add:l_comb_def l_comb_jointfun_fj)

lemma (in Module) linear_comb0_1Tr:"H ⊆ carrier M ⟹ 
      s ∈ {j. j ≤ (n::nat)} → {𝟬R} ∧  
      m ∈ {j. j ≤ n} → H ⟶ l_comb R M n s m = 𝟬M"
apply (induct_tac n)
 apply (rule impI) apply (erule conjE)
 apply (simp add:l_comb_def subsetD sc_0_m)

apply (rule impI) apply (erule conjE)
 apply (frule func_pre [of _ _ "{𝟬R}"])
 apply (frule func_pre [of _ _ "H"])
 apply simp
 apply (thin_tac "s ∈ {j. j ≤ n} → {𝟬R}",
        thin_tac "m ∈ {j. j ≤ n} → H")
 apply (simp add:l_comb_def)
 apply (frule_tac x = "Suc n" and f = s and A = "{j. j ≤ Suc n}" in 
        funcset_mem[of _ _ "{𝟬R}"], simp, simp,
        frule_tac x = "Suc n" and f = m and A = "{j. j ≤ Suc n}" in 
        funcset_mem[of _ _ H], simp,
        frule_tac c = "m (Suc n)" in subsetD[of H "carrier M"], assumption+,
        simp add:sc_0_m)
 apply (cut_tac ag_inc_zero)
 apply (simp add:ag_l_zero)
done

lemma (in Module) linear_comb0_1:"⟦H ⊆ carrier M; 
      s ∈ {j. j ≤ (n::nat)} → {𝟬R}; m ∈ {j. j ≤ n} → H ⟧ ⟹ 
      l_comb R M n s m = 𝟬M"
apply (simp add:linear_comb0_1Tr)
done

lemma (in Module) linear_comb0_2Tr:"ideal R A ⟹ s ∈ {j. j ≤ (n::nat)} → A 
      ∧  m ∈ {j. j ≤ n} → {𝟬M} ⟶ l_comb R M n s m = 𝟬M"
apply (induct_tac n )
 apply (rule impI) apply (erule conjE)
 apply (simp add:l_comb_def sc_a_0 sc_Ring Ring.ideal_subset)

apply (rule impI)
 apply (erule conjE)+
 apply (frule func_pre [of "s"],
        frule func_pre [of "m"], simp)
  apply (thin_tac "s ∈ {j. j ≤ n} → A",
         thin_tac "m ∈ {j. j ≤ n} → {𝟬}")
 apply (simp add:l_comb_def)
  apply (frule_tac A = "{j. j ≤ Suc n}" and x = "Suc n" in 
         funcset_mem [of "m" _ "{𝟬}"], simp+,
         frule_tac A = "{j. j ≤ Suc n}" and x = "Suc n" in 
         funcset_mem[of s _ A], simp+,
         cut_tac sc_Ring,
         frule_tac h = "s (Suc n)" in Ring.ideal_subset[of R A], assumption+)
  apply (cut_tac ag_inc_zero, simp add:sc_a_0)
 apply (simp add:ag_l_zero)
done

lemma (in Module) linear_comb0_2:"⟦ideal R A;  s ∈ {j. j ≤ (n::nat)} → A;
       m ∈ {j. j ≤ n} → {𝟬M} ⟧ ⟹  l_comb R M n s m = 𝟬M"
apply (simp add:linear_comb0_2Tr)
done

lemma (in Module) liear_comb_memTr:"⟦ideal R A; H ⊆ carrier M⟧ ⟹
 ∀s. ∀m. s ∈ {j. j ≤ (n::nat)} → A ∧ 
          m ∈ {j. j ≤ n} → H ⟶ l_comb R M n s m ∈ carrier M"
apply (induct_tac n)
 apply (rule allI)+ apply (rule impI) apply (erule conjE)
 apply (simp add: l_comb_def sc_mem Ring.ideal_subset[of R A] subsetD sc_Ring) 

apply (rule allI)+ apply (rule impI) apply (erule conjE)
 apply (frule func_pre [of _ _ "A"],
        frule func_pre [of _ _ "H"],
        drule_tac x = s in spec,
        drule_tac x = m in spec)

apply (simp add:l_comb_def)
 apply (rule ag_pOp_closed, assumption+)
 apply (rule sc_mem)
 apply (cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset subsetD)
 apply (simp add:Pi_def subsetD)
done

lemma (in Module) l_comb_mem:"⟦ideal R A; H ⊆ carrier M; 
       s ∈ {j. j ≤ (n::nat)} → A; m ∈ {j. j ≤ n} → H⟧ ⟹ 
      l_comb R M n s m ∈ carrier M"
apply (simp add:liear_comb_memTr)
done

lemma (in Module) l_comb_transpos:" ⟦ideal R A; H ⊆ carrier M;
      s ∈ {l. l ≤ Suc n} → A; f ∈ {l. l ≤ Suc n} → H;
      j < Suc n ⟧ ⟹ 
     Σe M (cmp (λk. s k ⋅s f k) (transpos j (Suc n))) (Suc n) =
       Σe M (λk. (cmp s (transpos j (Suc n))) k ⋅s
                  (cmp f (transpos j (Suc n))) k) (Suc n)"
apply (cut_tac sc_Ring)
apply (rule nsum_eq) 
 apply (rule allI, rule impI, simp add:cmp_def)
 apply (cut_tac l = ja in transpos_mem[of j "Suc n" "Suc n"],
        simp add:less_imp_le, simp, simp, assumption)
 apply (rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI, simp add:cmp_def)
 apply (frule less_imp_le[of j "Suc n"],
        frule_tac l = ja in transpos_mem[of j "Suc n" "Suc n"], simp,
        simp, assumption+)
 apply (rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI,
        simp add:cmp_def)
done

lemma (in Module) l_comb_transpos1:" ⟦ideal R A; H ⊆ carrier M;
      s ∈ {l. l ≤ Suc n} → A; f ∈ {l. l ≤ Suc n} → H; j < Suc n ⟧ ⟹ 
 l_comb R M (Suc n) s f = 
  l_comb R M (Suc n) (cmp s (transpos j (Suc n))) (cmp f (transpos j (Suc n)))"
apply (cut_tac sc_Ring)
apply (frule l_comb_transpos[THEN sym, of A H s n f j], assumption+)
 apply (simp del:nsum_suc add:l_comb_def,
        thin_tac e M (λk. (cmp s (transpos j (Suc n))) k ⋅s
               (cmp f (transpos j (Suc n))) k) (Suc n) =
     Σe M (cmp (λk. s k ⋅s f k) (transpos j (Suc n))) (Suc n)")
 apply (cut_tac addition2[of "λj. s j ⋅s f j" n "transpos j (Suc n)"],
         simp)
 apply (rule Pi_I, rule sc_mem,
          simp add:Pi_def Ring.ideal_subset,
          simp add:Pi_def subsetD)
 apply (rule_tac i = j and n = "Suc n" and j = "Suc n" in transpos_hom,
        simp add:less_imp_le, simp, simp)
 apply (rule_tac i = j and n = "Suc n" and j = "Suc n" in transpos_inj,
         simp add:less_imp_le, simp, simp)
done

lemma (in Module) sc_linear_span:"⟦ideal R A; H ⊆ carrier M; a ∈ A;
 h ∈ H⟧ ⟹ a ⋅s h ∈ linear_span R M A H"
apply (simp add:linear_span_def)
 apply (simp add:nonempty)
 apply (simp add:l_comb_def)
 apply (subgoal_tac "(λk∈{j. j ≤ (0::nat)}. a) ∈{j. j ≤ 0} → A")
 apply (subgoal_tac "(λk∈{j. j ≤ 0}. h) ∈ {j. j ≤ (0::nat)} → H")
 apply (subgoal_tac "a ⋅s h = 
 Σe M (λj. (λk∈{j. j ≤ (0::nat)}. a) j ⋅s (λk∈{j. j ≤ (0::nat)}. h) j) 0")
 apply blast
 apply simp+
done

lemma (in Module) l_span_cont_H:"H ⊆ carrier M ⟹ 
                      H ⊆ linear_span R M (carrier R) H"            
apply (rule subsetI)
apply (cut_tac sc_Ring,
       cut_tac Ring.whole_ideal[of R])
apply (frule_tac A = "carrier R" and H = H and a = "1rR" 
       and h = x in sc_linear_span, assumption+)
 apply (simp add:Ring.ring_one, assumption+)
 apply (frule_tac c = x in subsetD[of H "carrier M"], assumption+,
        simp add:sprod_one, assumption)
done

lemma (in Module) linear_span_inc_0:"⟦ideal R A; H ⊆ carrier M⟧  ⟹ 
                   𝟬 ∈ linear_span R M A H" 
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)

apply (frule nonempty_ex[of H], erule exE)
 apply (frule_tac h = x in sc_linear_span[of A H "𝟬R"], assumption)
 apply (cut_tac sc_Ring, simp add:Ring.ideal_zero, assumption)
 apply (frule_tac c = x in subsetD[of H "carrier M"], assumption,
        simp add:sc_0_m)
done

lemma (in Module) linear_span_iOp_closedTr1:"⟦ideal R A;
       s ∈ {j. j ≤ (n::nat)} → A⟧ ⟹
               (λx∈{j. j ≤ n}. -aR (s x)) ∈ {j. j ≤ n} → A"
apply (rule Pi_I)
 apply simp
 apply (cut_tac sc_Ring,
        rule Ring.ideal_inv1_closed, assumption+)
 apply (simp add:Pi_def)
done

lemma (in Module) l_span_gen_mono:"⟦K ⊆ H; H ⊆ carrier M; ideal R A⟧ ⟹
        linear_span R M A K ⊆ linear_span R M A H"
apply (rule subsetI)
apply (case_tac "K = {}", simp add:linear_span_def[of _ _ _ "{}"],
       simp add:linear_span_inc_0)
apply (frule nonempty_ex[of K], erule exE,
       frule_tac c = xa in subsetD[of K H], assumption+,
       frule nonempty[of _ H])
apply (simp add:linear_span_def[of _ _ _ K],
       erule exE, (erule bexE)+, simp,
       frule extend_fun[of _ _ K H], assumption+)
apply (simp add: l_comb_mem_linear_span)
done

lemma (in Module) l_comb_add:"⟦ideal R A; H ⊆ carrier M;
        s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ n} → H;
        t ∈ {j. j ≤ (m::nat)} → A; g ∈ {j. j ≤ m} → H⟧ ⟹
  l_comb R M (Suc (n + m)) (jointfun n s m t) (jointfun n f m g) =
                                  l_comb R M n s f ± l_comb R M m t g"
apply (cut_tac sc_Ring)       
apply (simp del:nsum_suc add:l_comb_def)
 apply (subst nsum_split)
 apply (rule allI, rule impI)
 apply (case_tac "j ≤ n", simp add:jointfun_def,
        rule sc_mem, simp add:Pi_def Ring.ideal_subset,
       simp add:Pi_def subsetD)
 apply (simp add:jointfun_def sliden_def) 
 apply (frule_tac m = j and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
        thin_tac "j ≤ Suc (n + m)", simp,
        rule sc_mem, simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD) 
 apply (simp add:l_comb_jointfun_jj[of H A s n f t m g])
 apply (cut_tac nsum_eq[of m "cmp (λj. jointfun n s m t j ⋅s 
        jointfun n f m g j) (slide (Suc n))" "λj. t j ⋅s g j"], simp)
 apply (rule allI, rule impI, simp add:cmp_def,
        simp add:jointfun_def sliden_def slide_def,
        rule sc_mem, simp add:Pi_def Ring.ideal_subset,
       simp add:Pi_def subsetD)
 apply (rule allI, rule impI,
        rule sc_mem, simp add:Pi_def Ring.ideal_subset,
       simp add:Pi_def subsetD)
 apply (simp add:cmp_def jointfun_def sliden_def slide_def)
done
       
lemma (in Module) l_comb_add1Tr:"⟦ideal R A; H ⊆ carrier M⟧ ⟹
  f ∈ {j. j ≤ (n::nat)} → H ∧ s ∈ {j. j ≤ n} → A ∧ t ∈ {j. j ≤ n} → A ⟶
    l_comb R M n (λx∈{j. j ≤ n}. (s x) ±R (t x)) f =
      l_comb R M n s f ± l_comb R M n t f"
apply (induct_tac n)
 apply (simp add:l_comb_def sc_Ring Ring.ideal_subset subsetD sc_l_distr)

 apply (rule impI, (erule conjE)+)
 apply (frule func_pre[of f], frule func_pre[of s], frule func_pre[of t],
        simp)
 apply (simp add:l_comb_def, cut_tac sc_Ring)
 apply (cut_tac n = n and f = "λj. (if j ≤ n then s j ±R t j else undefined) ⋅s f j" and g = "λj. (if j ≤ Suc n then s j ±R t j else undefined) ⋅s
                     f j" in nsum_eq)
 apply (rule allI, rule impI, simp,
         rule sc_mem, frule Ring.ring_is_ag,
         rule aGroup.ag_pOp_closed[of R], assumption,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def subsetD)
 apply (rule allI, rule impI, simp,
         rule sc_mem, frule Ring.ring_is_ag,
         rule aGroup.ag_pOp_closed[of R], assumption,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def subsetD)
 apply (simp)
 apply simp
 apply (thin_tac e M (λj. (if j ≤ n then s j ±R t j else undefined) ⋅s f j)
        n =  Σe M (λj. s j ⋅s f j) n ± Σe M (λj. t j ⋅s f j) n",
        thin_tac e M (λj. (if j ≤ Suc n then s j ±R t j else undefined) ⋅s 
        f j) n = Σe M (λj. s j ⋅s f j) n ± Σe M (λj. t j ⋅s f j) n")
 apply (frule_tac x = "Suc n" and A = "{j. j ≤ Suc n}" in 
        funcset_mem[of s _ A], simp,
        frule_tac x = "Suc n" and A = "{j. j ≤ Suc n}" in 
        funcset_mem[of t _ A], simp,
        frule_tac x = "Suc n" and A = "{j. j ≤ Suc n}" in
        funcset_mem[of f _ H], simp,
        cut_tac sc_Ring,
        frule_tac h = "s (Suc n)" in Ring.ideal_subset, assumption+,
        frule_tac h = "t (Suc n)" in Ring.ideal_subset, assumption+,
        frule_tac c = "f (Suc n)" in subsetD[of H "carrier M"], assumption+)
 apply (simp add:sc_l_distr)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" in nsum_mem,
        rule allI, rule impI,  rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (cut_tac n = n and f = "λj. t j ⋅s f j" in nsum_mem,
        rule allI, rule impI,  rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (cut_tac a = "s (Suc n)" and m = "f (Suc n)" in sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (cut_tac a = "t (Suc n)" and m = "f (Suc n)" in sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (subst pOp_assocTr41[THEN sym], assumption+,
        subst pOp_assocTr42, assumption+)
 apply (frule_tac x = e M (λj. t j ⋅s f j) n" and 
         y = "s (Suc n) ⋅s f (Suc n)" in ag_pOp_commute, assumption+, simp)
  apply (subst pOp_assocTr42[THEN sym], assumption+,
         subst pOp_assocTr41, assumption+, simp)
done

lemma (in Module) l_comb_add1:"⟦ideal R A; H ⊆ carrier M; 
 f ∈ {j. j ≤ (n::nat)} → H; s ∈ {j. j ≤ n} → A; t ∈ {j. j ≤ n} → A ⟧ ⟹ 
   l_comb R M n (λx∈{j. j ≤ n}. (s x) ±R (t x)) f =
                                l_comb R M n s f ± l_comb R M n t f"
apply (simp add:l_comb_add1Tr)
done

lemma (in Module) linear_span_iOp_closedTr2:"⟦ideal R A; H ⊆ carrier M; 
       f ∈ {j. j ≤ (n::nat)} → H; s ∈ {j. j ≤ n} → A⟧  ⟹
       -a (l_comb R M n s f) = 
           l_comb R M n (λx∈{j. j ≤ n}. -aR (s x)) f"
apply (frule_tac f = f and A = "{j. j ≤ n}" and B = H and x = 0 in 
       funcset_mem, simp)
apply (frule_tac A = A and s = s in linear_span_iOp_closedTr1, assumption+)
apply (frule l_comb_add1[of A H f n s "λx∈{j. j ≤ n}. -aR (s x)"], 
        assumption+)
apply (cut_tac linear_comb0_1[of H "λx∈{j. j ≤ n}. s x ±R 
                  (λx∈{j. j ≤ n}. -aR (s x)) x" n f])
 apply (simp,
       thin_tac "l_comb R M n
 (λx∈{j. j ≤ n}. s x ±R (if x ≤ n then -aR (s x) else undefined)) f = 𝟬")
 apply (frule l_comb_mem[of A H s n f], assumption+,
        frule l_comb_mem[of A H "λx∈{j. j ≤ n}. -aR (s x)" n f], assumption+)
 apply (frule ag_mOp_closed[of "l_comb R M n s f"])
 apply (frule ag_pOp_assoc[of "-a (l_comb R M n s f)" "l_comb R M n s f" "l_comb R M n (λx∈{j. j ≤ n}. -aR (s x)) f"], assumption+)
 apply (simp, simp add:ag_l_inv1, simp add:ag_l_zero, simp add:ag_r_zero)
 apply assumption+
 apply (rule Pi_I, simp)
 apply (frule_tac x = x in funcset_mem[of s "{j. j ≤ n}" A], simp,
        cut_tac sc_Ring,
        frule_tac h = "s x" in Ring.ideal_subset[of R A], assumption+)
 apply (frule Ring.ring_is_ag[of R],
        simp add:aGroup.ag_r_inv1[of R])
 apply assumption
done

lemma (in Module) linear_span_iOp_closed:"⟦ideal R A; H ⊆ carrier M; 
 a ∈ linear_span R M A H⟧ ⟹ -a a ∈ linear_span R M A H"
apply (case_tac "H = {}")
apply (simp add:linear_span_def)
apply (simp add:ag_inv_zero)
apply (simp add:linear_span_def, erule exE, (erule bexE)+)
apply simp
apply (frule_tac f = f and n = n and s = s in 
                 linear_span_iOp_closedTr2[of A H], assumption+)
apply (subgoal_tac "(λx∈{j. j ≤ n}. -aR (s x)) ∈ {j. j ≤ n} → A")
apply blast
apply (rule Pi_I, simp)
apply(cut_tac sc_Ring,
      rule Ring.ideal_inv1_closed, assumption+,
      simp add:Pi_def)
done

lemma (in Module) linear_span_pOp_closed:
 "⟦ideal R A; H ⊆ carrier M; a ∈ linear_span R M A H; b ∈ linear_span R M A H⟧
  ⟹ a ± b ∈ linear_span R M A H"
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)
 apply (cut_tac ag_inc_zero, simp add:ag_r_zero)
apply (simp add:linear_span_def) 
 apply ((erule exE)+, (erule bexE)+)
 apply (rename_tac n m f g s t)
 apply (simp add:l_comb_def)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" and m = m and 
                g = "λj. t j ⋅s g j" in nsum_add_nm)
 apply (rule allI, rule impI, rule sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI, rule sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rotate_tac -1, frule sym, 
        thin_tac e M (jointfun n (λj. s j ⋅s f j) m (λj. t j ⋅s g j)) 
                     (Suc (n + m)) =
                  Σe M (λj. s j ⋅s f j) n ± Σe M (λj. t j ⋅s g j) m",
         simp del:nsum_suc)
 apply (cut_tac n = "Suc (n + m)" and f = "jointfun n (λj. s j ⋅s f j) m 
  (λj. t j ⋅s g j)" and g = "λj. (jointfun n s m t) j ⋅s (jointfun n f m g) j"
   in nsum_eq)
 apply (rule allI, rule impI)
  apply (simp add:jointfun_def)
  apply (case_tac "j ≤ n", simp)
  apply (rule sc_mem,
         cut_tac sc_Ring,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def subsetD)  
  apply (simp, rule sc_mem)
  apply (simp add:sliden_def,
         frule_tac m = j and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
         thin_tac "j ≤ Suc (n + m)", simp,
         cut_tac sc_Ring,
         simp add:Pi_def Ring.ideal_subset) 
  apply (simp add:sliden_def,
         frule_tac m = j and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
         thin_tac "j ≤ Suc (n + m)", simp,
         cut_tac sc_Ring,
         simp add:Pi_def subsetD) 
 apply (rule allI, rule impI)
  apply (simp add:jointfun_def)
  apply (case_tac "j ≤ n", simp)
  apply (rule sc_mem,
         cut_tac sc_Ring,
         simp add:Pi_def Ring.ideal_subset,
         simp add:Pi_def subsetD)  
  apply (simp, simp add:sliden_def,
         rule sc_mem,
         frule_tac m = j and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
         thin_tac "j ≤ Suc (n + m)", simp,
         cut_tac sc_Ring,
         simp add:Pi_def Ring.ideal_subset) 
  apply (frule_tac m = j and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
         thin_tac "j ≤ Suc (n + m)", simp,
         cut_tac sc_Ring,
         simp add:Pi_def subsetD)
  apply (rule allI, rule impI,
         simp add:jointfun_def)
apply (simp del:nsum_suc,
       thin_tac e M (λj. s j ⋅s f j) n ± Σe M (λj. t j ⋅s g j) m =
        Σe M (λj. jointfun n s m t j ⋅s jointfun n f m g j) (Suc (n + m))",
       thin_tac e M (jointfun n (λj. s j ⋅s f j) m (λj. t j ⋅s g j))
                        (Suc (n + m)) =
        Σe M (λj. jointfun n s m t j ⋅s jointfun n f m g j) (Suc (n + m))")
 apply (frule_tac f = s and n = n and A = A and g = t and m = m and B = A in
                  jointfun_hom0, assumption+, simp del:nsum_suc,
        frule_tac f = f and n = n and A = H and g = g and m = m and B = H in
                  jointfun_hom0, assumption+, simp del:nsum_suc)
 apply blast
done

lemma (in Module) l_comb_scTr:"⟦ideal R A; H ⊆ carrier M;
 r ∈ carrier R; H ≠ {}⟧  ⟹ s ∈ {j. j ≤ (n::nat)} → A ∧ 
 g ∈ {j. j ≤ n} →  H  ⟶ r ⋅s (nsum M (λk. (s k) ⋅s (g k))  n) =  
                             nsum M (λk. r ⋅s ((s k) ⋅s (g k))) n" 
apply (induct_tac n)
 apply (rule impI, (erule conjE)+, simp)

apply (rule impI) apply (erule conjE)
 apply (frule func_pre [of _ _ "A"]) apply (frule func_pre [of _ _ "H"])
 apply (simp)
 apply (cut_tac n = n and f = "λk. s k ⋅s g k" in nsum_mem,
        rule allI, rule impI,
        cut_tac sc_Ring, rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)  
 apply (cut_tac a = "s (Suc n)" and m = "g (Suc n)" in sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)  
 apply (simp add:sc_r_distr)
done

lemma (in Module) l_comb_sc1Tr:"⟦ideal R A; H ⊆ carrier M;
 r ∈ carrier R; H ≠ {}⟧  ⟹ s ∈ {j. j ≤ (n::nat)} → A ∧ 
 g ∈ {j. j ≤ n} →  H  ⟶ r ⋅s (nsum M (λk. (s k) ⋅s (g k))  n) =  
                             nsum M (λk. (r ⋅rR (s k)) ⋅s (g k)) n"
apply (cut_tac sc_Ring) 
apply (induct_tac n)
 apply (rule impI, (erule conjE)+, simp)
 apply (subst sc_assoc, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD, simp)

apply (rule impI) apply (erule conjE)
 apply (frule func_pre [of _ _ "A"], frule func_pre [of _ _ "H"])
 apply simp
 apply (cut_tac n = n and f = "λk. s k ⋅s g k" in nsum_mem,
        rule allI, rule impI,
        cut_tac sc_Ring, rule sc_mem,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)  
 apply (cut_tac a = "s (Suc n)" and m = "g (Suc n)" in sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)  
 apply (simp add:sc_r_distr)
 apply (subst  sc_assoc, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD, simp)
done

lemma (in Module) l_comb_sc:"⟦ideal R A; H ⊆ carrier M; r ∈ carrier R; 
      s ∈ {j. j ≤ (n::nat)} → A;  g ∈ {j. j ≤ n} →  H⟧ ⟹
r ⋅s (nsum M (λk. (s k) ⋅s (g k)) n) = nsum M (λk. r ⋅s ((s k) ⋅s (g k))) n" 
apply (case_tac "H ≠ {}")
 apply (simp add:l_comb_scTr)
 apply simp
 apply (frule_tac x = 0 in funcset_mem[of g " {j. j ≤ n}" "{}"], simp)
 apply blast
done

lemma (in Module) l_comb_sc1:"⟦ideal R A; H ⊆ carrier M; r ∈ carrier R; 
      s ∈ {j. j ≤ (n::nat)} → A;  g ∈ {j. j ≤ n} →  H⟧ ⟹
r ⋅s (nsum M (λk. (s k) ⋅s (g k)) n) = nsum M (λk. (r ⋅rR (s k)) ⋅s (g k)) n" 
apply (case_tac "H ≠ {}")
 apply (simp add:l_comb_sc1Tr)
 apply simp
 apply (frule_tac x = 0 in funcset_mem[of g " {j. j ≤ n}" "{}"], simp)
 apply blast
done

lemma (in Module) linear_span_sc_closed:"⟦ideal R A; H ⊆ carrier M;
 r ∈ carrier R; x ∈ linear_span R M A H⟧ ⟹ r ⋅s x ∈ linear_span R M A H"
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)
 apply (simp add:sc_a_0)
apply (simp add:linear_span_def)
 apply (erule exE, (erule bexE)+)
 apply (simp add:l_comb_def) 
 apply (simp add:l_comb_sc)
 
apply (cut_tac n = n and f = "λj. r ⋅s (s j ⋅s f j)" and 
       g = "λj. (r ⋅rR (s j)) ⋅s f j" in nsum_eq)
 apply (rule allI, rule impI,
        rule sc_mem, assumption, rule sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI,
        rule sc_mem,
        cut_tac sc_Ring,
        rule Ring.ring_tOp_closed, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule allI, rule impI,
        subst sc_assoc, assumption,
        cut_tac sc_Ring, 
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD, simp,
        simp,
    thin_tac e M (λj. r ⋅s (s j ⋅s f j)) n = 
                                Σe M (λj. (r ⋅rR s j) ⋅s f j) n",
    thin_tac "x = Σe M (λj. s j ⋅s f j) n")

 apply (cut_tac n = n and f = "λj. (r ⋅rR s j) ⋅s f j" and 
       g = "λj. (λx∈{j. j ≤ n}. r ⋅rR (s x)) j ⋅s f j" in nsum_eq)
  apply (rule allI, rule impI,
        rule sc_mem,
        cut_tac sc_Ring,
        rule Ring.ring_tOp_closed, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
   apply (rule allI, rule impI,
         rule sc_mem, simp) apply (
          cut_tac sc_Ring,
        rule Ring.ring_tOp_closed, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
  apply (rule allI, rule impI)
         apply simp
  apply (subgoal_tac "(λx∈{j. j ≤ n}. r ⋅rR s x) ∈ {j. j ≤ n} → A",
         blast)
  apply (rule Pi_I, simp)
apply (thin_tac e M (λj. (r ⋅rR s j) ⋅s f j) n =
        Σe M (λj. (if j ≤ n then r ⋅rR s j else undefined) ⋅s f j) n",
        cut_tac sc_Ring,
        rule Ring.ideal_ring_multiple, assumption+, simp add:Pi_def,
        assumption)
done
    
lemma (in Module) mem_single_l_spanTr:"⟦ideal R A; h ∈ carrier M⟧ ⟹
      s ∈ {j. j ≤ (n::nat)} → A ∧
      f ∈ {j. j ≤ n} → {h} ∧ l_comb R M n s f ∈ linear_span R M A {h}
      ⟶ (∃a ∈ A. l_comb R M n s f = a ⋅s h)"
apply (cut_tac sc_Ring)  
apply (induct_tac n)
 apply (rule impI, (erule conjE)+)
 apply (simp add:l_comb_def Ring.ideal_subset[of R A] bexI[of _ "s 0"])
apply (rule impI, (erule conjE)+,
       frule func_pre[of _ _ A], frule func_pre[of _ _ "{h}"],
       frule_tac n = n in l_comb_mem_linear_span[of A "{h}" s _ f],
       rule subsetI, simp, assumption+, simp,
       erule bexE)
apply (frule singleton_sub[of h "carrier M"])
 apply (frule Ring.ideal_subset1[of R A], assumption)
 apply (frule extend_fun[of s _ A "carrier R"], assumption)
 apply (frule_tac n = n in l_comb_Suc[of "{h}" A s _ f], assumption+,
        simp)
 apply (frule_tac A = "{j. j ≤ Suc n}" and x = "Suc n" in 
        funcset_mem[of f _ "{h}"], simp, simp,
        frule_tac A = "{j. j ≤ Suc n}" and x = "Suc n" in 
        funcset_mem[of s _ A], simp,
        frule_tac h = "s (Suc n)" in Ring.ideal_subset[of R A], assumption+)
 apply (frule_tac h = a in Ring.ideal_subset[of R A], assumption+,
        frule_tac h = "s (Suc n)" in Ring.ideal_subset[of R A], assumption+,
        simp add:sc_l_distr[THEN sym],
        frule_tac x = a and y = "s (Suc n)" in Ring.ideal_pOp_closed[of R A],
        assumption+, blast)
done

lemma (in Module) mem_single_l_span:"⟦ideal R A; h ∈ carrier M; 
       s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ n} → {h}; 
       l_comb R M n s f ∈ linear_span R M A {h}⟧ ⟹
       ∃a ∈ A. l_comb R M n s f = a ⋅s h"
apply (simp add:mem_single_l_spanTr)
done

lemma (in Module) mem_single_l_span1:"⟦ideal R A; h ∈ carrier M; 
       x ∈ linear_span R M A {h}⟧ ⟹ ∃a ∈ A. x = a ⋅s h"
apply (simp add:linear_span_def, erule exE, (erule bexE)+, simp)
apply (frule_tac s = s and n = n and f = f in mem_single_l_span[of A h],
       assumption+)
apply (frule singleton_sub[of h "carrier M"],
      rule_tac s = s and f = f in l_comb_mem_linear_span[of A "{h}"],
      assumption+)
done

lemma (in Module) linear_span_subModule:"⟦ideal R A; H ⊆ carrier M⟧  ⟹ 
                  submodule R M (linear_span R M A H)"
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)
 apply (simp add:submodule_0)

apply (simp add:submodule_def)
apply (rule conjI)
 apply (simp add:linear_span_def)
 apply (rule subsetI)
 apply (simp add:CollectI)
 apply (erule exE, (erule bexE)+)
 apply simp
 apply (simp add:l_comb_mem)
apply (rule conjI)
 apply (rule asubg_test) 
 apply (rule subsetI) apply (simp add:linear_span_def)
 apply (erule exE, (erule bexE)+)
 apply (simp add:l_comb_mem)
 apply (frule linear_span_inc_0[of A H], assumption, blast)
 apply (rule ballI)+
 apply (rule linear_span_pOp_closed, assumption+)
 apply (rule linear_span_iOp_closed, assumption+)
apply (rule allI)+
 apply (simp add:linear_span_sc_closed)
done

lemma (in Module) l_comb_mem_submoduleTr:"⟦ideal R A; submodule R M N⟧ ⟹
 (s ∈ {j. j ≤ (n::nat)} → A ∧ f ∈ {j. j ≤ n} → carrier M ∧
 (∀j ≤ n.(s j) ⋅s (f j) ∈ N)) ⟶ l_comb R M n s f ∈ N"
apply (induct_tac n)
 apply (simp add:l_comb_def, rule impI, (erule conjE)+)
apply (frule func_pre[of _ _ A], frule func_pre[of _ _ "carrier M"], simp)
apply (simp add:l_comb_def)
apply (frule_tac a = "Suc n" in forall_spec, simp) 
apply (rule submodule_pOp_closed, assumption+)
done

lemma (in Module) l_span_sub_submodule:"⟦ideal R A; submodule R M N; H ⊆ N⟧ ⟹
       linear_span R M A H ⊆ N"
apply (cut_tac sc_Ring)
 apply (rule subsetI, simp add:linear_span_def)
 apply (case_tac "H = {}", simp)
 apply (simp add:submodule_inc_0)

 apply simp
 apply (erule exE, (erule bexE)+)
 apply (cut_tac s = s and A = A and f = f and N = N and n = n in 
        l_comb_mem_submoduleTr, assumption+,
        frule submodule_subset[of N],
        frule subset_trans[of H N "carrier M"], assumption+,
        frule_tac f = f and A = "{j. j ≤ n}" and B = H and ?B1.0 = "carrier M"
        in extend_fun, assumption+)
 apply (subgoal_tac "∀j≤n. s j ⋅s f j ∈ N", simp)
 apply (rule allI, rule impI)
 apply (rule submodule_sc_closed[of N], assumption,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
done

lemma (in Module) linear_span_sub:"⟦ideal R A; H ⊆ carrier M⟧  ⟹ 
                  (linear_span R M A H) ⊆ carrier M"
apply (frule linear_span_subModule[of A H], assumption+)
apply (simp add:submodule_subset)
done

definition
  smodule_ideal_coeff :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
       'r set] ⇒ 'a set" where
  "smodule_ideal_coeff R M A = linear_span R M A (carrier M)"

abbreviation
  SMLIDEALCOEFF  ("(3_/ ⊙_ _)" [64,64,65]64) where
  "A ⊙R M == smodule_ideal_coeff R M A"

lemma (in Module) smodule_ideal_coeff_is_Submodule:"ideal R A  ⟹
            submodule R M (A ⊙R M)"
apply (simp add:smodule_ideal_coeff_def)
apply (simp add:linear_span_subModule)
done

lemma (in Module) mem_smodule_ideal_coeff:"⟦ideal R A; x ∈ A ⊙R M⟧ ⟹
             ∃n. ∃s ∈ {j. j ≤ n} → A. ∃g ∈ {j. j ≤ n} → carrier M.
              x = l_comb R M n s g" 
apply (cut_tac ag_inc_zero,
       frule nonempty[of "𝟬" "carrier M"])
apply (simp add:smodule_ideal_coeff_def linear_span_def,
       erule exE, (erule bexE)+, blast)
done

definition
  quotient_of_submodules :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
            'a set, 'a set] ⇒ 'r set" where
  "quotient_of_submodules R M N P = {x | x. x∈carrier R ∧ 
                                    (linear_span R M (Rxa R x)  P) ⊆ N}"

definition
  Annihilator :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme]
    ⇒ 'r set" ("(Ann_ _)" [82,83]82) where
  "AnnR M = quotient_of_submodules R M {𝟬M} (carrier M)"

abbreviation
  QOFSUBMDS  ("(4_ _‡_ _)" [82,82,82,83]82) where
  "N R‡M P == quotient_of_submodules R M N P"

lemma (in Module) quotient_of_submodules_inc_0:
     "⟦submodule R M P; submodule R M Q⟧ ⟹ 𝟬R ∈ (P R‡M Q)"
apply (simp add:quotient_of_submodules_def)
apply (cut_tac sc_Ring, simp add:Ring.ring_zero)
apply (simp add:linear_span_def)
 apply (frule submodule_inc_0[of Q], simp add:nonempty)
apply (rule subsetI)
 apply (simp, erule exE, (erule bexE)+)
 apply (simp, thin_tac "x = l_comb R M n s f", simp add:l_comb_def)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" in nsum_zeroA)
 apply (rule allI, rule impI,
       frule_tac x = j and f = s and A = "{j. j ≤ n}" in 
       funcset_mem[of _ _ "R ♢p 𝟬R"], simp)
 apply (simp add:Rxa_def, erule bexE, simp) apply (
        simp add:Ring.ring_times_x_0,
        rule sc_0_m) apply (
        frule submodule_subset[of Q],
        simp add:Pi_def subsetD)
 apply (simp add:submodule_inc_0)
done
 
lemma (in Module) quotient_of_submodules_is_ideal:
      "⟦submodule R M P; submodule R M Q⟧ ⟹ ideal R (P R‡M Q)"
apply (frule quotient_of_submodules_inc_0 [of P Q], assumption+)
apply (cut_tac sc_Ring,
       rule Ring.ideal_condition[of R], assumption+)
apply (simp add:quotient_of_submodules_def)
 apply (rule subsetI)
 apply (simp add:CollectI)
apply (simp add:nonempty) apply (thin_tac "𝟬R ∈ P R‡M Q")
 apply (rule ballI)+
 apply (simp add:quotient_of_submodules_def) 
apply (erule conjE)+
 apply (rule conjI)
 apply (frule Ring.ring_is_ag,
        rule aGroup.ag_pOp_closed[of R], assumption+)
 apply (rule aGroup.ag_mOp_closed, assumption+)
apply (subst linear_span_def)
 apply (frule submodule_inc_0 [of Q], simp add:nonempty)
 apply (rule subsetI, simp,
        erule exE, (erule bexE)+, simp add:l_comb_def,
        thin_tac "xa = Σe M (λj. s j ⋅s f j) n")
 apply (cut_tac s = s and n = n and f = f in 
           l_comb_mem_submoduleTr[of "carrier R" P])
 apply (simp add:Ring.whole_ideal, assumption+)
 apply (frule Ring.ring_is_ag[of R],
        frule_tac x = y in aGroup.ag_mOp_closed[of R], assumption+,
        frule_tac x = x and y = "-aR y" in aGroup.ag_pOp_closed, assumption+,
        frule_tac a = "x ±R -aR y" in Ring.principal_ideal[of R], assumption+,
        frule_tac I = "R ♢p (x ±R -aR y)" in Ring.ideal_subset1, assumption+)
  apply (frule_tac f = s and A = "{j. j ≤ n}" and B = "R ♢p (x ±R -aR y)" 
         and ?B1.0 = "carrier R" in extend_fun, assumption+,
         frule_tac submodule_subset[of Q],
         frule_tac f = f and A = "{j. j ≤ n}" and B = Q  
         and ?B1.0 = "carrier M" in extend_fun, assumption+)        
  apply (subgoal_tac "∀j≤n. s j ⋅s f j ∈ P", simp add:l_comb_def,
         thin_tac "s ∈ {j. j ≤ n} → carrier R ∧
        f ∈ {j. j ≤ n} → carrier M ∧ (∀j≤n. s j ⋅s f j ∈ P) ⟶
        l_comb R M n s f ∈ P",
         thin_tac "s ∈ {j. j ≤ n} → carrier R",
         thin_tac "f ∈ {j. j ≤ n} → carrier M")
  apply (rule allI, rule impI,
         frule_tac x = j and f = s and A = "{j. j ≤ n}" and 
         B = "R ♢p (x ±R -aR y)" in funcset_mem, simp, 
         thin_tac "s ∈ {j. j ≤ n} → R ♢p (x ±R -aR y)",
         thin_tac "ideal R (R ♢p (x ±R -aR y))")
  apply (simp add:Rxa_def, fold Rxa_def, erule bexE, simp,
         thin_tac "s j = r ⋅rR (x ±R -aR y)")
  apply (simp add:Ring.ring_distrib1,
         frule_tac x = r and y = x in Ring.ring_tOp_closed, assumption+, 
         frule_tac x = r and y = "-aR y" in Ring.ring_tOp_closed, assumption+,
         frule_tac x = j and A = "{j. j ≤ n}" and B = Q in funcset_mem,
         simp,
         frule_tac c = "f j" in subsetD[of Q "carrier M"], assumption+,
         simp add:sc_l_distr)
  apply (subst Ring.ring_inv1_2[THEN sym], assumption+,
         subst Ring.ring_inv1_1, assumption+)
  apply (frule_tac a = x in Ring.principal_ideal[of R], assumption+,
         frule_tac a = x in Ring.principal_ideal[of R], assumption+,
         frule_tac A = "R ♢p x" and H = Q and a = "r ⋅rR x" and h = "f j" in
         sc_linear_span, assumption+, simp add:Rxa_def, blast,
         simp add:Pi_def)
  apply (frule_tac x = r in aGroup.ag_mOp_closed[of R], assumption+,
         frule_tac a = y in Ring.principal_ideal[of R], assumption+,
         frule_tac a = y in Ring.principal_ideal[of R], assumption+,
         frule_tac A = "R ♢p y" and H = Q and a = "(-aR r) ⋅rR y" and
         h = "f j" in sc_linear_span, assumption+, simp add:Rxa_def,
         blast,
         simp add:Pi_def)
  apply (frule_tac c = "(r ⋅rR x) ⋅s f j" and A = "linear_span R M (R ♢p x) Q" 
         and B = P in subsetD, assumption+) apply (
         frule_tac c = "((-aR r) ⋅rR y) ⋅s f j" and 
         A = "linear_span R M (R ♢p y) Q" and B = P in subsetD, assumption+)
  apply (rule submodule_pOp_closed, assumption+)

  apply ((rule ballI)+,
         thin_tac "𝟬R ∈ P R‡M Q",
         simp add:quotient_of_submodules_def, erule conjE)
  apply (simp add:Ring.ring_tOp_closed)
  apply (rule subsetI)
  apply (frule submodule_inc_0[of Q],
         simp add:linear_span_def nonempty)
  apply (erule exE, (erule bexE)+)
  apply (rule_tac c = xa and A = "{xa. ∃n. ∃f∈{j. j ≤ n} → Q.
                ∃s∈{j. j ≤ n} → R ♢p x. xa = l_comb R M n s f}" in
          subsetD[of _ P], assumption+,
          thin_tac "{xa. ∃n. ∃f∈{j. j ≤ n} → Q.
                ∃s∈{j. j ≤ n} → R ♢p x. xa = l_comb R M n s f} ⊆ P")
  apply simp
  apply (frule_tac a = r and b = x in Ring.Rxa_mult_smaller[of R], assumption+)
  apply (frule_tac f = s and A = "{j. j ≤ n}" and B = "R ♢p (r ⋅rR x)" and
          ?B1.0 = "R ♢p x" in extend_fun, assumption+)
  apply blast
done
 
lemma (in Module) Ann_is_ideal:"ideal R (AnnR M)"
apply (simp add:Annihilator_def)
apply (rule quotient_of_submodules_is_ideal)
apply (simp add:submodule_0)
apply (simp add:submodule_whole)
done

lemma (in Module) linmap_im_of_lincombTr:"⟦ideal R A; R module N; 
      f ∈ mHom R M N; H ⊆ carrier M⟧ ⟹  
      s ∈ {j. j ≤ (n::nat)} → A ∧ g ∈ {j. j ≤ n} → H ⟶
      f (l_comb R M n s g) = l_comb R N n s (cmp f g)"
apply (induct_tac n)
 apply (rule impI) apply (erule conjE)
 apply (simp add:l_comb_def)
 apply (cut_tac m = "g 0" and f = f and a = "s 0" in mHom_lin [of N],
        assumption+,
        simp add:Pi_def subsetD, assumption,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset, simp add:cmp_def)

apply (rule impI, erule conjE)
 apply (frule_tac f = s in func_pre,
        frule_tac f = g in func_pre, simp)
 apply (simp add:l_comb_def)
 apply (subst mHom_add[of N f], assumption+)
 apply (rule nsum_mem,
        rule allI, rule impI, rule sc_mem,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
 apply (rule sc_mem,
         cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD, simp,
        frule_tac x = "Suc n" and A = "{j. j ≤ Suc n}" and f = s and 
                  B = A in funcset_mem, simp,
        cut_tac sc_Ring,
        frule_tac h = "s (Suc n)" in Ring.ideal_subset[of R A], assumption+,
        frule_tac x = "Suc n" and A = "{j. j ≤ Suc n}" and f = g and 
                  B = H in funcset_mem, simp,
        frule_tac c = "g (Suc n)" in subsetD[of H "carrier M"], assumption+)
 apply (simp add:mHom_lin cmp_def)
done
 
lemma (in Module) linmap_im_lincomb:"⟦ideal R A; R module N; f ∈ mHom R M N; 
      H ⊆ carrier M; s ∈ {j. j ≤ (n::nat)} → A; g ∈ {j. j ≤ n} → H ⟧ ⟹ 
      f (l_comb R M n s g) = l_comb R N n s (cmp f g)"
apply (simp add:linmap_im_of_lincombTr)
done

lemma (in Module) linmap_im_linspan:"⟦ideal R A; R module N; f ∈ mHom R M N; 
       H ⊆ carrier M; s ∈ {j. j ≤ (n::nat)} → A; g ∈ {j. j ≤ n} → H ⟧ ⟹ 
            f (l_comb R M n s g) ∈ linear_span R N A (f ` H)"
apply (frule l_comb_mem_linear_span[of A H s n g], assumption+) 
 apply (simp add:linmap_im_lincomb)
 apply (rule Module.l_comb_mem_linear_span[of N R A "f ` H" s n "cmp f g"],
        assumption+,
        rule subsetI,
        simp add:image_def, erule bexE, simp,
        frule_tac c = xa in subsetD[of H "carrier M"], assumption+,
        simp add:mHom_mem[of N f], assumption+)
 apply (rule Pi_I, simp add:cmp_def)
 apply (frule_tac f = g and A = "{j. j ≤ n}" and B = H and x = x in 
        funcset_mem, simp, simp add:image_def) 
 apply blast
done

lemma (in Module) linmap_im_linspan1:"⟦ideal R A; R module N; f ∈ mHom R M N; 
      H ⊆ carrier M; h ∈ linear_span R M A H⟧ ⟹ 
                              f h ∈ linear_span R N A (f ` H)"
apply (simp add:linear_span_def [of "R" "M"])
 apply (case_tac "H = {}", simp add:linear_span_def)
 apply (simp add:mHom_0, simp)
apply (erule exE, (erule bexE)+)
 apply (simp add:linmap_im_linspan)
done

(*
section "A module over two rings"

record ('a, 'r, 's) bModule = "'a aGroup" +
  sc_l  :: "'r ⇒ 'a ⇒ 'a"    (infixl "⋅slı" 70)
  sc_r  :: "'a ⇒ 's ⇒ 'a"    (infixl "⋅srı" 70)

locale bModule = aGroup M +
  fixes R (structure)
  fixes S (structure)
  assumes  scl_Ring: "Ring R"
  and      scr_Ring: "Ring S" 
  and  scl_closed :
      "⟦ a ∈ carrier R; m ∈ carrier M⟧ ⟹ a ⋅sl m ∈ carrier M"
  and scr_closed :
      "⟦ b ∈ carrier S; m ∈ carrier M⟧ ⟹ m ⋅sr b ∈ carrier M" 
  and scl_l_distr:
      "⟦a ∈ carrier R; b ∈ carrier R; m ∈ carrier M⟧ ⟹
       (a ±R b) ⋅sl m = a ⋅sl m ± b ⋅sl m"
  and scr_l_distr:
      "⟦a ∈ carrier S; m ∈ carrier M; n ∈ carrier M ⟧ ⟹
        (m ± n) ⋅sr a = m ⋅sr a ±  n ⋅sr a"
  and scl_r_distr:
      "⟦ a ∈ carrier R; m ∈ carrier M; n ∈ carrier M ⟧ ⟹
      a ⋅sl (m ± n) = a ⋅sl m ± a ⋅sl n"
  and scr_r_distr:
        "⟦a ∈ carrier S; b ∈ carrier S; m ∈ carrier M⟧ ⟹
          m ⋅sr (a ±S b) = m ⋅sr a ±  m ⋅sr b"
  and scl_assoc:
      "⟦ a ∈ carrier R; b ∈ carrier R; m ∈ carrier M ⟧ ⟹
      (a ⋅rR b) ⋅sl m = a ⋅sl (b ⋅sl m)"
  and scr_assoc:
      "⟦a ∈ carrier S; b ∈ carrier S; m ∈ carrier M ⟧ ⟹
       m ⋅sr (a ⋅rS b)  =  (m ⋅sr a) ⋅sr b"
  and scl_one:
      "m ∈ carrier M ⟹ (1rR) ⋅sl m = m" 
  and scr_one:
       "m ∈ carrier M ⟹ m ⋅sr (1rS) = m" 

definition lModule :: "('a, 'r, 's, 'more) bModule_scheme ⇒ ('a, 'r) Module" where
       ("(_l)" [1000]999)
  "Ml == ⦇carrier = carrier M, pop = pop M, mop = mop M, 
    zero = zero M, sprod = sc_l M ⦈"

definition scr_re :: "('a, 'b, 'c, 'more) bModule_scheme ⇒ 'c ⇒ 'a ⇒ 'a" where 
                  
 "scr_re M r m == sc_r M m r"

definition rModule :: "('a, 'r, 's, 'more) bModule_scheme ⇒ ('a, 's) Module" where
        ("(_r)" [1000]999) 
  "Mr == ⦇carrier = carrier M, pop = pop M, mop = mop M, 
    zero = zero M, sprod = scr_re M ⦈"

lemma (in bModule) bmodule_is_ag:"aGroup M"  
apply assumption
done

lemma (in bModule) lModule_is_Module:"R module Ml"
apply (subgoal_tac "aGroup M")
apply (rule Module.intro)
 apply (rule aGroup.intro)
 apply (simp add:lModule_def, simp add:aGroup.pop_closed[of M])
 apply (simp add:lModule_def, simp add:aGroup.ag_pOp_assoc)
 apply (simp add:lModule_def, simp add:aGroup.ag_pOp_commute)
 apply (simp add:lModule_def, rule mop_closed)
 apply (simp add:lModule_def, rule l_m, assumption+)
 apply (simp add:lModule_def, rule ex_zero)
 apply (simp add:lModule_def, rule l_zero, assumption)
apply (rule Module_axioms.intro)
 apply (simp add:scl_Ring)
 apply (simp add:lModule_def, rule  scl_closed, assumption+)
 apply (simp add:lModule_def, rule  scl_l_distr, assumption+)
 apply (simp add:lModule_def, rule  scl_r_distr, assumption+)
 apply (simp add:lModule_def, rule  scl_assoc, assumption+)
 apply (simp add:lModule_def, rule scl_one, assumption+)
done


lemma (in bModule) rModule_is_Module:"S module Mr"
apply (subgoal_tac "aGroup M")
apply (rule Module.intro)
 apply (rule aGroup.intro)
 apply (simp add:rModule_def, simp add:aGroup.pop_closed[of M])
 apply (simp add:rModule_def, simp add:aGroup.ag_pOp_assoc)
 apply (simp add:rModule_def, simp add:aGroup.ag_pOp_commute)
 apply (simp add:rModule_def, rule mop_closed)
 apply (simp add:rModule_def, rule l_m, assumption+)
 apply (simp add:rModule_def, rule ex_zero)
 apply (simp add:rModule_def, rule l_zero, assumption)
apply (rule Module_axioms.intro,
       simp add:scr_Ring)
apply (simp add:rModule_def, simp add:scr_re_def scr_closed)
apply (simp add:rModule_def, simp add:scr_re_def, simp add:scr_r_distr)
apply (simp add:rModule_def, simp add:scr_re_def, rule scr_l_distr, 
        assumption+)
apply (simp add:rModule_def scr_re_def,
       subst scr_assoc[THEN sym], assumption+,
       cut_tac scr_Ring,
       simp add:Ring.ring_tOp_commute)
apply (simp add:rModule_def scr_re_def) 
apply (cut_tac m = m in scr_one, simp)
apply assumption+
done

lemma (in Module) sprodr_welldefTr1:"⟦ideal R A; A ⊆ AnnR M; a ∈ A;
       m ∈ carrier M⟧  ⟹ a ⋅s m = 𝟬" 
apply (simp add:Annihilator_def quotient_of_submodules_def)
apply (frule subsetD, assumption+)
 apply (simp add:CollectI, erule conjE, 
        thin_tac "A ⊆ {u ∈ carrier R.
                   linear_span R M (R ♢p u) (carrier M) ⊆ {𝟬}}")
 apply (cut_tac sc_Ring,
        cut_tac a = a and A = "Rxa R a" in 
                         sc_linear_span[of  _ "carrier M" _ "m"],
                simp add:Ring.principal_ideal, simp, 
                simp add:Ring.a_in_principal, assumption)
 apply (frule subsetD[of "linear_span R M (R ♢p a) (carrier M)" "{𝟬}"
                             "a ⋅s m"], assumption)
 apply simp
done

lemma (in Module) sprodr_welldefTr2:"⟦ideal R A; A ⊆ AnnR M; a ∈ carrier R; 
      x ∈ a ⊎R A; m ∈ carrier M⟧  ⟹ a ⋅s m = x ⋅s m"
apply (cut_tac sc_Ring,
       frule Ring.mem_ar_coset1 [of R A a x], assumption+, erule bexE,
       rotate_tac -1, frule sym, thin_tac "h ±R a = x", simp)
apply (subst sc_l_distr)
 apply (simp add:Ring.ideal_subset, assumption+)
apply (simp add:sprodr_welldefTr1)
apply (frule sc_mem [of a m], assumption+)
apply (simp add:ag_l_zero)
done

definition cos_scr :: "[('r, 'm) Ring_scheme, 'r set, ('a, 'r, 'm1) Module_scheme] ⇒
               'a ⇒ 'r set ⇒ 'a" where
  "cos_scr R A M == λm. λX. (SOME x. x ∈ X) ⋅sM m"

lemma (in Module) cos_scr_welldef:"⟦ideal R A; A ⊆ AnnR M; a ∈ carrier R; 
       X = a ⊎R A; m ∈ carrier M⟧  ⟹ cos_scr R A M m X = a ⋅s m" 
apply (cut_tac sc_Ring,
       frule Ring.a_in_ar_coset [of R A a], assumption+)
 apply (simp add:cos_scr_def,
        rule sprodr_welldefTr2[THEN sym], assumption+) 
 prefer 2 apply simp
apply (rule someI2_ex, blast, assumption)
done

definition r_qr_bmod :: "[('r, 'm) Ring_scheme, 'r set, ('a, 'r, 'm1) Module_scheme] ⇒ 
    ('a, 'r, 'r set) bModule" where 
 "r_qr_bmod R A M == ⦇carrier = carrier M, pop = pop M, mop = mop M, 
  zero = zero M, sc_l = sprod M, sc_r = cos_scr R A M ⦈" *)
 (* Remark. A should be an ideal contained in AnnR M. *)

definition
  faithful :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme]
                             ⇒ bool" where
  "faithful R M ⟷ AnnR M = {𝟬R}"

section "nsum and Generators"

definition
  generator :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
               'a set] ⇒ bool" where
 "generator R M H == H ⊆ carrier M ∧ 
                      linear_span R M (carrier R) H = carrier M"

definition
  finite_generator :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
               'a set] ⇒ bool" where
  "finite_generator R M H ⟷ finite H ∧ generator R M H"

definition
  fGOver :: "[('a, 'r, 'm1) Module_scheme, ('r, 'm) Ring_scheme]  ⇒  bool"
              (*(infixl 70)*)  where
  "fGOver M R ⟷ (∃H. finite_generator R M H)"

abbreviation
  FGENOVER  (infixl "fgover" 70) where
  "M fgover R == fGOver M R"

lemma (in Module) h_in_linear_span:"⟦H ⊆ carrier M; h ∈ H⟧ ⟹
                                   h ∈ linear_span R M (carrier R) H"
apply (subst sprod_one [THEN sym, of h])
 apply (simp add:subsetD)
 apply (cut_tac sc_Ring)
 apply (frule Ring.ring_one)
 apply (rule sc_linear_span [of "carrier R" "H" "1rR" "h"])
 apply (simp add:Ring.whole_ideal) apply assumption+
done                                                   

lemma (in Module) generator_sub_carrier:"generator R M H ⟹
                                              H ⊆ carrier M" 
apply (simp add:generator_def)
done 

lemma (in Module) lin_span_sub_carrier:"⟦ideal R A; 
       H ⊆ carrier M⟧ ⟹ linear_span R M A H ⊆ carrier M"
apply (cut_tac sc_Ring)
apply (rule subsetI)
 apply (simp add:linear_span_def)
 apply (case_tac "H = {}") apply simp
 apply (simp add:module_inc_zero) 
apply simp
apply (erule exE, (erule bexE)+, simp,
       thin_tac "x = l_comb R M n s f")
apply (simp add:l_comb_def) 
apply (rule_tac n = n in nsum_mem) 
 apply (rule allI, rule impI)
 apply (rule sc_mem)
 apply (simp add:Pi_def Ring.ideal_subset)
 apply (simp add:Pi_def subsetD)
done

lemma (in Module) lin_span_coeff_mono:"⟦ideal R A; H ⊆ carrier M⟧⟹  
                        linear_span R M A H ⊆ linear_span R M (carrier R) H"
apply (cut_tac sc_Ring)
apply (rule subsetI)
 apply (simp add:linear_span_def)
 apply (case_tac "H = {}") apply simp apply simp
 apply (erule exE, (erule bexE)+)
 apply (frule Ring.ideal_subset1 [of R A], assumption+)
apply (frule_tac  f = s in extend_fun, assumption+) 
 apply blast
done

lemma (in Module) l_span_sum_closedTr:"⟦ideal R A; H ⊆ carrier M⟧⟹ 
   ∀s. ∀f. s∈{j. j ≤ (n::nat)} → A ∧ 
   f ∈ {j. j ≤ n} → linear_span R M A H ⟶
   (nsum M (λj. s j ⋅s (f j)) n ∈ linear_span R M A H)"
apply (cut_tac sc_Ring)
apply (induct_tac n)
 apply ((rule allI)+, rule impI, simp) 
 apply (erule conjE)
 apply (rule linear_span_sc_closed, assumption+)
 apply (simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def)

apply ((rule allI)+, rule impI, erule conjE)
 apply (frule func_pre [of _ _ "A"],
        frule func_pre [of _ _ "linear_span R M A H"])
 apply (drule_tac x = s in spec,
        drule_tac x = f in spec)

 apply simp
 apply (rule linear_span_pOp_closed, assumption+)
 apply (rule linear_span_sc_closed, assumption+,
        simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def subsetD)
done

lemma (in Module) l_span_closed:"⟦ideal R A; H ⊆ carrier M; 
 s ∈ {j. j ≤ (n::nat)} → A;  f ∈ {j. j ≤ n} → linear_span R M A H ⟧ ⟹
 l_comb R M n s f ∈ linear_span R M A H"
apply (simp add:l_comb_def)
apply (simp add: l_span_sum_closedTr)
done 

lemma (in Module) l_span_closed1:"⟦H ⊆ carrier M; 
      s ∈ {j. j ≤ (n::nat)} → carrier R;  
      f ∈ {j. j ≤ n} → linear_span R M (carrier R) H ⟧ ⟹
      Σe M (λj.  s j ⋅s (f j)) n ∈ linear_span R M (carrier R) H"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal [of "R"])
apply (frule l_span_sum_closedTr[of "carrier R" H n], assumption+)
apply (drule_tac x = s in spec,
       drule_tac x = f in spec,
       simp)
done

lemma (in Module) l_span_closed2Tr0:"⟦ideal R A; H ⊆ carrier M; Ring R; s ∈ A;
     f ∈ linear_span R M (carrier R) H ⟧ ⟹ s ⋅s f ∈ linear_span R M A H"
apply (cut_tac sc_Ring)
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)
 apply (rule sc_a_0,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset) 

 apply (simp add:linear_span_def) 
 apply (erule exE, (erule bexE)+, simp,
        thin_tac "f = l_comb R M n sa fa")
 apply (frule Ring.whole_ideal[of R])
 apply (frule_tac h = s in Ring.ideal_subset[of R A], assumption+)
 apply (frule_tac s = sa and g = f in l_comb_sc1[of "carrier R" H s],
        assumption+, simp add:l_comb_def,
        thin_tac "s ⋅s Σe M (λk. sa k ⋅s f k) n = 
                                    Σe M (λk. (s ⋅rR sa k) ⋅s f k) n")
 apply (cut_tac n = n and f = "λj. (s ⋅rR sa j) ⋅s f j" and 
        g = "λj. ((λx∈{j. j ≤ n}. (s ⋅rR sa x)) j) ⋅s f j" in nsum_eq)
        apply (rule allI, rule impI, rule sc_mem,
               rule Ring.ring_tOp_closed, assumption+,
               simp add:Pi_def,
               simp add:Pi_def subsetD)
        apply (rule allI, rule impI, simp,
                rule sc_mem,
               rule Ring.ring_tOp_closed, assumption+,
               simp add:Pi_def,
               simp add:Pi_def subsetD)
        apply (rule allI, rule impI, simp)
 apply (subgoal_tac "(λx∈{j. j ≤ n}. (s ⋅rR sa x)) ∈ {j. j ≤ n} → A",
        blast,
        thin_tac e M (λj. (s ⋅rR sa j) ⋅s f j) n =
        Σe M (λj. (λx∈{j. j ≤ n}. s ⋅rR sa x) j ⋅s f j) n")
        apply (rule Pi_I, simp,
               rule_tac x = s and r = "sa x" in 
               Ring.ideal_ring_multiple1[of R A], assumption+)
               apply (simp add:Pi_def)
done

lemma (in Module) l_span_closed2Tr:"⟦ideal R A; H ⊆ carrier M⟧ ⟹ 
       s ∈ {j. j ≤ (n::nat)} → A ∧ 
       f ∈ {j. j ≤ n} → linear_span R M (carrier R) H ⟶
            l_comb R M n s f ∈ linear_span R M A H"
apply (cut_tac sc_Ring)
apply (induct_tac n)
apply (rule impI, (erule conjE)+)
apply (case_tac "H = {}")
 apply (simp add:linear_span_def)
 apply (simp add:l_comb_def)
 apply (rule sc_a_0,
        cut_tac sc_Ring,
        simp add:Pi_def Ring.ideal_subset) 
 apply (simp add:l_comb_def l_span_closed2Tr0)

apply (rule impI, erule conjE,
       frule func_pre[of s], frule func_pre[of f], simp)
 apply (simp add:l_comb_def) 
 apply (rule linear_span_pOp_closed, assumption+) 
 apply (rule_tac s = "s (Suc n)" and f = "f (Suc n)" in 
                 l_span_closed2Tr0[of A H], assumption+,
       (simp add:Pi_def)+)
done

lemma (in Module) l_span_closed2:"⟦ideal R A; H ⊆ carrier M;
       s ∈ {j. j ≤ (n::nat)} → A ; 
       f ∈ {j. j ≤ n} → linear_span R M (carrier R) H⟧ ⟹
       l_comb R M n s f ∈ linear_span R M A H"
apply (simp add:l_span_closed2Tr)
done

lemma (in Module) l_span_l_span:"H ⊆ carrier M ⟹
       linear_span R M (carrier R) (linear_span R M (carrier R) H) =
                                          linear_span R M (carrier R) H"
apply (cut_tac sc_Ring, frule Ring.whole_ideal[of R])
apply (rule equalityI)
 apply (rule subsetI)
 apply (frule linear_span_inc_0[of "carrier R" H], assumption+,
        frule nonempty[of _ "linear_span R M (carrier R) H"],
        simp add:linear_span_def[of R M "carrier R" 
                            "linear_span R M (carrier R) H"],
        erule exE, (erule bexE)+, simp)
 apply (frule_tac s = s and n = n and f = f in l_span_closed2[of "carrier R"],
        assumption+,
        frule lin_span_sub_carrier[of "carrier R" "H"], assumption+,
        rule subsetI)
 apply (rule_tac h = x in h_in_linear_span[of "linear_span R M (carrier R) H"],
        assumption+)
done

lemma (in Module) l_spanA_l_span:"⟦ideal R A; H ⊆ carrier M⟧ ⟹
       linear_span R M A (linear_span R M (carrier R) H) =
                                          linear_span R M A H"
apply (cut_tac sc_Ring, frule Ring.whole_ideal[of R])
apply (rule equalityI)
 apply (rule subsetI)
 apply (frule linear_span_inc_0[of "carrier R" H], assumption+,
        frule nonempty[of _ "linear_span R M (carrier R) H"],
        simp add:linear_span_def[of R M A 
                            "linear_span R M (carrier R) H"],
        erule exE, (erule bexE)+, simp)
 apply (frule_tac s = s and n = n and f = f in l_span_closed2[of A],
        assumption+)
 apply (frule l_span_cont_H[of H])
 apply (frule l_span_gen_mono[of "H" "linear_span R M (carrier R) H" A],
        simp add:lin_span_sub_carrier[of "carrier R" H], assumption)
 apply assumption
done 

lemma (in Module) l_span_zero:"ideal R A ⟹ linear_span R M A {𝟬} = {𝟬}"
apply (cut_tac sc_Ring)
apply (rule equalityI)
 apply (rule subsetI,
        frule_tac x = x in mem_single_l_span1[of A 𝟬],
        simp add:ag_inc_zero, assumption,
        erule bexE, frule_tac h = a in Ring.ideal_subset[of R A], assumption+,
        simp add:sc_a_0)
 apply (rule subsetI, simp, rule linear_span_inc_0, assumption,
        rule subsetI, simp add:ag_inc_zero)
done

lemma (in Module) l_span_closed3:"⟦ideal R A; generator R M H;
       A ⊙R M = carrier M⟧ ⟹ linear_span R M A H = carrier M"
apply (cut_tac sc_Ring)

apply (rule equalityI) 
 apply (cut_tac linear_span_subModule[of A H],
        simp add:submodule_subset, assumption,
        simp add:generator_def)

apply (rule subsetI) 
 apply (simp add:generator_def)
 apply (erule conjE) 
 apply (case_tac "H = {}", simp, simp add:linear_span_def)
apply (simp add:smodule_ideal_coeff_def)
 apply (rotate_tac -2, frule sym,
        thin_tac "linear_span R M (carrier R) H = carrier M")
 apply simp 
 apply (frule sym, 
        thin_tac "linear_span R M A (linear_span R M (carrier R) H) =
                  linear_span R M (carrier R) H")
 apply (frule_tac a = x in eq_set_inc[of _ "linear_span R M (carrier R) H"
        "linear_span R M A (linear_span R M (carrier R) H)"], assumption+,
        thin_tac "x ∈ linear_span R M (carrier R) H",
        thin_tac "linear_span R M (carrier R) H =
         linear_span R M A (linear_span R M (carrier R) H)")
 apply (frule sym, 
        thin_tac "carrier M = linear_span R M (carrier R) H",
        frule subset_trans[of H "linear_span R M (carrier R) H" "carrier M"],
        simp,
        thin_tac "linear_span R M (carrier R) H = carrier M")
 apply (frule Ring.whole_ideal,
        frule linear_span_inc_0 [of "carrier R" "H"], assumption+,
        frule nonempty [of "𝟬" "linear_span R M (carrier R) H"])
apply (simp add:linear_span_def [of _ _ _ "linear_span R M (carrier R) H"])
 apply (erule exE, (erule bexE)+)
apply (simp add:l_span_closed2) 
done

lemma (in Module) generator_generator:"⟦generator R M H; H1 ⊆ carrier M; 
           H ⊆ linear_span R M (carrier R) H1⟧  ⟹  generator R M H1"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal[of R],
       frule linear_span_subModule[of "carrier R" H1], assumption,
       frule l_span_sub_submodule[of "carrier R" 
            "linear_span R M (carrier R) H1" H], assumption+)
apply (simp add:generator_def)
apply (rule equalityI,
       simp add:submodule_subset, assumption)
done

lemma (in Module) generator_elimTr:
"f ∈ {j. j ≤ (n::nat)} → carrier M ∧ generator R M (f ` {j. j ≤ n}) ∧ 
(∀i∈nset (Suc 0) n. f i ∈ 
   linear_span R M (carrier R) (f ` {j. j ≤ (i - Suc 0)})) ⟶ 
 linear_span R M (carrier R) {f 0} = carrier M"
apply (induct_tac n)
 apply (rule impI, (erule conjE)+)
 apply (simp add:nset_def generator_def)

apply (rule impI)
 apply (erule conjE)+
 apply (frule func_pre [of _ _ "carrier M"], simp)
 apply (subgoal_tac "generator R M (f ` {j. j ≤ n})")
 apply (subgoal_tac "∀i∈nset (Suc 0) n.
         f i ∈ linear_span R M (carrier R) (f ` {j. j ≤ (i - Suc 0)})")
 apply simp
 apply (thin_tac "generator R M (f ` {j. j ≤ n}) ∧
     (∀i∈nset (Suc 0) n. f i ∈ linear_span R M (carrier R) 
              (f ` {j. j ≤ i - Suc 0})) ⟶
         linear_span R M (carrier R) {f 0} = carrier M")
 apply (rule ballI)
 apply (frule_tac x = i in bspec, simp add:nset_def, assumption)
 apply (thin_tac "generator R M (f ` {j. j ≤ n}) ∧
         (∀i∈nset (Suc 0) n.
         f i ∈ linear_span R M (carrier R) (f ` {j. j ≤ i - Suc 0})) ⟶
         linear_span R M (carrier R) {f 0} = carrier M")
 apply (frule_tac x = "Suc n" in bspec, simp add:nset_def,
        thin_tac "∀i∈nset (Suc 0) (Suc n).
            f i ∈ linear_span R M (carrier R) (f ` {j. j ≤ i - Suc 0})",
        simp)
 apply (subgoal_tac "f ` {j. j ≤ Suc n} ⊆ linear_span R M (carrier R) (f ` {j. j ≤ n})")
 apply (frule_tac H = "f ` {j. j ≤ Suc n}" and ?H1.0 = "f ` {j. j ≤ n}"
        in generator_generator,
        rule subsetI, simp add:image_def, erule exE, erule conjE, simp,
        simp add:Pi_def)
 apply assumption+
 apply (rule subsetI, simp add:image_def, erule exE, erule conjE)
 apply (case_tac "xa = Suc n", simp)
 apply (frule_tac m = xa and n = "Suc n" in noteq_le_less, assumption,
        thin_tac "xa ≤ Suc n",
        frule_tac x = xa and n = "Suc n" in less_le_diff, 
        thin_tac "xa < Suc n", simp)
 apply (rule_tac H = "{y. ∃x≤n. y = f x}" and h = "f xa" in 
                       h_in_linear_span,
        rule subsetI, simp add:image_def, erule exE, erule conjE,
        simp add:Pi_def)
 apply (simp, blast)
done

lemma (in Module) generator_generator_elim:
 "⟦f ∈ {j. j ≤ (n::nat)} → carrier M; generator R M (f ` {j. j ≤ n}); 
  (∀i∈nset (Suc 0) n. f i ∈ linear_span R M (carrier R) 
     (f ` {j. j ≤ (i - Suc 0)}))⟧ ⟹ 
   linear_span R M (carrier R) {f 0} = carrier M"
apply (simp add:generator_elimTr [of f n])
done

lemma (in Module) surjec_generator:"⟦R module N; f ∈ mHom R M N;
 surjecM,N f; generator R M H⟧ ⟹ generator R N (f ` H)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (simp add:generator_def, erule conjE)
 apply (simp add:surjec_def, (erule conjE)+)
 apply (simp add:aHom_def, (erule conjE)+)
 apply (simp add:image_sub [of "f" "carrier M" "carrier N" "H"])

apply (frule Module.lin_span_sub_carrier[of N R "carrier R" "f ` H"],
       assumption,
       simp add:image_sub [of "f" "carrier M" "carrier N" "H"])
apply (rule equalityI, assumption+)
 apply (rule subsetI)
 apply (simp add:surj_to_def,
        thin_tac "f ∈ extensional (carrier M)",
        thin_tac "∀a∈carrier M. ∀b∈carrier M. f (a ± b) = f a ±N f b")
 apply (frule sym, rotate_tac 6, frule sym,
        thin_tac "f ` carrier M = carrier N",
        frule_tac a = x and A = "carrier N" and B = "f ` carrier M" in
        eq_set_inc, assumption,
        thin_tac "carrier N = f ` carrier M", 
        thin_tac "carrier M = linear_span R M (carrier R) H")
 apply (simp add:image_def[of f "carrier M"], erule bexE)
 apply (frule sym, thin_tac "linear_span R M (carrier R) H = carrier M",
        frule_tac a = xa in eq_set_inc[of _ "carrier M" 
        "linear_span R M (carrier R) H"], assumption,
        thin_tac "carrier M = linear_span R M (carrier R) H",
        thin_tac "linear_span R N (carrier R) (f ` H) ⊆ carrier N")

 apply (simp add:linear_span_def)
apply (case_tac "H = {}", simp) 
 apply (simp add:mHom_0, simp,
        erule exE, (erule bexE)+)
 apply (cut_tac sc_Ring, frule Ring.whole_ideal[of R],
       frule_tac s = s and n = n and g = fa in 
       linmap_im_linspan[of "carrier R" N f H], assumption+,
       rotate_tac -5, frule sym,
       thin_tac "xa = l_comb R M n s fa", simp,
       thin_tac "l_comb R M n s fa = xa")
 apply (simp add:linear_span_def)
done   

lemma (in Module) surjec_finitely_gen:"⟦R module N; f ∈ mHom R M N;
       surjecM,N f; M fgover R⟧  ⟹ N fgover R"
apply (simp add:fGOver_def)
 apply (erule exE)
 apply (simp add:finite_generator_def [of "R" "M"],erule conjE)
 apply (frule_tac H = H in surjec_generator[of N f], assumption+)
apply (simp add:finite_generator_def [of "R" "N"])
 apply (frule_tac F = H and h = f in finite_imageI)  
 apply blast
done
    
subsection "Sum up coefficients" 
 text{* Symbolic calculation. *}    

lemma (in Module) similar_termTr:"⟦ideal R A; a ∈ A⟧ ⟹
 ∀s. ∀f. s ∈ {j. j ≤ (n::nat)} → A ∧ 
         f ∈ {j. j ≤ n} → carrier M ∧ 
         m ∈ f ` {j. j ≤ n} ⟶
       (∃t∈{j. j ≤ n} → A. nsum M (λj. s j ⋅s (f j)) n ± a ⋅s m = 
           nsum M (λj. t j ⋅s (f j)) n )"
apply (cut_tac sc_Ring)   
apply (induct_tac n)
 apply (rule allI)+ apply (rule impI) apply (erule conjE)+
 apply simp
 apply (rule_tac x = "λk∈{0::nat}. (s 0 ±R a)" in bexI)
 apply (simp add: Ring.ideal_subset sc_l_distr)
 apply simp
   apply (simp add:Ring.ideal_pOp_closed)

(** n **)
apply ((rule allI)+, rule impI, (erule conjE)+)
 apply (simp del:nsum_suc add:image_def)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" in nsum_mem,
        rule allI, rule impI, rule sc_mem,
        simp add:funcset_mem Ring.ideal_subset,
        simp add:funcset_mem,
        frule_tac x = "Suc n" and f = s and A = "{j. j ≤ Suc n}" and
        B = A in funcset_mem, simp,
        frule_tac h = "s (Suc n)" in Ring.ideal_subset, assumption+,
        frule_tac x = "Suc n" and f = f and A = "{j. j ≤ Suc n}" and
        B = "carrier M" in funcset_mem, simp,
        frule_tac a = "s (Suc n)" and m = "f (Suc n)" in sc_mem, assumption+,
        cut_tac a = a and m = m in sc_mem,
        simp add:Ring.ideal_subset, erule exE, simp add:Pi_def,
        erule exE, erule conjE)
 apply (case_tac "x = Suc n", simp)  (***** case x = Suc n ********)
 apply (subst ag_pOp_assoc, assumption+)
 apply (thin_tac e M (λj. s j ⋅s f j) n ∈ carrier M",
        thin_tac "s (Suc n) ⋅s f (Suc n) ∈ carrier M",
        thin_tac "a ⋅s f (Suc n) ∈ carrier M",
        thin_tac "∀s fa.
           s ∈ {j. j ≤ n} → A ∧
           fa ∈ {j. j ≤ n} → carrier M ∧ (∃x≤n. f (Suc n) = fa x) ⟶
           (∃t∈{j. j ≤ n} → A.
               Σe M (λj. s j ⋅s fa j) n ± a ⋅s f (Suc n) =
               Σe M (λj. t j ⋅s fa j) n)")
 apply (subst sc_l_distr[THEN sym], assumption+,
        simp add:Ring.ideal_subset, assumption+)
 apply (frule func_pre[of _ _ A],
        frule_tac f = s and n = n and g = "λk∈{0::nat}. (s (Suc n) ±R a)" and
        m = 0 and A = A and B = A in jointfun_hom0,
        simp add: Ring.ideal_pOp_closed)
 apply (subgoal_tac e M (λj. s j ⋅s f j) n ± (s (Suc n) ±R a) ⋅s f (Suc n) =
      Σe M (λj. (jointfun n s 0 (λk∈{0}. s (Suc n) ±R a)) j ⋅s f j) (Suc n)",
      simp,
      thin_tac e M (λj. s j ⋅s f j) n ± (s (Suc n) ±R a) ⋅s f (Suc n) =
      Σe M (λj. jointfun n s 0 (λk∈{0}. s (Suc n) ±R a) j ⋅s f j) n ±
      jointfun n s 0 (λk∈{0}. s (Suc n) ±R a) (Suc n) ⋅s f (Suc n)")
 apply blast
 apply simp
 apply (simp add:jointfun_def sliden_def)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" and g = "λj. (if j ≤ n then s j
        else (λk∈{0}. s (Suc n) ±R a) (sliden (Suc n) j)) ⋅s f j" in
        nsum_eq)
        apply (rule allI, rule impI, rule sc_mem,
               simp add:Pi_def Ring.ideal_subset,
               simp add:Pi_def)
        apply (rule allI, rule impI, simp, rule sc_mem,
               simp add:Pi_def Ring.ideal_subset,
               simp add:Pi_def)
        apply (rule allI, rule impI, simp)
  apply simp
  
  apply (frule_tac m = x and n = "Suc n" in noteq_le_less, assumption,
         thin_tac "x ≤ Suc n",
         frule_tac x = x and n = "Suc n" in less_le_diff,
         thin_tac "x < Suc n", simp)
  apply (frule func_pre[of _ _ A], frule func_pre[of _ _ "carrier M"])
  apply (drule_tac x = s in spec,
         drule_tac x = f in spec)
   apply (subgoal_tac "∃xa≤n. f x = f xa", simp,
          thin_tac "∃xa≤n. f x = f xa", erule bexE)
   apply (subst ag_pOp_assoc, assumption+,
          frule_tac x = "s (Suc n) ⋅s f (Suc n)" and y = "a ⋅s f x" in 
          ag_pOp_commute, assumption+, simp,
          thin_tac "s (Suc n) ⋅s f (Suc n) ± a ⋅s f x =
          a ⋅s f x ± s (Suc n) ⋅s f (Suc n)",
          subst ag_pOp_assoc[THEN sym], assumption+, simp,
    thin_tac e M (λj. s j ⋅s f j) n ± a ⋅s f x = Σe M (λj. t j ⋅s f j) n")
  apply (frule_tac f = t and n = n and g = "λk∈{0::nat}. s (Suc n)" and
         m = 0 and A = A and B = A in jointfun_hom0,
         simp, simp)
  apply (subgoal_tac e M (λj. t j ⋅s f j) n ± s (Suc n) ⋅s f (Suc n) =
         Σe M (λj. (jointfun n t 0 (λk∈{0}. s (Suc n))) j ⋅s f j) (Suc n)",
         simp,
         thin_tac e M (λj. t j ⋅s f j) n ± s (Suc n) ⋅s f (Suc n) =
        Σe M (λj. jointfun n t 0 (λk∈{0}. s (Suc n)) j ⋅s f j) n ±
        jointfun n t 0 (λk∈{0}. s (Suc n)) (Suc n) ⋅s f (Suc n)")
  apply blast
   apply (simp add:jointfun_def sliden_def)
   apply (cut_tac n = n and f = "λj. t j ⋅s f j" and 
          g = "λj. (if j ≤ n then t j  else (λk∈{0}. s (Suc n)) 
                (sliden (Suc n) j)) ⋅s f j" in nsum_eq)
   apply (rule allI, rule impI, rule sc_mem,
          simp add:Pi_def Ring.ideal_subset,
          simp add:Pi_def)
   apply (rule allI, rule impI, simp, rule sc_mem,
          simp add:Pi_def Ring.ideal_subset,
          simp add:Pi_def)   
   apply (rule allI, rule impI, simp, simp)
   apply blast
done

lemma (in Module) similar_term1:"⟦ideal R A; a ∈ A; s ∈ {j. j≤(n::nat)} → A;
       f ∈ {j. j ≤ n} → carrier M; m ∈ f ` {j. j ≤ n}⟧ ⟹ 
      ∃t∈{j. j ≤ n} → A. Σe M (λj. s j ⋅s (f j)) n ± a ⋅s m =
             Σe M (λj.  t j ⋅s (f j)) n" 
apply (simp add:similar_termTr)
done


lemma (in Module) same_togetherTr:"⟦ideal R A; H ⊆ carrier M ⟧ ⟹ 
 ∀s. ∀f. s∈{j. j ≤ (n::nat)} → A  ∧ f ∈ {j. j ≤ n} → H ⟶ 
 (∃t ∈ {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} → A. 
  ∃g ∈ {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} → f ` {j. j ≤ n}. 
   surj_to g {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} (f ` {j. j ≤ n}) ∧ 
  nsum M (λj. s j ⋅s (f j)) n = nsum M (λk. t k ⋅s (g k)) 
       (card (f ` {j. j ≤ n}) - Suc 0))"  
apply (induct_tac n)
 apply ((rule allI)+, rule impI, erule conjE)
 apply (simp del: Pi_split_insert_domain)
 apply (frule_tac f = f and A = "{0}" and B= H in func_to_img,
        frule_tac f = f and A = "{0}" and B= H in surj_to_image,
        fastforce simp add:image_def)

apply ((rule allI)+, rule impI, erule conjE)
 apply (frule func_pre [of _ _ "A"], frule func_pre [of _ _ "H"])
 apply (drule_tac x = s in spec,
        drule_tac x = f in spec,
        simp, (erule bexE)+ , (erule conjE)+, simp,
        thin_tac e M (λj. s j ⋅s f j) n =
        Σe M (λk. t k ⋅s g k) (card (f ` {j. j ≤ n}) - Suc 0)")

apply (case_tac "f (Suc n) ∈ f ` {j. j ≤ n}")
 apply (frule_tac a = "s (Suc n)" and s = t and 
        n = "card (f ` {j. j ≤ n}) - Suc 0" and f = g and m = "f (Suc n)" in 
        similar_term1[of A],
        simp add:Pi_def,
        assumption,
        frule_tac f = f and A = "{j. j ≤ n}" and B = H in image_sub0,
        frule_tac A = "f ` {j. j ≤ n}" and B = H and C = "carrier M" 
         in subset_trans, assumption,
        rule_tac f = g and A = "{j. j ≤ card (f ` {j. j ≤ n}) - Suc 0}" and 
                 B = "f ` {j. j ≤ n}" and ?B1.0 = "carrier M" in extend_fun, 
        assumption+)
        apply (simp add:surj_to_def)
  apply (erule bexE, simp,
         thin_tac e M (λj. t j ⋅s g j) (card (f ` {j. j ≤ n}) - Suc 0) ±
        s (Suc n) ⋅s f (Suc n) =
        Σe M (λj. ta j ⋅s g j) (card (f ` {j. j ≤ n}) - Suc 0)") 
  apply (simp add:Nset_img0)
  apply blast
  
  apply (frule_tac f = t and n = "card (f ` {j. j ≤ n}) - Suc 0" and A = A and
        g = "λk∈{0::nat}. s (Suc n)" and m = 0 and B = A in jointfun_hom0)
        apply (simp add:Pi_def,
               simp)
  apply (frule_tac f = g and n = "card (f ` {j. j ≤ n}) - Suc 0" and 
         A = "f ` {j. j ≤ n}" and g = "λk∈{0::nat}. f (Suc n)" and m = 0 and 
         B = "{f (Suc n)}" in jointfun_hom0)
        apply (simp add:Pi_def,
               simp)
  apply (subgoal_tac e M (λk. t k ⋅s g k) (card (f ` {j. j ≤ n}) - Suc 0) ±
                s (Suc n) ⋅s f (Suc n) =
        Σe M (λj. (jointfun (card (f ` {j. j ≤ n}) - Suc 0) t 0 (λk∈{0}. 
            s (Suc n))) j ⋅s (jointfun (card (f ` {j. j ≤ n}) - Suc 0) g 0 
        (λk∈{0}. f (Suc n))) j) (card (f ` {j. j ≤ (Suc n)}) - Suc 0)", simp, 
        thin_tac e M (λk. t k ⋅s g k) (card (f ` {j. j ≤ n}) - Suc 0) ±
        s (Suc n) ⋅s f (Suc n) =
        Σe M (λj. jointfun (card (f ` {j. j ≤ n}) - Suc 0) t 0
                   (λk∈{0}. s (Suc n)) j ⋅s
                  jointfun (card (f ` {j. j ≤ n}) - Suc 0) g 0
                   (λk∈{0}. f (Suc n))
                   j) (card (f ` {j. j ≤ Suc n}) - Suc 0)")
  apply (simp del:nsum_suc add:card_image_Nsetn_Suc)
  apply (simp del:nsum_suc add:image_Nset_Suc[THEN sym])
 apply (subgoal_tac "surj_to (jointfun (card (f ` {j. j ≤ n}) - Suc 0) g 0 
       (λk∈{0}. f (Suc n))) {l. l ≤ Suc (card (f ` {j. j ≤ n}) - Suc 0)} 
       (f ` {j. j ≤ Suc n})", blast)

   apply (simp add:surj_to_def)
   apply (frule_tac f = g and n = "card (f ` {j. j ≤ n}) - Suc 0" and A = "f ` {j. j ≤ n}" and g = "λk∈{0}. f (Suc n)" and m = 0 and B = "{f (Suc n)}" in
  im_jointfun)
   apply (simp add:Pi_def)
   apply simp
   apply (simp add:image_Nset_Suc[THEN sym])
   apply (simp add:card_image_Nsetn_Suc)
   apply (simp add:Nset_img)

   apply (frule_tac f = f and A = "{j. j ≤ Suc n}" and B = H in image_sub0)
   apply (frule_tac A = "f ` {j. j ≤ Suc n}" and B = H and C = "carrier M" in
          subset_trans, assumption+)
   apply (cut_tac H = H and s = t and n = "card (f ` {j. j ≤ n}) - Suc 0" 
         and f = g and t = "λk∈{0}. s (Suc n)" and m = 0 and 
         g = "λk∈{0}. f (Suc n)" in 
         l_comb_jointfun_jj[of _ A], assumption+) 
   apply (frule_tac f = f and A = "{j. j ≤ n}" and B = H in image_sub0)
   apply (rule_tac f = g and A = "{j. j ≤ card (f ` {j. j ≤ n}) - Suc 0}" and
          B = "f ` {j. j ≤ n}" in extend_fun[of _ _ _ H], assumption+,
          simp add:Pi_def, simp add:Pi_def)
   apply simp
   apply (simp add:jointfun_def sliden_def)
done

 (* H shall a generator *)
lemma (in Module) same_together:"⟦ideal R A; H ⊆ carrier M; 
       s ∈ {j. j ≤ (n::nat)} → A; f ∈ {j. j ≤ n} → H⟧ ⟹ 
 ∃t ∈ {j. j ≤ (card (f ` {j. j ≤ (n::nat)}) - Suc 0)} → A. 
 ∃g ∈ {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} → f ` {j. j ≤ n}. 
       surj_to g {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} (f ` {j. j ≤ n}) ∧ 
  Σe M (λj. s j ⋅s (f j)) n = 
                  Σe M (λk. t k ⋅s (g k)) (card (f ` {j. j ≤ n}) - Suc 0)"  
apply (simp add:same_togetherTr[of A H])
done

lemma (in Module) one_last:"⟦ideal R A; H ⊆ carrier M; 
      s ∈ {j. j ≤ (Suc n)} → A; f ∈ {j. j ≤ (Suc n)} → H; 
      bij_to f {j. j ≤ (Suc n)} H; j ≤ (Suc n); j ≠ (Suc n)⟧ ⟹ 
 ∃t ∈ {j. j ≤ (Suc n)} → A. ∃g ∈ {j. j ≤ (Suc n)} → H.  
  Σe M (λk. s k  ⋅s (f k)) (Suc n) =  Σe M (λk. t k  ⋅s (g k)) (Suc n) ∧
  g (Suc n) = f j ∧ t (Suc n) = s j ∧ bij_to g {j. j ≤ (Suc n)} H"  
apply (cut_tac sc_Ring)
apply (subgoal_tac "(λk. s k ⋅s (f k)) ∈ {j. j ≤ Suc n} → carrier M")
apply (frule transpos_hom[of j "Suc n" "Suc n"], simp, assumption,
       frule transpos_inj[of j "Suc n" "Suc n"], simp, assumption,
       frule_tac f1 = "λk.  s k ⋅s (f k)" and n1 = n and h1 = 
         "transpos j (Suc n)" in addition2 [THEN sym], assumption+,
       simp del:nsum_suc)
prefer 2  
    apply (rule Pi_I, rule sc_mem,
           simp add:Pi_def Ring.ideal_subset,
           simp add:Pi_def subsetD)
 apply (frule cmp_fun[of "transpos j (Suc n)" "{j. j ≤ Suc n}" 
                         "{j. j ≤ Suc n}" s A], assumption+,
        frule cmp_fun[of "transpos j (Suc n)" "{j. j ≤ Suc n}" 
                         "{j. j ≤ Suc n}" f H], assumption+)
 apply (simp del:nsum_suc add:l_comb_transpos[of A H])
 apply (subgoal_tac "bij_to (cmp f (transpos j (Suc n))) {j. j ≤ (Suc n)} H") 
 apply (subgoal_tac "(cmp f (transpos j (Suc n))) (Suc n) = f j")
 apply (subgoal_tac "(cmp s (transpos j (Suc n))) (Suc n) = s j")
 apply blast
 apply (simp add:cmp_def, simp add:transpos_ij_2,
        simp add:cmp_def, simp add:transpos_ij_2)
 apply (simp add:bij_to_def, rule conjI,
        rule cmp_surj[of "transpos j (Suc n)" "{j. j ≤ Suc n}" 
          "{j. j ≤ Suc n}" f H], assumption+,
        simp add:transpos_surjec, assumption+, simp)
 apply (rule cmp_inj[of "transpos j (Suc n)" "{j. j ≤ Suc n}" 
          "{j. j ≤ Suc n}" f H], assumption+, simp)
done
 
lemma (in Module) finite_lin_spanTr1:"⟦ideal R A; z ∈ carrier M⟧ ⟹
      h ∈ {j. j ≤ (n::nat)} → {z} ∧ t ∈ {j. j ≤ n} → A  ⟶ 
      (∃s∈{0::nat} → A. Σe M (λj. t j ⋅s (h j)) n =  s 0 ⋅s z)"
apply (induct_tac n)
 apply (rule impI)
 apply ((erule conjE)+, simp)
 apply blast 

apply (rule impI) apply (erule conjE)+
 apply (frule func_pre [of _ _ "{z}"], frule func_pre [of _ _ "A"])
apply (simp del:nsum_suc, erule bexE, simp,
       frule_tac f = h and A = "{j. j ≤ Suc n}" and B = "{z}" and x = "Suc n"
       in funcset_mem, simp, simp,
       frule_tac f = t and A = "{j. j ≤ Suc n}" and B = A and x = "Suc n" in
        funcset_mem, simp, cut_tac sc_Ring,
       frule_tac h = "s 0" in Ring.ideal_subset[of R A], assumption+,
       frule_tac h = "t (Suc n)" in Ring.ideal_subset[of R A], assumption+)
 apply (simp add:sc_l_distr[THEN sym])
 apply (subgoal_tac "(λl∈{0::nat}. (s 0 ±R (t (Suc n)))) ∈ {0} → A")
apply (subgoal_tac "(s 0 ±R t (Suc n)) ⋅s z = (λl∈{0::nat}. (s 0 ±R (t (Suc n)))) 0 ⋅s z ") apply blast
 apply simp 
 apply (rule Pi_I) apply simp
 apply (rule Ring.ideal_pOp_closed, assumption+)
done

lemma (in Module) single_span:"⟦ideal R A; z ∈ carrier M;
    h ∈ {j. j ≤ (n::nat)} → {z}; t ∈ {j. j ≤ n} → A⟧ ⟹ 
     ∃s∈{0::nat} → A. Σe M (λj. t j ⋅s (h j)) n =  s 0 ⋅s z"
apply (simp add:finite_lin_spanTr1)
done
(*
lemma (in Module) finite_lin_spanTr2:"⟦ideal R A; ∀m. 
(∃n1. ∃f∈{j. j ≤ n1} → h ` {j. j ≤ n}. ∃s∈{j. j ≤ n1} → A. 
  m = Σe M (λj. s j ⋅s (f j)) n1) ⟶ 
     (∃s∈{j. j ≤ n} → A. m = Σe M (λj. s j ⋅s (h j)) n); 
  h ∈ {j. j ≤ (Suc n)} → carrier M; f ∈ {j. j ≤ n1} → h ` {j. j ≤ n}; 
  s ∈ {j. j ≤ n1} → A; m = Σe M (λj. s j ⋅s (f j)) n1⟧ ⟹ 
  ∃sa∈{j. j ≤ (Suc n)} → A. Σe M (λj. s j ⋅s (f j)) n1 = 
    Σe M (λj. sa j ⋅s (h j)) n ± (sa (Suc n) ⋅s (h (Suc n)))"
 apply (frule_tac 
 apply (subgoal_tac "∃l∈{j. j ≤ n} → A. m = Σe M (λj. l j ⋅s (h j)) n")
 prefer 2 
 apply (thin_tac "h ∈ {j. j ≤ (Suc n)} → carrier M")
 apply blast
 apply (thin_tac " ∀m. (∃n1. ∃f∈Nset n1 → h ` Nset n.
  ∃s∈Nset n1 → A. m = eΣ M (λj. s j ⋆M (f j)) n1) ⟶
              (∃s∈Nset n → A. m = eΣ M (λj.  s j ⋆M (h j)) n)")
 apply (subgoal_tac "∀l∈Nset n → A. m = eΣ M (λj. l j ⋆M (h j)) n ⟶ (∃sa∈Nset (Suc n) → A. eΣ M (λj. s j ⋆M (f j)) n1 = eΣ M (λj. sa j ⋆M (h j)) n +M  (sa (Suc n) ⋆M (h (Suc n))))")
 apply blast
 apply (thin_tac "∃l∈Nset n → A. m = eΣ M (λj. l j ⋆M (h j)) n")
 apply (rule ballI) apply (rule impI)
 apply (frule sym) apply (thin_tac "m = eΣ M (λj. s j ⋆M (f j)) n1")
 apply simp
 apply (thin_tac "m = eΣ M (λj. l j ⋆M (h j)) n")
 apply (thin_tac "eΣ M (λj. s j ⋆M (f j)) n1 = eΣ M (λj. l j ⋆M (h j)) n")
 apply (subgoal_tac "jointfun n l 0 (λx∈Nset 0. (0R)) ∈ Nset (Suc n) → A")
 apply (subgoal_tac " eΣ M (λj. l j ⋆M (h j)) n =
  eΣ M (λj. (jointfun n l 0 (λx∈Nset 0. (0R))) j ⋆M (h j)) n +M  ((jointfun n l 0 (λx∈Nset 0. (0R))) (Suc n)) ⋆M (h (Suc n))")
 apply blast
 apply (subgoal_tac "jointfun n l 0 (λx∈Nset 0. 0R) (Suc n) ⋆M (h (Suc n)) =
  0M") apply simp
 apply (subgoal_tac "eΣ M (λj. jointfun n l 0 (λx∈Nset 0. 0R) j ⋆M (h j)) n =
 eΣ M (λj. l j ⋆M (h j)) n ") apply simp
 apply (frule module_is_ag [of "R" "M"], assumption+)
 apply (subst ag_r_zero, assumption+)
 apply (subgoal_tac "(λj. l j ⋆M (h j)) ∈ Nset n → carrier M")
 apply (rule eSum_mem, assumption+) apply (simp add:n_in_Nsetn)
 apply (rule univar_func_test) apply (rule ballI) 
 apply (rule sprod_mem, assumption+)
 apply (simp add:funcset_mem ideal_subset)
 apply (frule func_pre [of "h" _ "carrier M"])
 apply (simp add:funcset_mem) apply simp
 apply (rule eSum_eq)
 apply (rule module_is_ag [of "R" "M"], assumption+)
 apply (rule univar_func_test)
 apply (rule ballI)
 apply (frule_tac x = x and n = n in Nset_le)
 apply (insert Nset_nonempty[of "0"]) 
 apply (simp add:jointfun_def)
 apply (rule sprod_mem, assumption+)
 apply (simp add:funcset_mem ideal_subset)
 apply (frule func_pre [of "h" _ "carrier M"]) 
 apply (simp add:funcset_mem)
 apply (rule univar_func_test) apply (rule ballI)
 apply (rule sprod_mem, assumption+)
 apply (simp add:funcset_mem ideal_subset)
 apply (frule func_pre [of "h" _ "carrier M"])
 apply (simp add:funcset_mem)
apply (rule ballI)
 apply (frule_tac x = la and n = n in Nset_le)
 apply (simp add:jointfun_def)
 apply (subgoal_tac "0 ∈ Nset 0")
 apply (simp add:jointfun_def sliden_def slide_def)
 apply (rule sprod_0_m, assumption+) 
 apply (subgoal_tac "Suc n ∈ Nset (Suc n)")
 apply (simp add:funcset_mem) apply (simp add:n_in_Nsetn)+ 
apply (frule_tac f = l and n = n and A = A and g = "λx∈Nset 0. 0R" and m = 0
      and B = A in jointfun_hom0)
 apply (rule univar_func_test) apply (rule ballI) apply (simp add:Nset_def)
 apply (simp add:ideal_zero) apply simp
done *) 

definition
  coeff_at_k :: "[('r, 'm) Ring_scheme, 'r, nat] ⇒ (nat ⇒ 'r)" where
  "coeff_at_k R a k = (λj. if j = k then a else (𝟬R))" 

lemma card_Nset_im:"f ∈ {j. j ≤ (n::nat)} → A ⟹ 
                      (Suc 0) ≤ card (f `{j. j ≤ n})"
apply (cut_tac image_Nsetn_card_pos[of f n])
apply (frule_tac m = 0 and n = "card (f ` {i. i ≤ n})" in Suc_leI,
        assumption+)
done 

lemma (in Module) eSum_changeTr1:"⟦ideal R A; 
  t ∈ {k. k ≤ (card (f ` {j. j ≤ (n1::nat)}) - Suc 0)} → A; 
  g ∈ {k. k ≤ (card (f ` {j. j ≤ n1}) - Suc 0)} → f `{j. j ≤ n1}; 
  Suc 0 < card (f `{j. j ≤ n1}); g x = h (Suc n); x = Suc n; 
card (f `{j. j ≤ n1}) - Suc 0 =  Suc (card (f ` {j. j ≤ n1}) - Suc 0 - Suc 0)⟧
  ⟹ 
 Σe M (λk. t k  ⋅s (g k)) (card (f ` {j. j ≤ n1}) - Suc 0) =  
 Σe M (λk. t k  ⋅s (g k)) (card (f ` {j. j ≤ n1}) - Suc 0 - Suc 0) ±  
    (t (Suc (card (f ` {j. j ≤ n1}) - Suc 0 - Suc 0))  ⋅s 
                (g ( Suc (card (f ` {j. j ≤ n1}) - Suc 0 - Suc 0))))"  
apply simp
done

definition
  zeroi :: "[('r, 'm) Ring_scheme] ⇒ nat ⇒ 'r" where
  "zeroi R = (λj. 𝟬R)" 

lemma zeroi_func:"⟦Ring R; ideal R A⟧ ⟹  zeroi R ∈ {j. j ≤ 0} → A"
by (simp add:zeroi_def Ring.ideal_zero)

lemma (in Module) prep_arrTr1:"⟦ideal R A; h ∈ {j. j ≤ (Suc n)} → carrier M;
 f ∈ {j. j ≤ (n1::nat)} → h ` {j. j ≤ (Suc n)}; s ∈ {j. j ≤ n1}→ A; 
 m = l_comb R M n1 s f⟧ ⟹ 
 ∃l∈{j. j ≤ (Suc n)}. (∃s∈{j. j ≤ (l::nat)} → A. 
 ∃g∈ {j. j ≤ l} → h `{j. j ≤ (Suc n)}. m = l_comb R M l s g ∧ 
                      bij_to g {j. j ≤ l} (f ` {j. j ≤ n1}))"
apply (cut_tac sc_Ring)
apply (frule_tac s = s and n = n1 and f = f in  same_together[of A 
      "h ` {j. j ≤ (Suc n)}"]) 
 apply (simp add:image_sub0, assumption+)
 apply (erule bexE)+ 
 apply (simp add:l_comb_def, erule conjE)
 apply (thin_tac e M (λj. s j ⋅s f j) n1 =
           Σe M (λk. t k ⋅s g k) (card (f ` {j. j ≤ n1}) - Suc 0)")
 apply (subgoal_tac "(card (f ` {j. j ≤ n1}) - Suc 0) ∈ {j. j ≤ Suc n}")
 apply (subgoal_tac "g ∈ {k. k ≤ (card (f `{j. j ≤ n1}) - Suc 0)} →
                         h ` {j. j ≤ Suc n}")
 apply (subgoal_tac "bij_to g {k. k ≤ (card (f ` {j. j ≤ n1}) - Suc 0)} (f ` {j. j ≤ n1})")
 apply blast
 prefer 2 
  apply (frule_tac f = f and A = "{j. j ≤ n1}" and B = "h ` {j. j ≤ Suc n}" 
          in image_sub0, simp)
  apply (rule extend_fun, assumption+)
 apply (simp add:bij_to_def)
apply (rule_tac A = "f ` {j. j ≤ n1}" and n = "card (f `{j. j ≤ n1}) - Suc 0" and f = g in Nset2finite_inj)
 apply (rule finite_imageI, simp)
 apply (frule_tac f = f and n = n1 and A = "h ` {j. j ≤ (Suc n)}" in card_Nset_im)
 apply (simp, assumption)
apply (subgoal_tac "finite (h ` {j. j ≤ (Suc n)})")
apply (frule_tac f = f and A = "{j. j ≤ n1}" and B = "h ` {j. j ≤ (Suc n)}" 
       in image_sub0, simp)
 apply (cut_tac B = "h ` {j. j ≤ (Suc n)}" and A = "f ` {j. j ≤ n1}" in 
        card_mono, simp,  assumption+,
        insert finite_Collect_le_nat [of "Suc n"],
        frule card_image_le [of "{j. j ≤ (Suc n)}" "h"],
        frule_tac i = "card (f ` {j. j ≤ n1})" and 
         j = "card (h ` {j. j ≤ (Suc n)})" and k = "card {j. j ≤ (Suc n)}" in
        le_trans, assumption+)
 apply simp
 apply (rule finite_imageI, simp)
done

lemma two_func_imageTr:"⟦ h ∈ {j. j ≤ Suc n} → B; 
   f ∈ {j. j ≤ (m::nat)} → h ` {j. j ≤ Suc n};  h (Suc n) ∉ f ` {j. j ≤ m}⟧
       ⟹ f ∈ {j. j ≤ m} → h ` {j. j ≤ n}" 
apply (rule Pi_I)
    apply (frule_tac x = x and f = f and A = "{j. j ≤ m}" and 
           B = "h ` {j. j ≤ Suc n}" in funcset_mem, assumption)
   apply (thin_tac "h ∈ {j. j ≤ Suc n} → B")
   apply (rule contrapos_pp, simp+)
     apply (simp add:image_def[of h])
     apply (erule exE, erule conjE)
     apply (case_tac "xa ≠ Suc n",
            frule_tac m = xa and n = "Suc n" in noteq_le_less, assumption)
          apply (
            thin_tac "xa ≤ Suc n",
            frule_tac x = xa and n = "Suc n" in less_le_diff,
            thin_tac "xa < Suc n", simp) apply blast
     apply simp
     apply (subgoal_tac "(f x) ∈ f ` {j. j ≤ m}", simp)
       apply (thin_tac "h (Suc n) ∉ f ` {j. j ≤ m}",
                 thin_tac "∀x≤n. h (Suc n) ≠ h x",
                 thin_tac "f x = h (Suc n)",
                 thin_tac "xa = Suc n")
     apply (simp add:image_def, blast)
done

lemma (in Module) finite_lin_spanTr3_0:"⟦bij_to g {j. j ≤ l} (g `{j. j ≤ l});
      ideal R A; 
     ∀na. ∀s∈{j. j ≤ na} → A.
                ∀f∈{j. j ≤ na} → h ` {j. j ≤ n}.
                   ∃t∈{j. j ≤ n} → A. l_comb R M na s f = l_comb R M n t h;
     h ∈ {j. j ≤ Suc n} → carrier M; s ∈ {j. j ≤ m} → A;
     f ∈ {j. j ≤ m} → h ` {j. j ≤ Suc n}; 
     l ≤ Suc n; sa ∈ {j. j ≤ l} → A; g ∈ {j. j ≤ l} → h ` {j. j ≤ Suc n};
     0 < l; f ` {j. j ≤ m} = g ` {j. j ≤ l}; h (Suc n) = g l⟧
 ⟹ ∃t∈{j. j ≤ Suc n} → A. l_comb R M l sa g = l_comb R M (Suc n) t h"
  apply (cut_tac sc_Ring)
  apply (subgoal_tac "l_comb R M l sa g = l_comb R M (Suc (l - Suc 0)) sa g",
         simp del:Suc_pred,
         thin_tac "l_comb R M l sa g = l_comb R M (Suc (l - Suc 0)) sa g",
         simp del:Suc_pred add:l_comb_def)
  apply (drule_tac x = "l - Suc 0" in spec,
         drule_tac x = sa in bspec)
        
  apply (rule Pi_I, simp)
        apply (rule_tac x = x and f = sa and A = "{j. j ≤ l}"and B = A
               in funcset_mem, assumption, simp) (*
        apply (rule_tac i = x and j = "l - Suc 0" and k = l in le_trans)
apply (
               assumption, subst Suc_le_mono[THEN sym], simp) *)
  apply (drule_tac x = g in bspec,
         thin_tac "f ∈ {j. j ≤ m} → h ` {j. j ≤ Suc n}",
         thin_tac "sa ∈ {j. j ≤ l} → A",
         thin_tac "f ` {j. j ≤ m} = g ` {j. j ≤ l}")
      apply (rule Pi_I, simp)
      apply (frule_tac x = x and f = g and A = "{j. j ≤ l}" and 
         B = "h ` {j. j ≤ Suc n}" in funcset_mem)
      apply simp (*
      apply (rule_tac i = x and j = "l - Suc 0" and k = l in Nat.le_trans,
             assumption, subst Suc_le_mono[THEN sym], simp) *)
      apply (unfold bij_to_def, frule conjunct2, fold bij_to_def,
             thin_tac "bij_to g {j. j ≤ l} (g ` {j. j ≤ l})",
             thin_tac "g ∈ {j. j ≤ l} → h ` {j. j ≤ Suc n}")
      apply (simp add:image_def, erule exE, erule conjE)
      apply (case_tac "xa = Suc n", simp add:inj_on_def,
             drule_tac a = x in forall_spec) apply simp
(*
      apply (frule_tac i = x and j = "l - Suc 0" and k = l in Nat.le_trans,
              subst Suc_le_mono[THEN sym], simp, assumption) *)
      apply(drule_tac a = l in forall_spec, simp) 
      apply (cut_tac n1 = l and m1 = "l - Suc 0" in Suc_le_mono[THEN sym])
             apply simp
     apply (frule_tac m = xa and n = "Suc n" in noteq_le_less, assumption,
            thin_tac "xa ≤ Suc n",
            frule_tac x = xa and n = "Suc n" in less_le_diff,
            thin_tac "xa < Suc n", simp)
      apply blast
      apply (erule bexE, simp)
      apply (rotate_tac -4, frule sym, thin_tac "h (Suc n) = g l", simp)
   apply (frule_tac f = t and n = n and A = A and g = "λk∈{0::nat}. sa l"
          and m = 0 and B = A in jointfun_hom0,
          simp add:Pi_def, simp)
   apply (subgoal_tac " Σe M (λj. t j ⋅s h j) n ± sa l ⋅s h (Suc n) =
           Σe M (λj. (jointfun n t 0 (λk∈{0}. sa l)) j ⋅s h j) (Suc n)",
          simp, blast) 
   apply (cut_tac H = "carrier M" and A = A and s = t and f = h and n = n and
          m = 0 and t = "λk∈{0}. sa l" in l_comb_jointfun_jf)
          apply simp+ 
          apply (simp add:Pi_def)
          apply simp
   apply (simp add:jointfun_def sliden_def, simp)
done
 
lemma (in Module) finite_lin_spanTr3:"ideal R A ⟹ 
       h ∈ {j. j ≤ (n::nat)} → carrier M ⟶ 
      (∀na. ∀s ∈ {j. j ≤ (na::nat)} → A. 
       ∀f∈ {j. j ≤ na} → (h ` {j. j ≤ n}). (∃t ∈ {j. j ≤ n} → A. 
       l_comb R M na s f = l_comb R M n t h))"
apply (cut_tac sc_Ring)
apply (induct_tac n)
 apply (rule impI, rule allI, (rule ballI)+) 
 apply (insert Nset_nonempty [of "0"]) 
 apply (simp add:l_comb_def)
 apply (frule_tac z = "h 0" and h = f and t = s and n = na in 
          single_span [of A])
 apply (simp add:Pi_def)
 apply assumption+
(********** n = 0 done ***********)
apply (rule impI, rule allI, (rule ballI)+) 
 apply (frule func_pre, simp)
 apply (case_tac "h (Suc n) ∉  f ` {j. j ≤ na}")
  apply (frule_tac h = h and n = n and B = "carrier M" and f = f and
         m = na in two_func_imageTr, assumption+)
  apply (drule_tac x = na in spec,
         drule_tac x = s in bspec, assumption,
         drule_tac x = f in bspec, assumption)
        
  apply (erule bexE, simp )
  apply (thin_tac "l_comb R M na s f = l_comb R M n t h") 
apply (simp add:l_comb_def)
 apply (subgoal_tac e M (λj. t j  ⋅s (h j)) n =
        Σe M (λj. (jointfun n t 0 (zeroi R)) j ⋅s (h j)) (Suc n)", simp,
        thin_tac e M (λj. t j ⋅s h j) n =
        Σe M (λj. jointfun n t 0 (zeroi R) j ⋅s h j) n ±
        jointfun n t 0 (zeroi R) (Suc n) ⋅s h (Suc n)")
 apply (frule_tac f = t and n = n and g = "zeroi R" and m = 0 and A = A and 
        B = A in jointfun_hom)
        apply (rule zeroi_func, assumption+, simp, blast)
 apply (cut_tac H = "carrier M" and s = t and n = n and f = h and m = 0 and 
        t = "zeroi R" in l_comb_jointfun_jf[of _ A],
        simp, assumption+, simp,
        rule zeroi_func, assumption+)
 apply (simp,
       thin_tac e M (λj. jointfun n t 0 (zeroi R) j ⋅s h j) n =
        Σe M (λj. t j ⋅s h j) n",
        simp add:jointfun_def sliden_def zeroi_def,
        subst sc_0_m, simp add:Pi_def,
        subst ag_r_zero,
        rule nsum_mem, rule allI, rule impI, rule sc_mem,
               simp add:Pi_def Ring.ideal_subset,
               simp add:Pi_def,
         simp)

(*** case h (Suc n) ∉  f ` (Nset na) done ***)

apply simp
apply (frule_tac h = h and n = n and m = "l_comb R M na s f" in 
               prep_arrTr1 [of "A"], assumption+, simp)
apply (erule bexE)+
 apply (simp, (erule conjE)+)
 apply (case_tac "l = 0", simp)
 apply (unfold bij_to_def, frule conjunct1, frule conjunct2, fold bij_to_def)
 apply (thin_tac "l_comb R M na s f = l_comb R M 0 sa g")
 apply (simp add:l_comb_def)
 apply (simp add:surj_to_def, rotate_tac -1, frule sym, 
        thin_tac "{g 0} = f ` {j. j ≤ na}", simp,
        rotate_tac -6, frule sym, thin_tac "h (Suc n) = g 0", simp)
 apply (cut_tac f = "zeroi R" and n = n and g = "λj. sa 0" and m = 0 and 
         A = A and B = A in jointfun_hom0)
        apply (simp add:zeroi_def Ring.ideal_zero)
        apply (simp add:Pi_def)
        apply simp
 apply (subgoal_tac "sa 0 ⋅s h (Suc n) = nsum M (λj. (jointfun n (zeroi R) 0 
         (λj. sa 0) j ⋅s h j)) (Suc n)", simp,
        thin_tac "sa 0 ⋅s h (Suc n) =
        Σe M (λj. jointfun n (zeroi R) 0 (λj. sa 0) j ⋅s h j) n ±
        jointfun n (zeroi R) 0 (λj. sa 0) (Suc n) ⋅s h (Suc n)",
        blast)
 apply simp
 apply (cut_tac n = n and f = "λj. jointfun n (zeroi R) 0 (λj. sa 0) j ⋅s h j"
        in nsum_zeroA)
 apply (rule allI, rule impI,
        simp add:jointfun_def zeroi_def,
        rule sc_0_m, simp add:Pi_def, simp,
       thin_tac e M (λj. jointfun n (zeroi R) 0 (λj. sa 0) j ⋅s h j) n = 𝟬")
 apply (simp add:jointfun_def sliden_def,
        subst ag_l_zero,
        rule sc_mem, simp add:Pi_def Ring.ideal_subset,
        simp add:Pi_def, simp)
 (**** l = 0 done ***)
apply (simp)
 apply (thin_tac "l_comb R M na s f = l_comb R M l sa g")
 apply (unfold bij_to_def, frule conjunct1, frule conjunct2, fold bij_to_def)
 apply (simp add:surj_to_def, rotate_tac -2, frule sym,
        thin_tac "g ` {j. j ≤ l} = f ` {j. j ≤ na}", simp)
  apply (subgoal_tac "∃x∈{j. j ≤ l}. h (Suc n) = g x")  
  prefer 2  apply (simp add:image_def) 
apply (erule bexE)
  apply (case_tac "x = l", simp)
apply (frule_tac g = g and l = l and A = A and h = h and n = n and s = s and
       m = na and f = f and l = l and sa = sa in finite_lin_spanTr3_0,
       assumption+)

apply (subgoal_tac "l_comb R M l sa g = l_comb R M (Suc (l - Suc 0)) sa g")
   prefer 2 apply simp
  apply (simp del:nsum_suc Suc_pred,
          thin_tac "l_comb R M l sa g = l_comb R M (Suc (l - Suc 0)) sa g",
          simp del:nsum_suc Suc_pred add:l_comb_def)
   apply (cut_tac f1 = "λj. sa j ⋅s g j" and n1 = "l - Suc 0" and
          h1 = "transpos x (Suc (l - Suc 0))" in addition2[THEN sym],
          thin_tac "∀na. ∀s∈{j. j ≤ na} → A.
                ∀f∈{j. j ≤ na} → h ` {j. j ≤ n}.
                   ∃t∈{j. j ≤ n} → A.
                      Σe M (λj. s j ⋅s f j) na = Σe M (λj. t j ⋅s h j) n",
              rule Pi_I, simp)
       apply (rule sc_mem, 
              simp add:Pi_def Ring.ideal_subset,
              frule_tac f = h and A = "{j. j ≤ Suc n}" and B = "carrier M" in
              image_sub0,
              frule_tac x = xa and f = g and A = "{j. j ≤ l}" and 
              B = "h ` {j. j ≤ Suc n}" in funcset_mem, simp, simp add:subsetD)
       apply (simp,
             rule_tac i = x and n = l and j = l in transpos_hom,
             assumption+, simp, assumption+)
       apply (simp,
              rule_tac i = x and n = l and j = l in transpos_inj,
              assumption+, simp, assumption+)
    apply (simp del:Suc_pred nsum_suc)
    apply (subst l_comb_transpos[of A "carrier M"], assumption, simp,
           simp, simp,
           rule Pi_I,
           frule_tac f = h and A = "{j. j ≤ Suc n}" and B = "carrier M" in
              image_sub0,
              frule_tac x = xa and f = g and A = "{j. j ≤ l}" and 
              B = "h ` {j. j ≤ Suc n}" in funcset_mem, simp, simp add:subsetD,
            simp)  
     apply (simp del:Suc_pred, simp,
            thin_tac e M (λj. sa j ⋅s g j) (l - Suc 0) ± sa l ⋅s g l =
        Σe M (cmp (λj. sa j ⋅s g j) (transpos x l)) (l - Suc 0) ±
        cmp (λj. sa j ⋅s g j) (transpos x l) l")

apply (cut_tac g = "cmp g (transpos x l)" and l = l and A = A and
      h = h and n = n and s = s and m = na and f = f and l = l and
      sa = "cmp sa (transpos x l)" in finite_lin_spanTr3_0)

   apply (frule_tac i = x and n = l and j = l in transpos_hom,
          simp, assumption)
   apply (cut_tac n = l in Nat.le_refl)
   apply (frule_tac i = x and n = l and j = l in transpos_surjec, assumption+)
   apply (frule_tac f = "transpos x l" and A = "{j. j ≤ l}" and 
         B = "{j. j ≤ l}" and g = g and C = "g ` {j. j ≤ l}" in cmp_surj,
         assumption+)
   apply (rule_tac f = g and A = "{j. j ≤ l}" and B = "h ` {j. j ≤ Suc n}"
          in func_to_img, assumption)
   apply (simp add:bij_to_def)
   apply (subst bij_to_def, simp)
   apply (subgoal_tac "cmp g (transpos x l) ` {j. j ≤ l} = g ` {j. j ≤ l}",
          simp) 
   apply (frule_tac f = "transpos x l" and A = "{j. j ≤ l}" and 
         B = "{j. j ≤ l}" and g = g and C = "h ` {j. j ≤ Suc n}" in cmp_inj,
         assumption+)
   apply (rule_tac i = x and n = l and j = l in transpos_inj, assumption,
          simp, assumption, simp add:bij_to_def,
          assumption)
   apply (simp add:cmp_fun_image, simp add:surj_to_def)

   apply assumption+ 
   apply (simp add:l_comb_def)
   apply assumption+
   
   apply (rule Pi_I)
   apply (simp add:cmp_def)
   apply (cut_tac n = l in Nat.le_refl,
          frule_tac i = x and n = l and j = l and l = xa in transpos_mem,
          assumption+,
          simp add:Pi_def)
   apply (rule Pi_I,
          simp add:cmp_def,
          cut_tac n = l in Nat.le_refl,
          frule_tac i = x and n = l and j = l and l = xa in transpos_mem,
          assumption+,
          simp add:Pi_def)
   apply simp
   
   apply (cut_tac n = l in Nat.le_refl,
          frule_tac i = x and n = l and j = l in transpos_surjec, assumption+)
     apply (frule_tac i = x and n = l and j = l in transpos_hom,
           simp, assumption)
   apply (frule_tac f = "transpos x l" and A = "{i. i ≤ l}" and 
          B = "{i. i ≤ l}" and g = g and C = "h ` {j. j ≤ Suc n}" in
          cmp_fun_image, assumption+)
   apply (simp add:surj_to_def)
   apply (simp add:cmp_def)
   apply (simp add:transpos_ij_2)
   apply (erule bexE)
   apply (thin_tac "∀na. ∀s∈{j. j ≤ na} → A.
                ∀f∈{j. j ≤ na} → h ` {j. j ≤ n}.
                   ∃t∈{j. j ≤ n} → A.
                      Σe M (λj. s j ⋅s f j) na = Σe M (λj. t j ⋅s h j) n")
   apply (rename_tac n na s f l sa g x sb)
  apply (subgoal_tac "l_comb R M l (cmp sa (transpos x l)) 
   (cmp g (transpos x l)) = l_comb R M (Suc (l - Suc 0)) 
    (cmp sa (transpos x l)) (cmp g (transpos x l)) ",
     simp del:Suc_pred,
     thin_tac "l_comb R M l (cmp sa (transpos x l)) (cmp g (transpos x l)) =
        l_comb R M (Suc n) sb h")
  apply (simp del:Suc_pred add:l_comb_def, simp,
         thin_tac " Σe M (λj. cmp sa (transpos x l) j ⋅s
                  cmp g (transpos x l) j) (l - Suc 0) ±
         cmp sa (transpos x l) l ⋅s cmp g (transpos x l) l =
         Σe M (λj. sb j ⋅s h j) n ± sb (Suc n) ⋅s g x")
   apply (rotate_tac -3, frule sym, thin_tac "h (Suc n) = g x", simp)
   apply blast

   apply simp
done
     
lemma (in Module) finite_lin_span:
"⟦ideal R A;  h ∈ {j. j ≤ (n::nat)} → carrier M; s ∈ {j. j ≤ (n1::nat)} → A;
 f ∈ {j. j ≤ n1} → h ` {j. j ≤ n}⟧ ⟹ ∃t∈{j. j ≤ n} → A.
              l_comb R M n1 s f = l_comb R M n t h"
apply (simp add:finite_lin_spanTr3)
done

subsection "Free generators"

definition
  free_generator :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, 'a set]
        ⇒ bool" where
 "free_generator R M H ⟷ generator R M H ∧
      (∀n. (∀s f. (s ∈ {j. j ≤ (n::nat)} → carrier R ∧
                   f ∈ {j. j ≤ n} → H ∧ inj_on f {j. j ≤ n} ∧ 
         l_comb R M n s f = 𝟬M) ⟶ s ∈ {j. j ≤ n} → {𝟬R}))"

lemma (in Module) free_generator_generator:"free_generator R M H ⟹
                  generator R M H"
by (simp add:free_generator_def)

lemma (in Module) free_generator_sub:"free_generator R M H ⟹ 
                    H ⊆ carrier M"
by (simp add:free_generator_def generator_def)

lemma (in Module) free_generator_nonzero:"⟦¬ (zeroring R); 
                free_generator R M H; h ∈ H⟧ ⟹ h ≠ 𝟬"
apply (cut_tac sc_Ring)
apply (rule contrapos_pp, simp+)
 apply (simp add:free_generator_def, (erule conjE)+)
 apply (subgoal_tac "(λt. 1rR) ∈ {j. j ≤ (0::nat)} → carrier R")
 apply (subgoal_tac "(λt. 𝟬) ∈ {j. j ≤ (0::nat)} → H ∧ 
                     inj_on (λt. 𝟬) {j. j ≤ (0::nat)} ∧
        l_comb R M 0 (λt.  1rR) (λt.  𝟬) =  𝟬")
 apply (subgoal_tac "(λt.  1rR) ∈ {j. j ≤ (0::nat)} → {𝟬R}") 
 prefer 2 apply blast
 apply (frule_tac f = "λt. 1rR" and A = "{j. j ≤ (0::nat)}" and B = "{𝟬R}" 
        and x = 0 in funcset_mem, simp, simp)
 apply (frule Ring.Zero_ring1 [of "R"], assumption+, simp)
apply simp
 apply (thin_tac "∀n s. s ∈ {j. j ≤ n} → carrier R ∧
           (∃f. f ∈ {j. j ≤ n} → H ∧
                inj_on f {j. j ≤ n} ∧ l_comb R M n s f = 𝟬) ⟶
           s ∈ {j. j ≤ n} → {𝟬R}")
 apply (simp add:l_comb_def)
 apply (rule sc_a_0)
 apply (simp add:Ring.ring_one)
 apply (simp add:Ring.ring_one)
done

lemma (in Module) has_free_generator_nonzeroring:" ⟦free_generator R M H; 
      ∃p ∈ linear_span R M (carrier R) H. p ≠ 𝟬 ⟧  ⟹ ¬ zeroring R"
apply (erule bexE, simp add:linear_span_def)
 apply (case_tac "H = {}", simp, simp)
 apply (erule exE, (erule bexE)+, simp,
        thin_tac "p = l_comb R M n s f")
apply (rule contrapos_pp, simp+)
 apply (simp add:zeroring_def, erule conjE)
 apply (frule Ring.ring_one[of "R"], simp)
 apply (simp add:l_comb_def)
 apply (cut_tac n = n and f = "λj. s j ⋅s f j" in nsum_zeroA)
 apply (rule allI, rule impI)
 apply (simp add:free_generator_def generator_def, frule conjunct1,
        frule_tac x = j and f = f and A = "{j. j ≤ n}" and B = H in
        funcset_mem, simp,
        frule_tac c = "f j" in subsetD[of H "carrier M"], assumption+,
       frule_tac x = j and f = s and A = "{j. j ≤ n}" and B = "{𝟬R}" in
       funcset_mem, simp, simp add:sc_0_m)
 apply simp
done

lemma (in Module) unique_expression1:"⟦H ⊆ carrier M; free_generator R M H;
      s ∈ {j. j ≤ (n::nat)} → carrier R; m ∈ {j. j ≤ n} → H; 
      inj_on m {j. j ≤ n}; l_comb R M n s m = 𝟬⟧ ⟹ 
                                 ∀j∈{j. j ≤ n}. s j = 𝟬R" 
apply (rule ballI)
apply (simp add:free_generator_def, (erule conjE)+)
apply (subgoal_tac "s ∈ {j. j ≤ n} → {𝟬R}")
 apply (frule_tac f = s and A = "{j. j ≤ n}" and B = "{𝟬R}" and x = j in 
        funcset_mem, simp, simp)
apply blast
done

lemma (in Module) free_gen_coeff_zero:"⟦H ⊆ carrier M; free_generator R M H;
       h ∈ H; a ∈ carrier R; a ⋅s h = 𝟬⟧ ⟹ a = 𝟬R"
apply (frule unique_expression1[of H "λx∈{0::nat}. a" 0 "λx∈{0::nat}. h"],
        assumption+,
       simp,
       simp,
       simp add:inj_on_def,
       simp add:l_comb_def,
       simp)
done

lemma (in Module) unique_expression2:"⟦H ⊆ carrier M; 
      f ∈ {j. j ≤ (n::nat)} → H; s ∈ {j. j ≤ n} → carrier R⟧ ⟹
    ∃m g t. g ∈ ({j. j ≤ (m::nat)} → H) ∧ 
            bij_to g {j. j ≤ (m::nat)} (f ` {j. j ≤ n}) ∧ 
            t ∈ {j. j ≤ m} → carrier R ∧ 
            l_comb R M n s f = l_comb R M m t g" 
apply (cut_tac sc_Ring)
apply (frule Ring.whole_ideal [of "R"])
apply (frule_tac  A = "carrier R" and H = H and s = s and f = f in 
       same_together, assumption+)
apply ((erule bexE)+, erule conjE)
apply (frule_tac f = f and A = "{j. j ≤ n}" in image_sub0,
       frule_tac f = g and A = "{j. j ≤ card (f ` {j. j ≤ n}) - Suc 0}" 
       and B = "f ` {j. j ≤ n}" in extend_fun[of _ _ _ "H"], assumption)
apply (subgoal_tac "bij_to g {j. j ≤ (card (f ` {j. j ≤ n}) - Suc 0)} 
                                  (f ` {j. j ≤ n})")
 apply (simp add:l_comb_def, blast)
apply (simp add:bij_to_def)
apply (cut_tac finite_Collect_le_nat[of n],
        frule finite_imageI[of "{j. j ≤ n}" f])
apply (rule_tac A = "f ` {j. j ≤ n}" and n = "card (f ` {j. j ≤ n}) - 
        Suc 0" and f = g in Nset2finite_inj, assumption)
 using image_Nsetn_card_pos[of f n] apply simp
apply assumption
done

lemma (in Module) unique_expression3_1:"⟦H ⊆ carrier M; 
      f ∈ {l. l ≤ (Suc n)} → H; s ∈ {l. l ≤ (Suc n)} → carrier R; 
      (f (Suc n)) ∉ f `({l. l ≤ (Suc n)} - {Suc n})⟧ ⟹ 
     ∃g m t. g ∈ {l. l ≤ (m::nat)} → H ∧ 
             inj_on g {l. l ≤ (m::nat)} ∧ 
             t ∈ {l. l ≤ (m::nat)} → carrier R ∧ 
             l_comb R M (Suc n) s f = 
                 l_comb R M m t g ∧ t m = s (Suc n) ∧ g m = f (Suc n)"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal)
apply (simp add:Nset_pre1)
 apply (subst l_comb_Suc[of H "carrier R" s n f], assumption+)
 apply (frule func_pre[of _ _ H], frule func_pre[of _ _ "carrier R"])
 apply (frule unique_expression2[of H f n s], assumption+)
 apply ((erule exE)+, (erule conjE)+, simp,
        thin_tac "l_comb R M n s f = l_comb R M m t g")
 apply (frule_tac f = g and n = m and A = H and g = "λk∈{0::nat}. f (Suc n)"
         and m = 0 and B = H in jointfun_hom0,
        simp add:Pi_def, simp)
 apply (frule_tac f = t and n = m and A = "carrier R" and 
        g = "λk∈{0::nat}. s (Suc n)"  and m = 0 and B = "carrier R" in 
        jointfun_hom0,
        simp add:Pi_def, simp)
 apply (subgoal_tac "inj_on (jointfun m g 0 (λk∈{0}. f (Suc n))) 
                       {l. l ≤ Suc m}",
    subgoal_tac "l_comb R M m t g ± s (Suc n) ⋅s f (Suc n) =
        l_comb R M (Suc m) (jointfun m t 0 (λk∈{0}. s (Suc n))) 
                             (jointfun m g 0 (λk∈{0}. f (Suc n)))",
    subgoal_tac "(jointfun m t 0 (λk∈{0}. s (Suc n))) (Suc m) = s (Suc n) ∧
                 (jointfun m g 0 (λk∈{0}. f (Suc n))) (Suc m) = f (Suc n)",
    simp, blast)
 apply (simp add:jointfun_def sliden_def)
  apply (frule_tac s = t and n = m and f = g and t = "λk∈{0}. s (Suc n)" and
         m = 0 and g = "λk∈{0}. f (Suc n)" in l_comb_jointfun_jj[of H 
        "carrier R"], assumption+,
         simp add:Pi_def, simp,
         simp add:Pi_def)
  apply (simp add:l_comb_def, simp add:jointfun_def sliden_def)
  apply (thin_tac "jointfun m g 0 (λk∈{0}. f (Suc n)) ∈ {l. l ≤ Suc m} → H",
  thin_tac "jointfun m t 0 (λk∈{0}. s (Suc n)) ∈ {l. l ≤ Suc m} → carrier R",
  thin_tac "t ∈ {j. j ≤ m} → carrier R", 
  thin_tac "s ∈ {j. j ≤ n} → carrier R")
 apply (rule_tac f = g and n = m and b = "f (Suc n)" and B = H in jointfun_inj,
        assumption+)
  apply (simp add:bij_to_def)
  apply (unfold bij_to_def, frule conjunct1, fold bij_to_def,
         simp add:surj_to_def)
done
(*
lemma (in Module) unique_expression3_1:"⟦H ⊆ carrier M; 
      f ∈ {l. l ≤ (Suc n)} → H; s ∈ {l. l ≤ (Suc n)} → carrier R; 
      (f (Suc n)) ∉ f `({l. l ≤ (Suc n)} - {Suc n})⟧ ⟹ 
     ∃g m t. g ∈ {l. l ≤ (m::nat)} → H ∧ 
             inj_on g {l. l ≤ (m::nat)} ∧ 
             t ∈ {l. l ≤ (m::nat)} → carrier R ∧ 
             l_comb R M (Suc n) s f = l_comb R M m t g ∧ 
              t m = s (Suc n)"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal)
apply (simp add:Nset_pre1)
 apply (subst l_comb_Suc[of H "carrier R" s n f], assumption+)
 apply (frule func_pre[of _ _ H], frule func_pre[of _ _ "carrier R"])
 apply (frule unique_expression2[of H f n s], assumption+)
 apply ((erule exE)+, (erule conjE)+, simp,
        thin_tac "l_comb R M n s f = l_comb R M m t g")
 apply (frule_tac f = g and n = m and A = H and g = "λk∈{0::nat}. f (Suc n)"
         and m = 0 and B = H in jointfun_hom0,
        rule univar_func_test, rule ballI, simp add:funcset_mem, simp)
 apply (frule_tac f = t and n = m and A = "carrier R" and 
        g = "λk∈{0::nat}. s (Suc n)"  and m = 0 and B = "carrier R" in 
        jointfun_hom0,
        rule univar_func_test, rule ballI, simp add:funcset_mem, simp)
 apply (subgoal_tac "inj_on (jointfun m g 0 (λk∈{0}. f (Suc n))) 
                       {l. l ≤ Suc m}",
    subgoal_tac "l_comb R M m t g ± s (Suc n) ⋅s f (Suc n) =
        l_comb R M (Suc m) (jointfun m t 0 (λk∈{0}. s (Suc n))) 
                             (jointfun m g 0 (λk∈{0}. f (Suc n)))",
    subgoal_tac "(jointfun m t 0 (λk∈{0}. s (Suc n))) (Suc m) = s (Suc n)",
    simp, blast)
 apply (simp add:jointfun_def sliden_def)
  apply (frule_tac s = t and n = m and f = g and t = "λk∈{0}. s (Suc n)" and
         m = 0 and g = "λk∈{0}. f (Suc n)" in l_comb_jointfun_jj[of H 
        "carrier R"], assumption+,
         rule univar_func_test, rule ballI, simp add:funcset_mem, simp,
         rule univar_func_test, rule ballI, simp add:funcset_mem)
  apply (simp add:l_comb_def, simp add:jointfun_def sliden_def)
  apply (thin_tac "jointfun m g 0 (λk∈{0}. f (Suc n)) ∈ {l. l ≤ Suc m} → H",
  thin_tac "jointfun m t 0 (λk∈{0}. s (Suc n)) ∈ {l. l ≤ Suc m} → carrier R",
  thin_tac "t ∈ {j. j ≤ m} → carrier R", 
  thin_tac "s ∈ {j. j ≤ n} → carrier R")
 apply (rule_tac f = g and n = m and b = "f (Suc n)" and B = H in jointfun_inj,
        assumption+)
  apply (simp add:bij_to_def)
  apply (unfold bij_to_def, frule conjunct1, fold bij_to_def,
         simp add:surj_to_def)
done     *)

lemma (in Module) unique_expression3_2:"⟦H ⊆ carrier M; 
      f ∈ {k. k ≤ (Suc n)} → H; s ∈ {k. k ≤ (Suc n)} → carrier R; 
      l ≤ (Suc n); (f l) ∉ f ` ({k. k ≤ (Suc n)} - {l}); l ≠ Suc n⟧ ⟹ 
    ∃g m t. g ∈ {l. l ≤ (m::nat)} → H ∧ inj_on g {l. l ≤ (m::nat)} ∧ 
            t ∈ {l. l ≤ m} → carrier R ∧ 
            l_comb R M (Suc n) s f = l_comb R M m t g ∧ 
             t m = s l ∧ g m = f l"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal)
 apply (subst l_comb_transpos1[of "carrier R" H s n f l], assumption+,
        rule noteq_le_less[of l "Suc n"], assumption+) 
 apply (cut_tac unique_expression3_1[of H "cmp f (transpos l (Suc n))" n 
        "cmp s (transpos l (Suc n))"])
 apply ((erule exE)+, (erule conjE)+, simp)
 apply (subgoal_tac "t m = s l ∧ g m = f l", blast)
 apply (thin_tac "l_comb R M (Suc n) (cmp s (transpos l (Suc n)))
         (cmp f (transpos l (Suc n))) = l_comb R M m t g")
 apply (simp add:cmp_def)
 apply (subst transpos_ij_2[of l "Suc n" "Suc n"], simp+,
        subst transpos_ij_2[of l "Suc n" "Suc n"], simp+) 
 apply (rule Pi_I, simp add:cmp_def,
        frule_tac l = x in transpos_mem[of l "Suc n" "Suc n"], simp,
         assumption+, simp add:Pi_def)
 apply (rule Pi_I, simp add:cmp_def,
        frule_tac l = x in transpos_mem[of l "Suc n" "Suc n"], simp,
         assumption+, simp add:Pi_def)
 apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in transpos_hom,
           simp, assumption)
 apply (frule cmp_fun_sub_image[of "transpos l (Suc n)" "{i. i ≤ Suc n}" 
       "{i. i ≤ Suc n}" f H "{l. l ≤ Suc n} - {Suc n}"], assumption+)
       apply (rule subsetI, simp)
       apply simp
       apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in transpos_inj,
              simp, assumption+)
       apply (subst injfun_elim_image[of "transpos l (Suc n)" "{i. i ≤ Suc n}"
        "{i. i ≤ Suc n}" "Suc n"], assumption+, simp)
       apply (thin_tac "cmp f (transpos l (Suc n)) ` ({l. l ≤ Suc n} - 
            {Suc n}) = f ` transpos l (Suc n) ` ({l. l ≤ Suc n} - {Suc n})")
       apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in 
              transpos_surjec, simp, assumption+)
       apply (simp add:surj_to_def cmp_def)
    apply (simp add:transpos_ij_2)
done

(*
lemma (in Module) unique_expression3_2:"⟦H ⊆ carrier M; 
      f ∈ {k. k ≤ (Suc n)} → H; s ∈ {k. k ≤ (Suc n)} → carrier R; 
      l ≤ (Suc n); (f l) ∉ f ` ({k. k ≤ (Suc n)} - {l}); l ≠ Suc n⟧ ⟹ 
    ∃g m t. g ∈ {l. l ≤ (m::nat)} → H ∧ inj_on g {l. l ≤ (m::nat)} ∧ 
            t ∈ {l. l ≤ m} → carrier R ∧ 
            l_comb R M (Suc n) s f = l_comb R M m t g ∧ t m = s l"
apply (cut_tac sc_Ring,
       frule Ring.whole_ideal)
 apply (subst l_comb_transpos1[of "carrier R" H s n f l], assumption+,
        rule noteq_le_less[of l "Suc n"], assumption+) 
 apply (cut_tac unique_expression3_1[of H "cmp f (transpos l (Suc n))" n 
        "cmp s (transpos l (Suc n))"])
 apply ((erule exE)+, (erule conjE)+, simp)
 apply (subgoal_tac "t m = s l", blast)
 apply (thin_tac "l_comb R M (Suc n) (cmp s (transpos l (Suc n)))
         (cmp f (transpos l (Suc n))) = l_comb R M m t g")
 apply (simp add:cmp_def)
 apply (subst transpos_ij_2[of l "Suc n" "Suc n"], assumption+,
        simp, assumption, simp, assumption)
 apply (rule univar_func_test, rule ballI, simp add:cmp_def,
        frule_tac l = x in transpos_mem[of l "Suc n" "Suc n"], simp,
         assumption+, simp add:funcset_mem)
 apply (rule univar_func_test, rule ballI, simp add:cmp_def,
        frule_tac l = x in transpos_mem[of l "Suc n" "Suc n"], simp,
         assumption+, simp add:funcset_mem)
 apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in transpos_hom,
           simp, assumption)
 apply (frule cmp_fun_sub_image[of "transpos l (Suc n)" "{i. i ≤ Suc n}" 
       "{i. i ≤ Suc n}" f H "{l. l ≤ Suc n} - {Suc n}"], assumption+)
       apply (rule subsetI, simp)
       apply simp
       apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in transpos_inj,
              simp, assumption+)
       apply (subst injfun_elim_image[of "transpos l (Suc n)" "{i. i ≤ Suc n}"
        "{i. i ≤ Suc n}" "Suc n"], assumption+, simp)
       apply (thin_tac "cmp f (transpos l (Suc n)) ` ({l. l ≤ Suc n} - 
            {Suc n}) = f ` transpos l (Suc n) ` ({l. l ≤ Suc n} - {Suc n})")
       apply (frule_tac i = l and n = "Suc n" and j = "Suc n" in 
              transpos_surjec, simp, assumption+)
       apply (simp add:surj_to_def cmp_def)
    apply (simp add:transpos_ij_2)
done  *)

lemma (in Module) unique_expression3:
   "⟦H ⊆ carrier M; f ∈ {k. k ≤ (Suc n)} → H;
     s ∈ {k. k ≤ (Suc n)} → carrier R; l ≤ (Suc n);
    (f l) ∉ f ` ({k. k ≤ (Suc n)} - {l})⟧ ⟹ 
   ∃g m t. g ∈ {k. k ≤ (m::nat)} → H ∧ 
        inj_on g {k. k ≤ m} ∧ 
        t ∈ {k. k ≤ m} → carrier R ∧ 
        l_comb R M (Suc n) s f = l_comb R M m t g ∧ t m = s l ∧ g m = f l"
apply (case_tac "l = Suc n", simp)
 apply (cut_tac unique_expression3_1[of H f n s], blast,
        assumption+)
 apply (rule unique_expression3_2[of H f n s l], assumption+)
done

lemma (in Module) unique_expression4:"free_generator R M H ⟹
     f ∈ {k. k ≤ (n::nat)} → H ∧ inj_on f {k. k ≤ n} ∧ 
     s ∈ {k. k ≤ n} → carrier R ∧ l_comb R M n s f ≠ 𝟬  ⟶ 
(∃m g t. (g ∈ {k. k ≤ m} → H) ∧ inj_on g {k. k ≤ m} ∧ 
        (g ` {k. k ≤ m} ⊆ f ` {k. k ≤ n}) ∧ (t ∈ {k. k ≤ m} → carrier R) ∧
        (∀l ∈ {k. k ≤ m}. t l ≠ 𝟬R) ∧ l_comb R M n s f = l_comb R M m t g)"
apply (cut_tac sc_Ring)
apply (frule free_generator_sub[of H])
apply (induct_tac n)
 apply (rule impI, (erule conjE)+)
 apply (frule has_free_generator_nonzeroring[of H])
   apply (frule Ring.whole_ideal,
         frule_tac s = s and n = 0 and f = f in 
             l_comb_mem_linear_span[of "carrier R" H], assumption+)
   apply blast
 apply (simp add:l_comb_def)
 apply (subgoal_tac "f ∈ {j. j ≤ (0::nat)} → H ∧ 
        inj_on f {j. j ≤ 0} ∧ f ` {j. j ≤ 0} ⊆ f ` {0} ∧ 
        s ∈ {j. j ≤ 0} → carrier R ∧ (∀l ≤ 0. s l ≠ 𝟬R) ∧  
        s 0 ⋅s (f 0) = Σe M (λj. s j ⋅s (f j)) 0",
        (erule conjE)+, blast)
 apply simp
 apply (rule contrapos_pp, simp+)
 apply (cut_tac m = "f 0" in sc_0_m,
           simp add:Pi_def subsetD, simp)

apply (rule impI) apply (erule conjE)+
 apply (frule func_pre[of _ _ H],
        frule_tac f = f and A = "{k. k ≤ Suc n}" and ?A1.0 = "{k. k ≤ n}" in
        restrict_inj, rule subsetI, simp,
        frule func_pre[of _ _ "carrier R"], simp)
 apply (frule Ring.whole_ideal)
 apply (frule free_generator_sub[of H], 
         simp add:l_comb_Suc[of H "carrier R" s _ f])

 apply (case_tac "s (Suc n) = 𝟬R", simp)
       apply (frule_tac x = "Suc n" and f = f and A = "{k. k ≤ Suc n}" and
              B = H in funcset_mem, simp,
             frule_tac c = "f (Suc n)" in subsetD[of H "carrier M"], simp)
       apply (frule_tac m = "f (Suc n)" in sc_0_m, simp)
       apply (frule_tac n = n in l_comb_mem[of "carrier R" H s _ f],
               assumption+, simp add:ag_r_zero)
   apply ((erule exE)+, (erule conjE)+)
   apply (frule_tac f = f and A = "{k. k ≤ Suc n}" and B = H and 
          ?A1.0 = "{k. k ≤ n}" and ?A2.0 = "{k. k ≤ Suc n}" in im_set_mono,
          rule subsetI, simp, simp,
          frule_tac A = "g ` {k. k ≤ m}" and B = "f ` {k. k ≤ n}" and 
          C = "f ` {k. k ≤ Suc n}" in subset_trans, assumption+)
   apply blast

  apply (case_tac "l_comb R M n s f = 𝟬M", simp,
         frule_tac x = "Suc n" and f = s and A = "{k. k ≤ Suc n}" and 
            B = "carrier R" in funcset_mem, simp,
         frule_tac x = "Suc n" and f = f and A = "{k. k ≤ Suc n}" and 
         B = H in funcset_mem, simp,
         frule_tac c = "f (Suc n)" in subsetD[of H "carrier M"], assumption+,
         frule_tac a = "s (Suc n)" and m = "f (Suc n)" in sc_mem, assumption+,
         simp add:ag_l_zero)
  apply (subgoal_tac "(λj∈{0::nat}. f (Suc n)) ∈ {j. j ≤ (0::nat)} → H ∧ 
     inj_on (λj∈{0::nat}. f (Suc n)) {j. j ≤ (0::nat)} ∧ 
     (λj∈{0::nat}. f (Suc n)) ` {j. j ≤ (0::nat)} ⊆  f  ` {k. k ≤ (Suc n)} ∧ 
      (λj∈{0::nat}. s (Suc n))∈ {k. k ≤ 0} → carrier R ∧ 
      (∀l≤0. (λj∈{0::nat}. s (Suc n)) l ≠ 𝟬R) ∧
        s (Suc n) ⋅s f (Suc n) = 
           l_comb R M 0 (λj∈{0::nat}. s (Suc n)) (λj∈{0::nat}. f (Suc n))")
 apply ((erule conjE)+, blast) 
 apply simp
 apply (simp add:l_comb_def)
 
 apply simp
 apply ((erule exE)+, (erule conjE)+, erule exE, (erule conjE)+, simp)
 apply (thin_tac "l_comb R M m t g ≠ 𝟬",
        thin_tac "l_comb R M m t g ± s (Suc n) ⋅s f (Suc n) ≠ 𝟬",
        thin_tac "l_comb R M n s f = l_comb R M m t g")
 apply (frule_tac f = g and n = m and A = H and g = "λj∈{0::nat}. f (Suc n)"
        and m = 0 and B = H in jointfun_hom,
        rule Pi_I, simp add:Pi_def,
        frule_tac f = t and n = m and A = "carrier R" and 
         g = "λj∈{0::nat}. s (Suc n)" and m = 0 and B = "carrier R" in 
         jointfun_hom, simp add:Pi_def, simp)
 apply (subgoal_tac "inj_on (jointfun m g 0 (λj∈{0}. f (Suc n)))
    {k. k ≤ Suc m} ∧ 
 (jointfun m g 0 (λj∈{0}. f (Suc n))) ` {k. k ≤ Suc m} ⊆ f ` {k. k ≤ Suc n} ∧
 (∀l ≤ (Suc m). (jointfun m t 0 (λj∈{0}. s (Suc n))) l ≠ 𝟬R) ∧
 l_comb R M m t g ± s (Suc n) ⋅s f (Suc n) =
    l_comb R M (Suc m) (jointfun m t 0 (λj∈{0}. s (Suc n)))
                            (jointfun m g 0 (λj∈{0}. f (Suc n)))") 
 apply (erule conjE)+ apply blast

 apply (rule conjI) 
  apply (rule_tac f = g and n = m and b = "f (Suc n)" and B = H in 
         jointfun_inj, assumption+)
  apply (rule contrapos_pp, simp+)   
  apply (frule_tac c = "f (Suc n)" and A = "g ` {k. k ≤ m}" and 
       B = "f ` {k. k ≤ n}" in subsetD, assumption+)

  apply (thin_tac "inj_on f {k. k ≤ n}",
         thin_tac "g ` {k. k ≤ m} ⊆ f ` {k. k ≤ n}",
         thin_tac "f (Suc n) ∈ g ` {j. j ≤ m}", simp add:image_def,
         erule exE, erule conjE)
  apply (simp add:inj_on_def,
         drule_tac a = "Suc n" in forall_spec, simp,
         thin_tac "∀x≤m. ∀y≤m. g x = g y ⟶ x = y",
         thin_tac "∀l≤m. t l ≠ 𝟬R",
         drule_tac a = x in forall_spec, simp, simp)

  apply (rule conjI, rule subsetI)
  apply (simp add:image_def, erule exE, erule conjE) 
   apply (case_tac "xa = Suc m", simp add:jointfun_def sliden_def)
   apply (cut_tac n = "Suc n" in Nat.le_refl, blast)
   apply (frule_tac m = xa and n = "Suc m" in noteq_le_less, assumption,
            thin_tac "xa ≤ Suc m",
            frule_tac x = xa and n = "Suc m" in less_le_diff,
            thin_tac "xa < Suc m", simp,
      thin_tac "jointfun m g 0 (λj∈{0}. f (Suc n)) ∈ {j. j ≤ Suc m} → H",
  thin_tac "jointfun m t 0 (λj∈{0}. s (Suc n)) ∈ {j. j ≤ Suc m} → carrier R",
  simp add:jointfun_def)
  apply (subgoal_tac "g xa ∈ {y. ∃x≤n. y = f x}", simp, erule exE)
  apply (erule conjE, frule_tac i = xb and j = n and k = "Suc n" in
         le_trans, simp, blast)
  apply (rule_tac c = "g xa" and A = "{y. ∃x≤m. y = g x}" and 
         B = "{y. ∃x≤n. y = f x}" in subsetD, assumption+,
         simp, blast)
  apply (rule conjI, rule allI, rule impI)
  apply (case_tac "l = Suc m", simp add:jointfun_def sliden_def)
    apply (frule_tac m = l and n = "Suc m" in noteq_le_less, assumption,
            thin_tac "l ≤ Suc m",
            frule_tac x = l and n = "Suc m" in less_le_diff,
            thin_tac "l < Suc m", simp,
  thin_tac "jointfun m g 0 (λj∈{0}. f (Suc n)) ∈ {j. j ≤ Suc m} → H",
  thin_tac "jointfun m t 0 (λj∈{0}. s (Suc n)) ∈ {j. j ≤ Suc m} → carrier R",
   simp add:jointfun_def)  
  apply (simp add:l_comb_def,
        subst l_comb_jointfun_jj[of H "carrier R"], assumption+,
        simp add:Pi_def,
        simp add:Pi_def)
  apply (simp add:jointfun_def sliden_def)
done

lemma (in Module) unique_prepression5_0:"⟦free_generator R M H; 
       f ∈ {j. j ≤ n} → H; inj_on f {j. j ≤ n};
       s ∈ {j. j ≤ n} → carrier R; g ∈ {j. j ≤ m} → H; 
       inj_on g {j. j ≤ m}; t ∈ {j. j ≤ m} → carrier R; 
       l_comb R M n s f = l_comb R M m t g;∀j≤n. s j ≠ 𝟬R; ∀k≤m. t k ≠ 𝟬R;
       f n ∉ g ` {j. j ≤ m}; 0 < n⟧  ⟹ False" 
apply (cut_tac sc_Ring,
       frule Ring.ring_is_ag,
       frule Ring.whole_ideal,
       frule free_generator_sub[of H])
 apply (cut_tac l_comb_Suc[of H "carrier R" s "n - Suc 0" f],
         simp,
         thin_tac "l_comb R M n s f = l_comb R M (n - Suc 0) s f ± s n ⋅s f n")
  apply (frule free_generator_sub[of H],
         frule l_comb_mem[of "carrier R" H t m g], assumption+,
         frule l_comb_mem[of "carrier R" H s "n - Suc 0" f], assumption+,
         rule func_pre, simp, rule func_pre, simp,
         cut_tac sc_mem[of "s n" "f n"])
  apply (frule ag_pOp_closed[of "l_comb R M (n - Suc 0) s f" "s n ⋅s f n"],
          assumption+,
         frule ag_mOp_closed[of "l_comb R M (n - Suc 0) s f"])
  apply (frule ag_pOp_add_l[of "l_comb R M m t g" "l_comb R M (n - Suc 0) s f ± s n ⋅s f n" "-a (l_comb R M (n - Suc 0) s f)"], assumption+,
        thin_tac "l_comb R M m t g = l_comb R M (n - Suc 0) s f ± s n ⋅s f n")
  apply (simp add:ag_pOp_assoc[THEN sym, of "-a (l_comb R M (n - Suc 0) s f)"
         "l_comb R M (n - Suc 0) s f" "s n ⋅s f n"],
         simp add:ag_l_inv1 ag_l_zero)
  apply (cut_tac func_pre[of f "n - Suc 0" H],
         cut_tac func_pre[of s "n - Suc 0" "carrier R"])
  apply (frule linear_span_iOp_closedTr2[of "carrier R" "H" f "n - Suc 0" s],
         assumption+)
  apply (simp, 
          thin_tac "-a (l_comb R M (n - Suc 0) s f) =
         l_comb R M (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) f")
  apply (subgoal_tac "(λx∈{j. j ≤ n - Suc 0}. -aR (s x)) 
         ∈ {j. j ≤ n - Suc 0} → carrier R")
  apply (simp add:l_comb_add[THEN sym, of "carrier R" H
          "λx∈{j. j ≤ n - Suc 0}. -aR (s x)" "n - Suc 0" f t m g],
        thin_tac "l_comb R M m t g ∈ carrier M",
        thin_tac "l_comb R M (n - Suc 0) s f ∈ carrier M",
        thin_tac "l_comb R M (n - Suc 0) s f ± s n ⋅s f n ∈ carrier M",
        thin_tac "l_comb R M (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) f
         ∈ carrier M")
  apply (frule jointfun_hom[of f "n - Suc 0" H g m H], assumption+,
         frule jointfun_hom[of "λx∈{j. j ≤ n - Suc 0}. -aR (s x)" "n - Suc 0"
          "carrier R" t m "carrier R"], assumption+, simp)
 (* to apply unique_expression3_1, we show
     f n ∉ (jointfun (n - Suc 0) f m g) ` {j. j ≤ n + m} *)
 apply (frule im_jointfun[of f "n - Suc 0" H g m H], assumption+)
 apply (frule unique_expression3_1[of H 
  "jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0::nat}. (f n))"
  "n + m"
  "jointfun (n + m) (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) 
  m t) 0 (λx∈{0::nat}. -aR (s n))"])
 apply (rule Pi_I,
        case_tac "x ≤ (n + m)", simp,
        simp add:jointfun_def[of "n+m"], simp add:Pi_def,
        simp add:jointfun_def[of "n+m"] sliden_def, simp add:Pi_def)
  apply (rule Pi_I,
        case_tac "x ≤ (n + m)", simp,
        simp add:jointfun_def[of "n+m"], simp add:Pi_def)
  apply (simp add:jointfun_def[of "n+m"] sliden_def,
         frule Ring.ring_is_ag[of R], rule aGroup.ag_mOp_closed, assumption,
         simp add:Pi_def)
  apply (thin_tac "s ∈ {j. j ≤ n} → carrier R",
         thin_tac "t ∈ {j. j ≤ m} → carrier R",
         thin_tac "∀j≤n. s j ≠ 𝟬R",
         thin_tac "∀k≤m. t k ≠ 𝟬R",
         thin_tac "l_comb R M (n + m)
          (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t)
          (jointfun (n - Suc 0) f m g) =
         s n ⋅s f n",
         thin_tac "s ∈ {j. j ≤ n - Suc 0} → carrier R")
 apply (thin_tac "(λx∈{j. j ≤ n - Suc 0}. -aR (s x))
         ∈ {j. j ≤ n - Suc 0} → carrier R",
        thin_tac "jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t
         ∈ {j. j ≤ n + m} → carrier R")
 apply (simp add:Nset_pre1,
        simp add:im_jointfunTr1[of "n + m" "jointfun (n - Suc 0) f m g" 0 
        "λx∈{0}. f n"],
        thin_tac "jointfun (n - Suc 0) f m g ∈ {j. j ≤ n + m} → H",
        thin_tac "jointfun (n - Suc 0) f m g ` {j. j ≤ n + m} =
         f ` {j. j ≤ n - Suc 0} ∪ g ` {j. j ≤ m}",
        simp add:jointfun_def[of "n+m"] sliden_def)
 apply (rule contrapos_pp, simp+, simp add:image_def, erule exE,erule conjE,
        simp add:inj_on_def[of f],
        drule_tac a = n in forall_spec, simp,
        thin_tac "∀xa≤m. f x ≠ g xa",
        drule_tac a = x in forall_spec,
        rule_tac i = x and j = "n - Suc 0" and k = n in Nat.le_trans,
        assumption+, subst Suc_le_mono[THEN sym], simp,
        simp,
        cut_tac n1 = x and m1 = "x - Suc 0" in 
               Suc_le_mono[THEN sym], simp)

defer
 apply (rule Pi_I, simp,
        rule aGroup.ag_mOp_closed, assumption,
        cut_tac  i = x and j = "n - Suc 0" and k = n in Nat.le_trans,
        assumption, subst Suc_le_mono[THEN sym], simp,
        simp add:Pi_def, simp, simp, simp add:Pi_def,
        simp add:Pi_def,
        simp add:Pi_def subsetD, assumption+, simp, simp)
 apply ((erule exE)+, (erule conjE)+, erule exE, (erule conjE)+) 
 apply (cut_tac l_comb_Suc[of H "carrier R" "jointfun (n + m)
           (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
           (λx∈{0}. -aR (s n))" "n + m"
           "jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)"],
        simp) apply (
       thin_tac "l_comb R M (Suc (n + m))
         (jointfun (n + m)
           (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
           (λx∈{0}. -aR (s n)))
         (jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)) =
        l_comb R M ma ta ga")
 apply (subgoal_tac "l_comb R M (n + m)
         (jointfun (n + m)
           (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
           (λx∈{0}. -aR (s n)))
         (jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)) ±
        jointfun (n + m)
         (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
         (λx∈{0}. -aR (s n)) (Suc (n + m)) ⋅s
        jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)
         (Suc (n + m)) = 𝟬M", simp,
       thin_tac "l_comb R M (n + m)
         (jointfun (n + m)
           (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
           (λx∈{0}. -aR (s n)))
         (jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)) ±
        jointfun (n + m)
         (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
         (λx∈{0}. -aR (s n)) (Suc (n + m)) ⋅s
        jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)
         (Suc (n + m)) =
        l_comb R M ma ta ga",
       thin_tac "l_comb R M (n + m)
         (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t)
         (jointfun (n - Suc 0) f m g) =
        s n ⋅s f n",
       thin_tac "jointfun (n - Suc 0) f m g ∈ {j. j ≤ n + m} → H",
       thin_tac "jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t
        ∈ {j. j ≤ n + m} → carrier R",
       thin_tac "jointfun (n - Suc 0) f m g ` {j. j ≤ n + m} =
        f ` {j. j ≤ n - Suc 0} ∪ g ` {j. j ≤ m}")
    apply (simp add:jointfun_def[of "n+m"] sliden_def)
    apply (rotate_tac -3, frule sym, thin_tac "𝟬 = l_comb R M ma ta ga")
    apply (frule_tac s = ta and n = ma and m = ga in unique_expression1[of H],
           assumption+)
    apply (rotate_tac -1, 
           drule_tac x = ma in bspec, simp)
    apply (frule_tac funcset_mem[of s "{j. j ≤ n}" "carrier R" n], simp,
           frule sym, thin_tac "ta ma = -aR (s n)",
           frule aGroup.ag_inv_inv[of R "s n"], assumption+, simp,
           thin_tac " -aR (s n) = 𝟬R",
           rotate_tac -1, frule sym, thin_tac " -aR 𝟬R = s n",
           simp add:aGroup.ag_inv_zero[of R])

   apply (thin_tac "l_comb R M (n + m)
         (jointfun (n + m)
           (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
           (λx∈{0}. -aR (s n)))
         (jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)) ±
        jointfun (n + m)
         (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
         (λx∈{0}. -aR (s n)) (Suc (n + m)) ⋅s
        jointfun (n + m) (jointfun (n - Suc 0) f m g) 0 (λx∈{0}. f n)
         (Suc (n + m)) =
        l_comb R M ma ta ga",
        thin_tac "ta ma =
        jointfun (n + m)
         (jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) 0
         (λx∈{0}. -aR (s n)) (Suc (n + m))")
  apply (subst l_comb_jointfun_jj1[of H "carrier R"], assumption+,
         rule Pi_I, simp,
         rule aGroup.ag_mOp_closed, assumption, simp add:Pi_def,
         simp add:Pi_def)
  apply (simp,
        thin_tac "l_comb R M (n + m) (jointfun (n - Suc 0) 
       (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t) (jointfun (n - Suc 0) f m g) =
        s n ⋅s f n",
       thin_tac "jointfun (n - Suc 0) f m g ∈ {j. j ≤ n + m} → H",
       thin_tac "jointfun (n - Suc 0) (λx∈{j. j ≤ n - Suc 0}. -aR (s x)) m t
        ∈ {j. j ≤ n + m} → carrier R",
       thin_tac "jointfun (n - Suc 0) f m g ` {j. j ≤ n + m} =
        f ` {j. j ≤ n - Suc 0} ∪ g ` {j. j ≤ m}")
  apply (simp add:jointfun_def[of "n+m"] sliden_def,
         subst sc_minus_am1[THEN sym],
         simp add:Pi_def, simp add:Pi_def subsetD,
         simp add:ag_r_inv1,  simp add:free_generator_sub) 
  apply (assumption+,
         rule Pi_I,
         case_tac "x ≤ n + m", simp add:jointfun_def[of "n+m"],
         simp add:Pi_def,
         simp add:jointfun_def[of "n+m"] sliden_def,
         rule aGroup.ag_mOp_closed, assumption, simp add:Pi_def,
         rule Pi_I, simp,
          case_tac "x ≤ n+m", simp add:jointfun_def[of "n+m"],
          simp add:Pi_def, 
          simp add:jointfun_def[of "n+m"] sliden_def,
          simp add:Pi_def)
done
   
lemma (in Module) unique_expression5:"⟦free_generator R M H; 
      f ∈ {j. j ≤ (n::nat)} → H; inj_on f {j. j ≤ n}; 
      s ∈ {j. j ≤ n} → carrier R; g ∈ {j. j ≤ (m::nat)} → H; 
      inj_on g {j. j ≤ m}; t ∈ {j. j ≤ m} → carrier R; 
      l_comb R M n s f = l_comb R M m t g; 
     ∀j ∈ {j. j ≤ n}. s j ≠ 𝟬R; ∀k ∈ {j. j ≤ m}. t k ≠ 𝟬R⟧ ⟹
      f ` {j. j ≤ n} ⊆ g ` {j. j ≤ m}"
apply (cut_tac sc_Ring, frule Ring.ring_is_ag[of R],
       frule Ring.whole_ideal, 
       frule free_generator_sub[of H]) 
apply (rule contrapos_pp, simp+, simp add:subset_eq)
 apply (erule exE, erule conjE) 
 apply (case_tac "n = 0", simp)
  apply (frule_tac f = t and n = m and A = "carrier R" and 
        g = "λk∈{0::nat}. -aR (s 0)"  and m = 0 and B = "carrier R" in 
        jointfun_hom0,
        simp add:Pi_def,
        rule aGroup.ag_mOp_closed, assumption, simp add:Pi_def,
        frule_tac f = g and n = m and A = H and 
        g = "λk∈{0::nat}. (f 0)" and m = 0 and B = H in 
        jointfun_hom0,
        simp add:Pi_def subsetD,
        simp)
  apply (frule sym, thin_tac "l_comb R M 0 s f = l_comb R M m t g")
  apply (frule_tac n = 0 in l_comb_mem[of "carrier R" H s _ f],
         simp add:free_generator_sub, simp+,
         frule_tac n = m in l_comb_mem[of "carrier R" H t _ g],
         simp add:free_generator_sub, assumption+)
  apply (simp add:ag_eq_diffzero[of "l_comb R M m t g" "l_comb R M 0 s f"],
         simp add:l_comb_def[of R M 0 s f],
         frule free_generator_sub[of H],
          frule_tac c = "f 0" in subsetD[of H "carrier M"], assumption+,
          simp add:sc_minus_am1)
  apply (subgoal_tac "l_comb R M m t g ± (-aR (s 0)) ⋅s f 0 = 
          l_comb R M (Suc m) (jointfun m t 0 (λk∈{0}. (-aR (s 0))))
          (jointfun m g 0 (λk∈{0}. f 0))", simp)
  apply (frule_tac f = g and n = m and B = H and b = "f 0" in jointfun_inj,
          assumption+)
  apply (frule unique_expression1[of H "jointfun m t 0 (λk∈{0}. (-aR (s 0)))" 
        "Suc m" "jointfun m g 0 (λk∈{0}. f 0)"], assumption+)
 apply (frule_tac x = "Suc m" in bspec, simp,
        thin_tac "∀j∈{j. j ≤ Suc m}. jointfun m t 0 (λk∈{0}. -aR (s 0)) j 
          = 𝟬R")
  apply (simp add:jointfun_def sliden_def)
  apply (frule aGroup.ag_inv_inv[THEN sym, of R "s 0"], assumption,
         simp add:aGroup.ag_inv_zero)
        
  apply (thin_tac "l_comb R M m t g ± (-aR (s 0)) ⋅s f 0 = 𝟬",
         simp del:nsum_suc add:l_comb_def)
  apply (cut_tac l_comb_jointfun_jj[of H "carrier R" t m g "λk∈{0}. -aR (s 0)"
               0 "λk∈{0}. f 0"], simp,
         thin_tac e M (λj. jointfun m t 0 (λk∈{0}. -aR (s 0)) j ⋅s
                   jointfun m g 0 (λk∈{0}. f 0) j) m =
         Σe M (λj. t j ⋅s g j) m",
         simp add:jointfun_def sliden_def, simp add:free_generator_sub,
         assumption+,
         rule Pi_I, simp,
         rule aGroup.ag_mOp_closed, assumption+,
         simp)
 apply (case_tac "x = n", simp,
        rule unique_prepression5_0[of H f n s g m t], assumption+)
 apply (frule_tac j = x in l_comb_transpos1[of "carrier R" H s "n - Suc 0" f],
        rule subsetI, simp,
        simp+,
        rotate_tac -1, frule sym,
        thin_tac "l_comb R M m t g = 
        l_comb R M n (cmp s (transpos x n)) (cmp f (transpos x n))",
        frule_tac i = x and n = n and j = n in transpos_hom, simp,
           assumption,
        frule_tac i = x and n = n and j = n in transpos_inj, simp,
           assumption+,
        rule_tac f = "cmp f (transpos x n)" and s = "cmp s (transpos x n)" in 
        unique_prepression5_0[of H _ n _ g m t], assumption+,
        simp add:cmp_fun, simp add:cmp_fun, simp add:cmp_inj,
        simp add:cmp_fun, assumption+,
        rule allI, rule impI, simp add:cmp_def,
        frule_tac i = x and n = n and j = n and l = j in transpos_mem,
        simp, assumption+, blast, assumption)
  apply (simp add:cmp_def transpos_ij_2) 
  apply simp
done
 
lemma (in Module) unique_expression6:"⟦free_generator R M H;
      f ∈ {j. j ≤ (n::nat)} → H; inj_on f {j. j ≤ n}; 
      s ∈ {j. j ≤ n} → carrier R; 
      g ∈ {j. j ≤ (m::nat)} → H; inj_on g {j. j ≤ m}; 
      t ∈ {j. j ≤ m} → carrier R;
      l_comb R M n s f = l_comb R M m t g;
      ∀j∈{j. j ≤ n}. s j ≠ 𝟬R; ∀k∈ {j. j ≤ m}. t k ≠ 𝟬R⟧ ⟹ 
      f `{j. j ≤ n} = g `  {j. j ≤ m}"
apply (rule equalityI)
apply (rule_tac  H = H and f = f and n = n and s = s and g = g and m = m and 
       t = t in unique_expression5, assumption+)
apply (rule_tac  H = H and f = g and n = m and s = t and g = f and m = n and 
       t = s in unique_expression5, assumption+)
apply (rule sym, assumption, blast, blast)
done

lemma (in Module) unique_expression7_1:"⟦free_generator R M H; 
    f ∈ {j. j ≤ (n::nat)} → H; inj_on f {j. j ≤ n}; 
    s ∈ {j. j ≤ n} → carrier R; 
    g ∈ {j. j ≤ (m::nat)} → H; inj_on g {j. j ≤ m}; 
    t ∈ {j. j ≤ m} → carrier R; 
    l_comb R M n s f = l_comb R M m t g; 
   ∀j ∈ {j. j ≤ n}. s j ≠ 𝟬R; ∀k∈{j. j ≤ m}. t k ≠ 𝟬R⟧ ⟹ n = m"
apply (frule_tac A = "{j. j ≤ n}" and f = f in card_image,
       frule_tac A = "{j. j ≤ m}" and f = g in card_image)
apply (frule_tac H = H and f = f and n = n and s = s and g = g and t = t and 
       m = m in unique_expression6, assumption+)
apply (rotate_tac -3, frule sym, 
       thin_tac "card (f ` {j. j ≤ n}) = card ({j. j ≤ n})")
apply simp
done

lemma (in Module) unique_expression7_2:"⟦free_generator R M H;
      f ∈ {j. j ≤ (n::nat)} → H;  inj_on f {j. j ≤ n};
      s ∈ {j. j ≤ n} → carrier R; t ∈ {j. j ≤ n} → carrier R; 
      l_comb R M n s f = l_comb R M n t f⟧ ⟹ (∀l ∈ {j. j ≤ n}. s l = t l)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
 apply (frule free_generator_sub[of H])
 apply (frule l_comb_mem[of "carrier R" H s n f], assumption+,
        frule l_comb_mem[of "carrier R" H t n f], assumption+)
 apply (simp add:ag_eq_diffzero[of "l_comb R M n s f" "l_comb R M n t f"])
 apply (simp add:linear_span_iOp_closedTr2[of "carrier R" H f n t])
 apply (frule l_comb_add1[THEN sym, of "carrier R" H f n s "λj∈{k. k ≤ n}. -aR (t j)"],
            assumption+)
       apply (rule Pi_I)
       apply (simp, frule Ring.ring_is_ag[of R],
              rule aGroup.ag_mOp_closed[of R], simp add:Pi_def)
       apply (simp add:Pi_def)
       apply simp
 apply (frule_tac s = "λx∈{x. x ≤ n}. s x ±R (if x ≤ n then -aR (t x) else 
        undefined)" in unique_expression1[of H _ n f], assumption+)
  apply (rule Pi_I, simp)
  apply (frule Ring.ring_is_ag[of R], rule aGroup.ag_pOp_closed, assumption,
         simp add:Pi_def,
         rule aGroup.ag_mOp_closed, assumption,
         simp add:Pi_def, assumption+)
  apply (rule allI, rule impI)
  apply (subst aGroup.ag_eq_diffzero[of R],
         simp add:Ring.ring_is_ag,
         simp add:Pi_def, simp add:Pi_def)
 apply (drule_tac x = l in bspec, simp)
  apply simp
done

end