Library mathcomp.algebra.ssralg
(* (c) Copyright 2006-2016 Microsoft Corporation and Inria.
Distributed under the terms of CeCILL-B. *)
From HB Require Import structures.
From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq.
From mathcomp Require Import choice fintype finfun bigop prime binomial.
Distributed under the terms of CeCILL-B. *)
From HB Require Import structures.
From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq.
From mathcomp Require Import choice fintype finfun bigop prime binomial.
The base hierarchy of algebraic structures
NB: See CONTRIBUTING.md for an introduction to HB concepts and commands.
Reference: Francois Garillot, Georges Gonthier, Assia Mahboubi, Laurence
Rideau, Packaging mathematical structures, TPHOLs 2009
This file defines the following algebraic structures:
nmodType == additive abelian monoid
The HB class is called Nmodule.
zmodType == additive abelian group (Nmodule with an opposite)
The HB class is called Zmodule.
semiRingType == non-commutative semi rings
(NModule with a multiplication)
The HB class is called SemiRing.
comSemiRingType == commutative semi rings
The HB class is called ComSemiRing.
ringType == non-commutative rings (semi rings with an opposite)
The HB class is called Ring.
comRingType == commutative rings
The HB class is called ComRing.
lmodType R == module with left multiplication by external scalars
in the ring R
The HB class is called Lmodule.
lalgType R == left algebra, ring with scaling
that associates on the left
The HB class is called Lalgebra.
algType R == ring with scaling that associates both left and right
The HB class is called Algebra.
comAlgType R == commutative algType
The HB class is called ComAlgebra.
unitRingType == Rings whose units have computable inverses
The HB class is called UnitRing.
comUnitRingType == commutative UnitRing
The HB class is called ComUnitRing.
unitAlgType R == algebra with computable inverses
The HB class is called UnitAlgebra.
comUnitAlgType R == commutative UnitAlgebra
The HB class is called ComUnitAlgebra.
idomainType == integral, commutative, ring with partial inverses
The HB class is called IntegralDomain.
fieldType == commutative fields
The HB class is called Field.
decFieldType == fields with a decidable first order theory
The HB class is called DecidableField.
closedFieldType == algebraically closed fields
The HB class is called ClosedField.
and their joins with subType:
subNmodType V P == join of nmodType and subType (P : pred V) such
that val is semi_additive
The HB class is called SubNmodule.
subZmodType V P == join of zmodType and subType (P : pred V)
such that val is additive
The HB class is called SubZmodule.
subSemiRingType R P == join of semiRingType and subType (P : pred R)
such that val is a semiring morphism
The HB class is called SubSemiRing.
subComSemiRingType R P == join of comSemiRingType and subType (P : pred R)
such that val is a morphism
The HB class is called SubComSemiRing.
subRingType R P == join of ringType and subType (P : pred R)
such that val is a morphism
The HB class is called SubRing.
subComRingType R P == join of comRingType and subType (P : pred R)
such that val is a morphism
The HB class is called SubComRing.
subLmodType R V P == join of lmodType and subType (P : pred V)
such that val is scalable
The HB class is called SubLmodule.
subLalgType R V P == join of lalgType and subType (P : pred V)
such that val is linear
The HB class is called SubLalgebra.
subAlgType R V P == join of algType and subType (P : pred V)
such that val is linear
The HB class is called SubAlgebra.
subUnitRingType R P == join of unitRingType and subType (P : pred R)
such that val is a ring morphism
The HB class is called SubUnitRing.
subComUnitRingType R P == join of comUnitRingType and subType (P : pred R)
such that val is a ring morphism
The HB class is called SubComUnitRing.
subIdomainType R P == join of idomainType and subType (P : pred R)
such that val is a ring morphism
The HB class is called SubIntegralDomain.
subField R P == join of fieldType and subType (P : pred R)
such that val is a ring morphism
The HB class is called SubField.
Morphisms between the above structures (see below for details):
{additive U -> V} == semi additive (resp. additive) functions between
nmodType (resp. zmodType) instances U and V.
The HB class is called Additive.
{rmorphism R -> S} == semi ring (resp. ring) morphism between
semiRingType (resp. ringType) instances R and S.
The HB class is called RMorphism.
{linear U -> V | s} == linear functions of type U -> V, where U is a
left module over a ring R, V is a Z-module, and
s is a scaling operator of type R -> V -> V.
The HB class is called Linear.
{lrmorphism A -> B | s} == linear ring morphisms of type A -> B, where A
is a left algebra over a ring R, B is a ring,
and s is a scaling operator of type R -> B -> B.
The HB class is called LRMorphism.
- > The scaling operator s above should be one of *:%R, *%R, or a combination nu \; *:%R or nu \; *%R with a ring morphism nu; otherwise some of the theory (e.g., the linearZ rule) will not apply. To enable the overloading of the scaling operator, we use the following structures: GRing.Scale.law R V == scaling morphisms of type R -> V -> V The HB class is called Scale.Law.
Nmodule (additive abelian monoids):
0 == the zero (additive identity) of a Nmodule x + y == the sum of x and y (in a Nmodule) x *+ n == n times x, with n in nat (non-negative), i.e., x + (x + .. (x + x)..) (n terms); x *+ 1 is thus convertible to x, and x *+ 2 to x + x \sum_<range> e == iterated sum for a Zmodule (cf bigop.v) e`_i == nth 0 e i, when e : seq M and M has a zmodType structure support f == 0.-support f, i.e., [pred x | f x != 0] addr_closed S <-> collective predicate S is closed under finite sums (0 and x + y in S, for x, y in S) [SubChoice_isSubNmodule of U by <: ] == nmodType mixin for a subType whose base type is a nmodType and whose predicate's is a nmodClosedZmodule (additive abelian groups):
- x == the opposite (additive inverse) of x
SemiRing (non-commutative semirings):
R^c == the converse (semi)ring for R: R^c is convertible to R but when R has a canonical (semi)RingType structure R^c has the converse one: if x y : R^c, then x * y = (y : R) * (x : R) 1 == the multiplicative identity element of a semiring n%:R == the semiring image of an n in nat; this is just notation for 1 *+ n, so 1%:R is convertible to 1 and 2%:R to 1 + 1 <number> == <number>%:R with <number> a sequence of digits x * y == the semiring product of x and y \prod_<range> e == iterated product for a semiring (cf bigop.v) x ^+ n == x to the nth power with n in nat (non-negative), i.e., x * (x * .. (x * x)..) (n factors); x ^+ 1 is thus convertible to x, and x ^+ 2 to x * x GRing.comm x y <-> x and y commute, i.e., x * y = y * x GRing.lreg x <-> x if left-regular, i.e., *%R x is injective GRing.rreg x <-> x if right-regular, i.e., *%R^~ x is injective [char R] == the characteristic of R, defined as the set of prime numbers p such that p%:R = 0 in R The set [char R] has at most one element, and is implemented as a pred_nat collective predicate (see prime.v); thus the statement p \in [char R] can be read as `R has characteristic p', while [char R] =i pred0 means `R has characteristic 0' when R is a field. Frobenius_aut chRp == the Frobenius automorphism mapping x in R to x ^+ p, where chRp : p \in [char R] is a proof that R has (non-zero) characteristic p mulr_closed S <-> collective predicate S is closed under finite products (1 and x * y in S for x, y in S) semiring_closed S <-> collective predicate S is closed under semiring operations (0, 1, x + y and x * y in S) [SubNmodule_isSubSemiRing of R by <: ] == [SubChoice_isSubSemiRing of R by <: ] == semiRingType mixin for a subType whose base type is a semiRingType and whose predicate's is a semiringClosedRing (non-commutative rings):
GRing.sign R b := (-1) ^+ b in R : ringType, with b : bool This is a parsing-only helper notation, to be used for defining more specific instances. smulr_closed S <-> collective predicate S is closed under products and opposite (-1 and x * y in S for x, y in S) subring_closed S <-> collective predicate S is closed under ring operations (1, x - y and x * y in S) [SubZmodule_isSubRing of R by <: ] == [SubChoice_isSubRing of R by <: ] == ringType mixin for a subType whose base type is a ringType and whose predicate's is a subringClosedComSemiRing (commutative SemiRings):
[SubNmodule_isSubComSemiRing of R by <: ] == [SubChoice_isSubComSemiRing of R by <: ] == comSemiRingType mixin for a subType whose base type is a comSemiRingType and whose predicate's is a semiringClosedComRing (commutative Rings):
[SubZmodule_isSubComRing of R by <: ] == [SubChoice_isSubComRing of R by <: ] == comRingType mixin for a subType whose base type is a comRingType and whose predicate's is a subringClosedUnitRing (Rings whose units have computable inverses):
x \is a GRing.unit <=> x is a unit (i.e., has an inverse) x^-1 == the ring inverse of x, if x is a unit, else x x / y == x divided by y (notation for x * y^-1) x ^- n := notation for (x ^+ n)^-1, the inverse of x ^+ n invr_closed S <-> collective predicate S is closed under inverse divr_closed S <-> collective predicate S is closed under division (1 and x / y in S) sdivr_closed S <-> collective predicate S is closed under division and opposite (-1 and x / y in S, for x, y in S) divring_closed S <-> collective predicate S is closed under unitRing operations (1, x - y and x / y in S) [SubRing_isSubUnitRing of R by <: ] == [SubChoice_isSubUnitRing of R by <: ] == unitRingType mixin for a subType whose base type is a unitRingType and whose predicate's is a divringClosed and whose ring structure is compatible with the base type'sComUnitRing (commutative rings with computable inverses):
[SubChoice_isSubComUnitRing of R by <: ] == comUnitRingType mixin for a subType whose base type is a comUnitRingType and whose predicate's is a divringClosed and whose ring structure is compatible with the base type'sIntegralDomain (integral, commutative, ring with partial inverses):
[SubComUnitRing_isSubIntegralDomain R by <: ] == [SubChoice_isSubIntegralDomain R by <: ] == mixin axiom for a idomain subTypeField (commutative fields):
GRing.Field.axiom inv == field axiom: x != 0 -> inv x * x = 1 for all x This is equivalent to the property above, but does not require a unitRingType as inv is an explicit argument. [SubIntegralDomain_isSubField of R by <: ] == mixin axiom for a field subTypeDecidableField (fields with a decidable first order theory):
GRing.term R == the type of formal expressions in a unit ring R with formal variables 'X_k, k : nat, and manifest constants x%:T, x : R The notation of all the ring operations is redefined for terms, in scope %T. GRing.formula R == the type of first order formulas over R; the %T scope binds the logical connectives /\, \/, ~, ==>, ==, and != to formulae; GRing.True/False and GRing.Bool b denote constant formulae, and quantifiers are written 'forall/'exists 'X_k, f GRing.Unit x tests for ring units GRing.If p_f t_f e_f emulates if-then-else GRing.Pick p_f t_f e_f emulates fintype.pick foldr GRing.Exists/Forall q_f xs can be used to write iterated quantifiers GRing.eval e t == the value of term t with valuation e : seq R (e maps 'X_i to e`_i) GRing.same_env e1 e2 <-> environments e1 and e2 are extensionally equal GRing.qf_form f == f is quantifier-free GRing.holds e f == the intuitionistic CiC interpretation of the formula f holds with valuation e GRing.qf_eval e f == the value (in bool) of a quantifier-free f GRing.sat e f == valuation e satisfies f (only in a decField) GRing.sol n f == a sequence e of size n such that e satisfies f, if one exists, or [:: ] if there is no such e 'exists 'X_i, u1 == 0 /\ ... /\ u_m == 0 /\ v1 != 0 ... /\ v_n != 0Lmodule (module with left multiplication by external scalars).
a *: v == v scaled by a, when v is in an Lmodule V and a is in the scalar Ring of V scaler_closed S <-> collective predicate S is closed under scaling linear_closed S <-> collective predicate S is closed under linear combinations (a *: u + v in S when u, v in S) submod_closed S <-> collective predicate S is closed under lmodType operations (0 and a *: u + v in S) [SubZmodule_isSubLmodule of V by <: ] == [SubChoice_isSubLmodule of V by <: ] == mixin axiom for a subType of an lmodTypeLalgebra (left algebra, ring with scaling that associates on the left):
R^o == the regular algebra of R: R^o is convertible to R, but when R has a ringType structure then R^o extends it to an lalgType structure by letting R act on itself: if x : R and y : R^o then x *: y = x * (y : R) k%:A == the image of the scalar k in an L-algebra; this is simply notation for k *: 1 subalg_closed S <-> collective predicate S is closed under lalgType operations (1, a *: u + v and u * v in S) [SubRing_SubLmodule_isSubLalgebra of V by <: ] == [SubChoice_isSubLalgebra of V by <: ] == mixin axiom for a subType of an lalgTypeAlgebra (ring with scaling that associates both left and right):
[SubLalgebra_isSubAlgebra of V by <: ] == [SubChoice_isSubAlgebra of V by <: ] == mixin axiom for a subType of an algTypeUnitAlgebra (algebra with computable inverses):
divalg_closed S <-> collective predicate S is closed under all unitAlgType operations (1, a *: u + v and u / v are in S fo u, v in S) In addition to this structure hierarchy, we also develop a separate, parallel hierarchy for morphisms linking these structures:Additive (semi additive or additive functions):
semi_additive f <-> f of type U -> V is semi additive, i.e., f maps the Nmodule structure of U to that of V, 0 to 0 and + to + := (f 0 = 0) * {morph f : x y / x + y} additive f <-> f of type U -> V is additive, i.e., f maps the Zmodule structure of U to that of V, 0 to 0,- to - and + to + (equivalently, binary - to -)
RMorphism (semiring or ring morphisms):
multiplicative f <-> f of type R -> S is multiplicative, i.e., f maps 1 and * in R to 1 and * in S, respectively R and S must have canonical semiRingType instances {rmorphism R -> S} == the interface type for semiring morphisms; both R and S must have semiRingType instances When both R and S have ringType instances, it is a ring morphism. := GRing.RMorphism.type R S- > If R and S are UnitRings the f also maps units to units and inverses of units to inverses; if R is a field then f is a field isomorphism between R and its image.
- > Additive properties (raddf_suffix, see below) are duplicated and specialised for RMorphism (as rmorph_suffix). This allows more precise rewriting and cleaner chaining: although raddf lemmas will recognize RMorphism functions, the converse will not hold (we cannot add reverse inheritance rules because of incomplete backtracking in the Canonical Projection unification), so one would have to insert a /= every time one switched from additive to multiplicative rules.
Linear (linear functions):
scalable_for s f <-> f of type U -> V is scalable for the scaling operator s of type R -> V -> V, i.e., f morphs a *: _ to s a _; R, U, and V must be a ringType, an lmodType R, and a zmodType, respectively. := forall a, {morph f : u / a *: u >-> s a u} scalable f <-> f of type U -> V is scalable, i.e., f morphs scaling on U to scaling on V, a *: _ to a *: _; U and V must be an lmodType R for the same ringType R. := scalable_for *:%R f linear_for s f <-> f of type U -> V is linear for s of type R -> V -> V, i.e., f (a *: u + v) = s a (f u) + f v; R, U, and V must be a ringType, an lmodType R, and a zmodType, respectively. linear f <-> f of type U -> V is linear, i.e., f (f *: u + v) = a *: f u + f v; U and V must be lmodType R for the same ringType R. := linear_for *:%R f scalar f <-> f of type U -> R is a scalar function, i.e., f (a *: u + v) = a * f u + f v; R and U must be a ringType and an lmodType R, respectively. := linear_for *%R f {linear U -> V | s} == the interface type for functions linear for the scaling operator s of type R -> V -> V, i.e., a structure that encapsulates two properties semi_additive f and scalable_for s f for functions f : U -> V; R, U, and V must be a ringType, an lmodType R, and a zmodType, respectively := @GRing.Linear.type R U V s {linear U -> V} == the interface type for linear functions, of type U -> V where both U and V must be lmodType R for the same ringType R := {linear U -> V | *:%R} {scalar U} == the interface type for scalar functions of type U -> R where U must be an lmodType R := {linear U -> V | *%R} (a *: u)%Rlin == transient forms that simplify to a *: u, a * u, (a * u)%Rlin nu a *: u, and nu a * u, respectively, and are (a *:^nu u)%Rlin created by rewriting with the linearZ lemma (a *^nu u)%Rlin The forms allows the RHS of linearZ to be matched reliably, using the GRing.Scale.law structure.- > Similarly to Ring morphisms, additive properties are specialized for linear functions.
- > Although {scalar U} is convertible to {linear U -> R^o}, it does not actually use R^o, so that rewriting preserves the canonical structure of the range of scalar functions.
- > The generic linearZ lemma uses a set of bespoke interface structures to ensure that both left-to-right and right-to-left rewriting work even in the presence of scaling functions that simplify non-trivially (e.g., idfun \; *%R). Because most of the canonical instances and projections are coercions the machinery will be mostly invisible (with only the {linear ...} structure and %Rlin notations showing), but users should beware that in (a *: f u)%Rlin, a actually occurs in the f u subterm.
- > The simpler linear_LR, or more specialized linearZZ and scalarZ rules should be used instead of linearZ if there are complexity issues, as well as for explicit forward and backward application, as the main parameter of linearZ is a proper sub-interface of {linear fUV | s}.
LRMorphism (linear ring morphisms, i.e., algebra morphisms):
{lrmorphism A -> B | s} == the interface type for ring morphisms linear for the scaling operator s of type R -> B -> B, i.e., the join of ring morphisms {rmorphism A -> B} and linear functions {linear A -> B | s}; R, A, and B must be a ringType, an lalgType R, and a ringType, respectively := @GRing.LRMorphism.type R A B s {lrmorphism A -> B} == the interface type for algebra morphisms, where A and B must be lalgType R for the same ringType R := {lrmorphism A -> B | *:%R}- > Linear and rmorphism properties do not need to be specialized for as we supply inheritance join instances in both directions.
Set Implicit Arguments.
Declare Scope ring_scope.
Declare Scope term_scope.
Declare Scope linear_ring_scope.
Reserved Notation "+%R" (at level 0).
Reserved Notation "-%R" (at level 0).
Reserved Notation "*%R" (at level 0, format " *%R").
Reserved Notation "*:%R" (at level 0, format " *:%R").
Reserved Notation "n %:R" (at level 2, left associativity, format "n %:R").
Reserved Notation "k %:A" (at level 2, left associativity, format "k %:A").
Reserved Notation "[ 'char' F ]" (at level 0, format "[ 'char' F ]").
Reserved Notation "x %:T" (at level 2, left associativity, format "x %:T").
Reserved Notation "''X_' i" (at level 8, i at level 2, format "''X_' i").
Patch for recurring Coq parser bug: Coq seg faults when a level 200
notation is used as a pattern.
Reserved Notation "''exists' ''X_' i , f"
(at level 199, i at level 2, right associativity,
format "'[hv' ''exists' ''X_' i , '/ ' f ']'").
Reserved Notation "''forall' ''X_' i , f"
(at level 199, i at level 2, right associativity,
format "'[hv' ''forall' ''X_' i , '/ ' f ']'").
Reserved Notation "x ^f" (at level 2, left associativity, format "x ^f").
Reserved Notation "\0" (at level 0).
Reserved Notation "f \+ g" (at level 50, left associativity).
Reserved Notation "f \- g" (at level 50, left associativity).
Reserved Notation "\- f" (at level 35, f at level 35).
Reserved Notation "a \*o f" (at level 40).
Reserved Notation "a \o* f" (at level 40).
Reserved Notation "a \*: f" (at level 40).
Reserved Notation "f \* g" (at level 40, left associativity).
Reserved Notation "'{' 'additive' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'additive' U -> V }").
Reserved Notation "'{' 'rmorphism' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'rmorphism' U -> V }").
Reserved Notation "'{' 'lrmorphism' U '->' V '|' s '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'lrmorphism' U -> V | s }").
Reserved Notation "'{' 'lrmorphism' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'lrmorphism' U -> V }").
Reserved Notation "'{' 'linear' U '->' V '|' s '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'linear' U -> V | s }").
Reserved Notation "'{' 'linear' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'linear' U -> V }").
Declare Scope ring_scope.
Delimit Scope ring_scope with R.
Declare Scope term_scope.
Delimit Scope term_scope with T.
Local Open Scope ring_scope.
Module Import GRing.
Import Monoid.Theory.
#[short(type="nmodType")]
HB.structure Definition Nmodule := {V of isNmodule V & Choice V}.
Module NmodExports.
Bind Scope ring_scope with Nmodule.sort.
End NmodExports.
Local Notation "0" := (@zero _) : ring_scope.
Local Notation "+%R" := (@add _) : function_scope.
Local Notation "x + y" := (add x y) : ring_scope.
Definition natmul V x n := iterop n +%R x (@zero V).
Arguments natmul : simpl never.
Local Notation "x *+ n" := (natmul x n) : ring_scope.
Local Notation "\sum_ ( i <- r | P ) F" := (\big[+%R/0]_(i <- r | P) F).
Local Notation "\sum_ ( m <= i < n ) F" := (\big[+%R/0]_(m ≤ i < n) F).
Local Notation "\sum_ ( i < n ) F" := (\big[+%R/0]_(i < n) F).
Local Notation "\sum_ ( i 'in' A ) F" := (\big[+%R/0]_(i in A) F).
Local Notation "s `_ i" := (nth 0 s i) : ring_scope.
Section NmoduleTheory.
Variable V : nmodType.
Implicit Types x y : V.
Lemma addr0 : @right_id V V 0 +%R.
#[export]
HB.instance Definition _ := Monoid.isComLaw.Build V 0 +%R addrA addrC add0r.
Lemma addrCA : @left_commutative V V +%R.
Lemma addrAC : @right_commutative V V +%R.
Lemma addrACA : @interchange V +%R +%R.
Lemma mulr0n x : x *+ 0 = 0.
Lemma mulr1n x : x *+ 1 = x.
Lemma mulr2n x : x *+ 2 = x + x.
Lemma mulrS x n : x *+ n.+1 = x + x *+ n.
Lemma mulrSr x n : x *+ n.+1 = x *+ n + x.
Lemma mulrb x (b : bool) : x *+ b = (if b then x else 0).
Lemma mul0rn n : 0 *+ n = 0 :> V.
Lemma mulrnDl n : {morph (fun x ⇒ x *+ n) : x y / x + y}.
Lemma mulrnDr x m n : x *+ (m + n) = x *+ m + x *+ n.
Lemma mulrnA x m n : x *+ (m × n) = x *+ m *+ n.
Lemma mulrnAC x m n : x *+ m *+ n = x *+ n *+ m.
Lemma iter_addr n x y : iter n (+%R x) y = x *+ n + y.
Lemma iter_addr_0 n x : iter n (+%R x) 0 = x *+ n.
Lemma sumrMnl I r P (F : I → V) n :
\sum_(i <- r | P i) F i *+ n = (\sum_(i <- r | P i) F i) *+ n.
Lemma sumrMnr x I r P (F : I → nat) :
\sum_(i <- r | P i) x *+ F i = x *+ (\sum_(i <- r | P i) F i).
Lemma sumr_const (I : finType) (A : pred I) x : \sum_(i in A) x = x *+ #|A|.
Lemma sumr_const_nat m n x : \sum_(n ≤ i < m) x = x *+ (m - n).
Definition addr_closed (S : {pred V}) :=
0 \in S ∧ {in S &, ∀ u v, u + v \in S}.
End NmoduleTheory.
#[short(type="zmodType")]
HB.structure Definition Zmodule := {V of Nmodule_isZmodule V & Nmodule V}.
Module ZmodExports.
Bind Scope ring_scope with Zmodule.sort.
End ZmodExports.
Local Notation "-%R" := (@opp _) : ring_scope.
Local Notation "- x" := (opp x) : ring_scope.
Local Notation "x - y" := (x + - y) : ring_scope.
Local Notation "x *- n" := (- (x *+ n)) : ring_scope.
Section ZmoduleTheory.
Variable V : zmodType.
Implicit Types x y : V.
Lemma addrN : @right_inverse V V V 0 -%R +%R.
Definition subrr := addrN.
Lemma addKr : @left_loop V V -%R +%R.
Lemma addNKr : @rev_left_loop V V -%R +%R.
Lemma addrK : @right_loop V V -%R +%R.
Lemma addrNK : @rev_right_loop V V -%R +%R.
Definition subrK := addrNK.
Lemma subKr x : involutive (fun y ⇒ x - y).
Lemma addrI : @right_injective V V V +%R.
Lemma addIr : @left_injective V V V +%R.
Lemma subrI : right_injective (fun x y ⇒ x - y).
Lemma subIr : left_injective (fun x y ⇒ x - y).
Lemma opprK : @involutive V -%R.
Lemma oppr_inj : @injective V V -%R.
Lemma oppr0 : -0 = 0 :> V.
Lemma oppr_eq0 x : (- x == 0) = (x == 0).
Lemma subr0 x : x - 0 = x.
Lemma sub0r x : 0 - x = - x.
Lemma opprB x y : - (x - y) = y - x.
Lemma opprD : {morph -%R: x y / x + y : V}.
Lemma addrKA z x y : (x + z) - (z + y) = x - y.
Lemma subrKA z x y : (x - z) + (z + y) = x + y.
Lemma addr0_eq x y : x + y = 0 → - x = y.
Lemma subr0_eq x y : x - y = 0 → x = y.
Lemma subr_eq x y z : (x - z == y) = (x == y + z).
Lemma subr_eq0 x y : (x - y == 0) = (x == y).
Lemma addr_eq0 x y : (x + y == 0) = (x == - y).
Lemma eqr_opp x y : (- x == - y) = (x == y).
Lemma eqr_oppLR x y : (- x == y) = (x == - y).
Lemma mulNrn x n : (- x) *+ n = x *- n.
Lemma mulrnBl n : {morph (fun x ⇒ x *+ n) : x y / x - y}.
Lemma mulrnBr x m n : n ≤ m → x *+ (m - n) = x *+ m - x *+ n.
Lemma sumrN I r P (F : I → V) :
(\sum_(i <- r | P i) - F i = - (\sum_(i <- r | P i) F i)).
Lemma sumrB I r (P : pred I) (F1 F2 : I → V) :
\sum_(i <- r | P i) (F1 i - F2 i)
= \sum_(i <- r | P i) F1 i - \sum_(i <- r | P i) F2 i.
Lemma telescope_sumr n m (f : nat → V) : n ≤ m →
\sum_(n ≤ k < m) (f k.+1 - f k) = f m - f n.
Lemma telescope_sumr_eq n m (f u : nat → V) : n ≤ m →
(∀ k, (n ≤ k < m)%N → u k = f k.+1 - f k) →
\sum_(n ≤ k < m) u k = f m - f n.
Section ClosedPredicates.
Variable S : {pred V}.
Definition oppr_closed := {in S, ∀ u, - u \in S}.
Definition subr_2closed := {in S &, ∀ u v, u - v \in S}.
Definition zmod_closed := 0 \in S ∧ subr_2closed.
Lemma zmod_closedN : zmod_closed → oppr_closed.
Lemma zmod_closedD : zmod_closed → addr_closed S.
End ClosedPredicates.
End ZmoduleTheory.
Arguments addrI {V} y [x1 x2].
Arguments addIr {V} x [x1 x2].
Arguments opprK {V}.
Arguments oppr_inj {V} [x1 x2].
Arguments telescope_sumr_eq {V n m} f u.
#[short(type="semiRingType")]
HB.structure Definition SemiRing := { R of Nmodule_isSemiRing R & Nmodule R }.
Module SemiRingExports.
Bind Scope ring_scope with SemiRing.sort.
End SemiRingExports.
Definition exp R x n := iterop n (@mul R) x (@one R).
Arguments exp : simpl never.
Definition comm R x y := @mul R x y = mul y x.
Definition lreg R x := injective (@mul R x).
Definition rreg R x := injective ((@mul R)^~ x).
Local Notation "1" := (@one _) : ring_scope.
Local Notation "n %:R" := (1 *+ n) : ring_scope.
Local Notation "*%R" := (@mul _) : function_scope.
Local Notation "x * y" := (mul x y) : ring_scope.
Local Notation "x ^+ n" := (exp x n) : ring_scope.
Local Notation "\prod_ ( i <- r | P ) F" := (\big[*%R/1]_(i <- r | P) F).
Local Notation "\prod_ ( i | P ) F" := (\big[*%R/1]_(i | P) F).
Local Notation "\prod_ ( i 'in' A ) F" := (\big[*%R/1]_(i in A) F).
Local Notation "\prod_ ( m <= i < n ) F" := (\big[*%R/1%R]_(m ≤ i < n) F%R).
(at level 199, i at level 2, right associativity,
format "'[hv' ''exists' ''X_' i , '/ ' f ']'").
Reserved Notation "''forall' ''X_' i , f"
(at level 199, i at level 2, right associativity,
format "'[hv' ''forall' ''X_' i , '/ ' f ']'").
Reserved Notation "x ^f" (at level 2, left associativity, format "x ^f").
Reserved Notation "\0" (at level 0).
Reserved Notation "f \+ g" (at level 50, left associativity).
Reserved Notation "f \- g" (at level 50, left associativity).
Reserved Notation "\- f" (at level 35, f at level 35).
Reserved Notation "a \*o f" (at level 40).
Reserved Notation "a \o* f" (at level 40).
Reserved Notation "a \*: f" (at level 40).
Reserved Notation "f \* g" (at level 40, left associativity).
Reserved Notation "'{' 'additive' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'additive' U -> V }").
Reserved Notation "'{' 'rmorphism' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'rmorphism' U -> V }").
Reserved Notation "'{' 'lrmorphism' U '->' V '|' s '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'lrmorphism' U -> V | s }").
Reserved Notation "'{' 'lrmorphism' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'lrmorphism' U -> V }").
Reserved Notation "'{' 'linear' U '->' V '|' s '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'linear' U -> V | s }").
Reserved Notation "'{' 'linear' U '->' V '}'"
(at level 0, U at level 98, V at level 99,
format "{ 'linear' U -> V }").
Declare Scope ring_scope.
Delimit Scope ring_scope with R.
Declare Scope term_scope.
Delimit Scope term_scope with T.
Local Open Scope ring_scope.
Module Import GRing.
Import Monoid.Theory.
#[short(type="nmodType")]
HB.structure Definition Nmodule := {V of isNmodule V & Choice V}.
Module NmodExports.
Bind Scope ring_scope with Nmodule.sort.
End NmodExports.
Local Notation "0" := (@zero _) : ring_scope.
Local Notation "+%R" := (@add _) : function_scope.
Local Notation "x + y" := (add x y) : ring_scope.
Definition natmul V x n := iterop n +%R x (@zero V).
Arguments natmul : simpl never.
Local Notation "x *+ n" := (natmul x n) : ring_scope.
Local Notation "\sum_ ( i <- r | P ) F" := (\big[+%R/0]_(i <- r | P) F).
Local Notation "\sum_ ( m <= i < n ) F" := (\big[+%R/0]_(m ≤ i < n) F).
Local Notation "\sum_ ( i < n ) F" := (\big[+%R/0]_(i < n) F).
Local Notation "\sum_ ( i 'in' A ) F" := (\big[+%R/0]_(i in A) F).
Local Notation "s `_ i" := (nth 0 s i) : ring_scope.
Section NmoduleTheory.
Variable V : nmodType.
Implicit Types x y : V.
Lemma addr0 : @right_id V V 0 +%R.
#[export]
HB.instance Definition _ := Monoid.isComLaw.Build V 0 +%R addrA addrC add0r.
Lemma addrCA : @left_commutative V V +%R.
Lemma addrAC : @right_commutative V V +%R.
Lemma addrACA : @interchange V +%R +%R.
Lemma mulr0n x : x *+ 0 = 0.
Lemma mulr1n x : x *+ 1 = x.
Lemma mulr2n x : x *+ 2 = x + x.
Lemma mulrS x n : x *+ n.+1 = x + x *+ n.
Lemma mulrSr x n : x *+ n.+1 = x *+ n + x.
Lemma mulrb x (b : bool) : x *+ b = (if b then x else 0).
Lemma mul0rn n : 0 *+ n = 0 :> V.
Lemma mulrnDl n : {morph (fun x ⇒ x *+ n) : x y / x + y}.
Lemma mulrnDr x m n : x *+ (m + n) = x *+ m + x *+ n.
Lemma mulrnA x m n : x *+ (m × n) = x *+ m *+ n.
Lemma mulrnAC x m n : x *+ m *+ n = x *+ n *+ m.
Lemma iter_addr n x y : iter n (+%R x) y = x *+ n + y.
Lemma iter_addr_0 n x : iter n (+%R x) 0 = x *+ n.
Lemma sumrMnl I r P (F : I → V) n :
\sum_(i <- r | P i) F i *+ n = (\sum_(i <- r | P i) F i) *+ n.
Lemma sumrMnr x I r P (F : I → nat) :
\sum_(i <- r | P i) x *+ F i = x *+ (\sum_(i <- r | P i) F i).
Lemma sumr_const (I : finType) (A : pred I) x : \sum_(i in A) x = x *+ #|A|.
Lemma sumr_const_nat m n x : \sum_(n ≤ i < m) x = x *+ (m - n).
Definition addr_closed (S : {pred V}) :=
0 \in S ∧ {in S &, ∀ u v, u + v \in S}.
End NmoduleTheory.
#[short(type="zmodType")]
HB.structure Definition Zmodule := {V of Nmodule_isZmodule V & Nmodule V}.
Module ZmodExports.
Bind Scope ring_scope with Zmodule.sort.
End ZmodExports.
Local Notation "-%R" := (@opp _) : ring_scope.
Local Notation "- x" := (opp x) : ring_scope.
Local Notation "x - y" := (x + - y) : ring_scope.
Local Notation "x *- n" := (- (x *+ n)) : ring_scope.
Section ZmoduleTheory.
Variable V : zmodType.
Implicit Types x y : V.
Lemma addrN : @right_inverse V V V 0 -%R +%R.
Definition subrr := addrN.
Lemma addKr : @left_loop V V -%R +%R.
Lemma addNKr : @rev_left_loop V V -%R +%R.
Lemma addrK : @right_loop V V -%R +%R.
Lemma addrNK : @rev_right_loop V V -%R +%R.
Definition subrK := addrNK.
Lemma subKr x : involutive (fun y ⇒ x - y).
Lemma addrI : @right_injective V V V +%R.
Lemma addIr : @left_injective V V V +%R.
Lemma subrI : right_injective (fun x y ⇒ x - y).
Lemma subIr : left_injective (fun x y ⇒ x - y).
Lemma opprK : @involutive V -%R.
Lemma oppr_inj : @injective V V -%R.
Lemma oppr0 : -0 = 0 :> V.
Lemma oppr_eq0 x : (- x == 0) = (x == 0).
Lemma subr0 x : x - 0 = x.
Lemma sub0r x : 0 - x = - x.
Lemma opprB x y : - (x - y) = y - x.
Lemma opprD : {morph -%R: x y / x + y : V}.
Lemma addrKA z x y : (x + z) - (z + y) = x - y.
Lemma subrKA z x y : (x - z) + (z + y) = x + y.
Lemma addr0_eq x y : x + y = 0 → - x = y.
Lemma subr0_eq x y : x - y = 0 → x = y.
Lemma subr_eq x y z : (x - z == y) = (x == y + z).
Lemma subr_eq0 x y : (x - y == 0) = (x == y).
Lemma addr_eq0 x y : (x + y == 0) = (x == - y).
Lemma eqr_opp x y : (- x == - y) = (x == y).
Lemma eqr_oppLR x y : (- x == y) = (x == - y).
Lemma mulNrn x n : (- x) *+ n = x *- n.
Lemma mulrnBl n : {morph (fun x ⇒ x *+ n) : x y / x - y}.
Lemma mulrnBr x m n : n ≤ m → x *+ (m - n) = x *+ m - x *+ n.
Lemma sumrN I r P (F : I → V) :
(\sum_(i <- r | P i) - F i = - (\sum_(i <- r | P i) F i)).
Lemma sumrB I r (P : pred I) (F1 F2 : I → V) :
\sum_(i <- r | P i) (F1 i - F2 i)
= \sum_(i <- r | P i) F1 i - \sum_(i <- r | P i) F2 i.
Lemma telescope_sumr n m (f : nat → V) : n ≤ m →
\sum_(n ≤ k < m) (f k.+1 - f k) = f m - f n.
Lemma telescope_sumr_eq n m (f u : nat → V) : n ≤ m →
(∀ k, (n ≤ k < m)%N → u k = f k.+1 - f k) →
\sum_(n ≤ k < m) u k = f m - f n.
Section ClosedPredicates.
Variable S : {pred V}.
Definition oppr_closed := {in S, ∀ u, - u \in S}.
Definition subr_2closed := {in S &, ∀ u v, u - v \in S}.
Definition zmod_closed := 0 \in S ∧ subr_2closed.
Lemma zmod_closedN : zmod_closed → oppr_closed.
Lemma zmod_closedD : zmod_closed → addr_closed S.
End ClosedPredicates.
End ZmoduleTheory.
Arguments addrI {V} y [x1 x2].
Arguments addIr {V} x [x1 x2].
Arguments opprK {V}.
Arguments oppr_inj {V} [x1 x2].
Arguments telescope_sumr_eq {V n m} f u.
#[short(type="semiRingType")]
HB.structure Definition SemiRing := { R of Nmodule_isSemiRing R & Nmodule R }.
Module SemiRingExports.
Bind Scope ring_scope with SemiRing.sort.
End SemiRingExports.
Definition exp R x n := iterop n (@mul R) x (@one R).
Arguments exp : simpl never.
Definition comm R x y := @mul R x y = mul y x.
Definition lreg R x := injective (@mul R x).
Definition rreg R x := injective ((@mul R)^~ x).
Local Notation "1" := (@one _) : ring_scope.
Local Notation "n %:R" := (1 *+ n) : ring_scope.
Local Notation "*%R" := (@mul _) : function_scope.
Local Notation "x * y" := (mul x y) : ring_scope.
Local Notation "x ^+ n" := (exp x n) : ring_scope.
Local Notation "\prod_ ( i <- r | P ) F" := (\big[*%R/1]_(i <- r | P) F).
Local Notation "\prod_ ( i | P ) F" := (\big[*%R/1]_(i | P) F).
Local Notation "\prod_ ( i 'in' A ) F" := (\big[*%R/1]_(i in A) F).
Local Notation "\prod_ ( m <= i < n ) F" := (\big[*%R/1%R]_(m ≤ i < n) F%R).
The ``field'' characteristic; the definition, and many of the theorems,
has to apply to rings as well; indeed, we need the Frobenius automorphism
results for a non commutative ring in the proof of Gorenstein 2.6.3.
Definition char (R : semiRingType) : nat_pred :=
[pred p | prime p & p%:R == 0 :> R].
Local Notation has_char0 L := (char L =i pred0).
[pred p | prime p & p%:R == 0 :> R].
Local Notation has_char0 L := (char L =i pred0).
Converse ring tag.
Definition converse R : Type := R.
Local Notation "R ^c" := (converse R) (at level 2, format "R ^c") : type_scope.
Section SemiRingTheory.
Variable R : semiRingType.
Implicit Types x y : R.
Lemma oner_eq0 : (1 == 0 :> R) = false.
#[export]
HB.instance Definition _ := Monoid.isLaw.Build R 1 *%R mulrA mul1r mulr1.
#[export]
HB.instance Definition _ := Monoid.isMulLaw.Build R 0 *%R mul0r mulr0.
#[export]
HB.instance Definition _ := Monoid.isAddLaw.Build R *%R +%R mulrDl mulrDr.
Lemma mulr_suml I r P (F : I → R) x :
(\sum_(i <- r | P i) F i) × x = \sum_(i <- r | P i) F i × x.
Lemma mulr_sumr I r P (F : I → R) x :
x × (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) x × F i.
Lemma mulrnAl x y n : (x *+ n) × y = (x × y) *+ n.
Lemma mulrnAr x y n : x × (y *+ n) = (x × y) *+ n.
Lemma mulr_natl x n : n%:R × x = x *+ n.
Lemma mulr_natr x n : x × n%:R = x *+ n.
Lemma natrD m n : (m + n)%:R = m%:R + n%:R :> R.
Lemma natr1 n : n%:R + 1 = n.+1%:R :> R.
Lemma nat1r n : 1 + n%:R = n.+1%:R :> R.
Definition natr_sum := big_morph (natmul 1) natrD (mulr0n 1).
Lemma natrM m n : (m × n)%:R = m%:R × n%:R :> R.
Lemma expr0 x : x ^+ 0 = 1.
Lemma expr1 x : x ^+ 1 = x.
Lemma expr2 x : x ^+ 2 = x × x.
Lemma exprS x n : x ^+ n.+1 = x × x ^+ n.
Lemma expr0n n : 0 ^+ n = (n == 0%N)%:R :> R.
Lemma expr1n n : 1 ^+ n = 1 :> R.
Lemma exprD x m n : x ^+ (m + n) = x ^+ m × x ^+ n.
Lemma exprSr x n : x ^+ n.+1 = x ^+ n × x.
Lemma expr_sum x (I : Type) (s : seq I) (P : pred I) F :
x ^+ (\sum_(i <- s | P i) F i) = \prod_(i <- s | P i) x ^+ F i :> R.
Lemma commr_sym x y : comm x y → comm y x.
Lemma commr_refl x : comm x x.
Lemma commr0 x : comm x 0.
Lemma commr1 x : comm x 1.
Lemma commrD x y z : comm x y → comm x z → comm x (y + z).
Lemma commr_sum (I : Type) (s : seq I) (P : pred I) (F : I → R) x :
(∀ i, P i → comm x (F i)) → comm x (\sum_(i <- s | P i) F i).
Lemma commrMn x y n : comm x y → comm x (y *+ n).
Lemma commrM x y z : comm x y → comm x z → comm x (y × z).
Lemma commr_prod (I : Type) (s : seq I) (P : pred I) (F : I → R) x :
(∀ i, P i → comm x (F i)) → comm x (\prod_(i <- s | P i) F i).
Lemma commr_nat x n : comm x n%:R.
Lemma commrX x y n : comm x y → comm x (y ^+ n).
Lemma exprMn_comm x y n : comm x y → (x × y) ^+ n = x ^+ n × y ^+ n.
Lemma exprMn_n x m n : (x *+ m) ^+ n = x ^+ n *+ (m ^ n) :> R.
Lemma exprM x m n : x ^+ (m × n) = x ^+ m ^+ n.
Lemma exprAC x m n : (x ^+ m) ^+ n = (x ^+ n) ^+ m.
Lemma expr_mod n x i : x ^+ n = 1 → x ^+ (i %% n) = x ^+ i.
Lemma expr_dvd n x i : x ^+ n = 1 → n %| i → x ^+ i = 1.
Lemma natrX n k : (n ^ k)%:R = n%:R ^+ k :> R.
Lemma lastr_eq0 (s : seq R) x : x != 0 → (last x s == 0) = (last 1 s == 0).
Lemma mulrI_eq0 x y : lreg x → (x × y == 0) = (y == 0).
Lemma lreg_neq0 x : lreg x → x != 0.
Lemma lreg1 : lreg (1 : R).
Lemma lregM x y : lreg x → lreg y → lreg (x × y).
Lemma lregMl (a b: R) : lreg (a × b) → lreg b.
Lemma rregMr (a b: R) : rreg (a × b) → rreg a.
Lemma lregX x n : lreg x → lreg (x ^+ n).
Lemma iter_mulr n x y : iter n ( *%R x) y = x ^+ n × y.
Lemma iter_mulr_1 n x : iter n ( *%R x) 1 = x ^+ n.
Lemma prodr_const (I : finType) (A : pred I) x : \prod_(i in A) x = x ^+ #|A|.
Lemma prodr_const_nat n m x : \prod_(n ≤ i < m) x = x ^+ (m - n).
Lemma prodrXr x I r P (F : I → nat) :
\prod_(i <- r | P i) x ^+ F i = x ^+ (\sum_(i <- r | P i) F i).
Lemma prodrM_comm {I : eqType} r (P : pred I) (F G : I → R) :
(∀ i j, P i → P j → comm (F i) (G j)) →
\prod_(i <- r | P i) (F i × G i) =
\prod_(i <- r | P i) F i × \prod_(i <- r | P i) G i.
Lemma prodrMl_comm {I : finType} (A : pred I) (x : R) F :
(∀ i, A i → comm x (F i)) →
\prod_(i in A) (x × F i) = x ^+ #|A| × \prod_(i in A) F i.
Lemma prodrMr_comm {I : finType} (A : pred I) (x : R) F :
(∀ i, A i → comm x (F i)) →
\prod_(i in A) (F i × x) = \prod_(i in A) F i × x ^+ #|A|.
Lemma prodrMn (I : Type) (s : seq I) (P : pred I) (F : I → R) (g : I → nat) :
\prod_(i <- s | P i) (F i *+ g i) =
\prod_(i <- s | P i) (F i) *+ \prod_(i <- s | P i) g i.
Lemma prodrMn_const n (I : finType) (A : pred I) (F : I → R) :
\prod_(i in A) (F i *+ n) = \prod_(i in A) F i *+ n ^ #|A|.
Lemma natr_prod I r P (F : I → nat) :
(\prod_(i <- r | P i) F i)%:R = \prod_(i <- r | P i) (F i)%:R :> R.
Lemma exprDn_comm x y n (cxy : comm x y) :
(x + y) ^+ n = \sum_(i < n.+1) (x ^+ (n - i) × y ^+ i) *+ 'C(n, i).
Lemma exprD1n x n : (x + 1) ^+ n = \sum_(i < n.+1) x ^+ i *+ 'C(n, i).
Lemma sqrrD1 x : (x + 1) ^+ 2 = x ^+ 2 + x *+ 2 + 1.
Definition Frobenius_aut p of p \in char R := fun x ⇒ x ^+ p.
Section FrobeniusAutomorphism.
Variable p : nat.
Hypothesis charFp : p \in char R.
Lemma charf0 : p%:R = 0 :> R.
Lemma charf_prime : prime p.
Hint Resolve charf_prime : core.
Lemma mulrn_char x : x *+ p = 0.
Lemma natr_mod_char n : (n %% p)%:R = n%:R :> R.
Lemma dvdn_charf n : (p %| n)%N = (n%:R == 0 :> R).
Lemma charf_eq : char R =i (p : nat_pred).
Lemma bin_lt_charf_0 k : 0 < k < p → 'C(p, k)%:R = 0 :> R.
Local Notation "x ^f" := (Frobenius_aut charFp x).
Lemma Frobenius_autE x : x^f = x ^+ p.
Local Notation fE := Frobenius_autE.
Lemma Frobenius_aut0 : 0^f = 0.
Lemma Frobenius_aut1 : 1^f = 1.
Lemma Frobenius_autD_comm x y (cxy : comm x y) : (x + y)^f = x^f + y^f.
Lemma Frobenius_autMn x n : (x *+ n)^f = x^f *+ n.
Lemma Frobenius_aut_nat n : (n%:R)^f = n%:R.
Lemma Frobenius_autM_comm x y : comm x y → (x × y)^f = x^f × y^f.
Lemma Frobenius_autX x n : (x ^+ n)^f = x^f ^+ n.
End FrobeniusAutomorphism.
Section Char2.
Hypothesis charR2 : 2 \in char R.
Lemma addrr_char2 x : x + x = 0.
End Char2.
Section ClosedPredicates.
Variable S : {pred R}.
Definition mulr_2closed := {in S &, ∀ u v, u × v \in S}.
Definition mulr_closed := 1 \in S ∧ mulr_2closed.
Definition semiring_closed := addr_closed S ∧ mulr_closed.
Lemma semiring_closedD : semiring_closed → addr_closed S.
Lemma semiring_closedM : semiring_closed → mulr_closed.
End ClosedPredicates.
End SemiRingTheory.
#[short(type="ringType")]
HB.structure Definition Ring := { R of SemiRing R & Zmodule R }.
Local Notation "1" := one.
Local Notation "x * y" := (mul x y).
Lemma mul0r : @left_zero R R 0 mul.
Lemma mulr0 : @right_zero R R 0 mul.
Module RingExports.
Bind Scope ring_scope with Ring.sort.
End RingExports.
Notation sign R b := (exp (- @one R) (nat_of_bool b)) (only parsing).
Local Notation "- 1" := (- (1)) : ring_scope.
Section RingTheory.
Variable R : ringType.
Implicit Types x y : R.
Lemma mulrN x y : x × (- y) = - (x × y).
Lemma mulNr x y : (- x) × y = - (x × y).
Lemma mulrNN x y : (- x) × (- y) = x × y.
Lemma mulN1r x : -1 × x = - x.
Lemma mulrN1 x : x × -1 = - x.
Lemma mulrBl x y z : (y - z) × x = y × x - z × x.
Lemma mulrBr x y z : x × (y - z) = x × y - x × z.
Lemma natrB m n : n ≤ m → (m - n)%:R = m%:R - n%:R :> R.
Lemma commrN x y : comm x y → comm x (- y).
Lemma commrN1 x : comm x (-1).
Lemma commrB x y z : comm x y → comm x z → comm x (y - z).
Lemma commr_sign x n : comm x ((-1) ^+ n).
Lemma signr_odd n : (-1) ^+ (odd n) = (-1) ^+ n :> R.
Lemma signr_eq0 n : ((-1) ^+ n == 0 :> R) = false.
Lemma mulr_sign (b : bool) x : (-1) ^+ b × x = (if b then - x else x).
Lemma signr_addb b1 b2 : (-1) ^+ (b1 (+) b2) = (-1) ^+ b1 × (-1) ^+ b2 :> R.
Lemma signrE (b : bool) : (-1) ^+ b = 1 - b.*2%:R :> R.
Lemma signrN b : (-1) ^+ (~~ b) = - (-1) ^+ b :> R.
Lemma mulr_signM (b1 b2 : bool) x1 x2 :
((-1) ^+ b1 × x1) × ((-1) ^+ b2 × x2) = (-1) ^+ (b1 (+) b2) × (x1 × x2).
Lemma exprNn x n : (- x) ^+ n = (-1) ^+ n × x ^+ n :> R.
Lemma sqrrN x : (- x) ^+ 2 = x ^+ 2.
Lemma sqrr_sign n : ((-1) ^+ n) ^+ 2 = 1 :> R.
Lemma signrMK n : @involutive R ( *%R ((-1) ^+ n)).
Lemma mulrI0_lreg x : (∀ y, x × y = 0 → y = 0) → lreg x.
Lemma lregN x : lreg x → lreg (- x).
Lemma lreg_sign n : lreg ((-1) ^+ n : R).
Lemma prodrN (I : finType) (A : pred I) (F : I → R) :
\prod_(i in A) - F i = (- 1) ^+ #|A| × \prod_(i in A) F i.
Lemma exprBn_comm x y n (cxy : comm x y) :
(x - y) ^+ n =
\sum_(i < n.+1) ((-1) ^+ i × x ^+ (n - i) × y ^+ i) *+ 'C(n, i).
Lemma subrXX_comm x y n (cxy : comm x y) :
x ^+ n - y ^+ n = (x - y) × (\sum_(i < n) x ^+ (n.-1 - i) × y ^+ i).
Lemma subrX1 x n : x ^+ n - 1 = (x - 1) × (\sum_(i < n) x ^+ i).
Lemma sqrrB1 x : (x - 1) ^+ 2 = x ^+ 2 - x *+ 2 + 1.
Lemma subr_sqr_1 x : x ^+ 2 - 1 = (x - 1) × (x + 1).
Section FrobeniusAutomorphism.
Variable p : nat.
Hypothesis charFp : p \in char R.
Hint Resolve charf_prime : core.
Local Notation "x ^f" := (Frobenius_aut charFp x).
Lemma Frobenius_autN x : (- x)^f = - x^f.
Lemma Frobenius_autB_comm x y : comm x y → (x - y)^f = x^f - y^f.
End FrobeniusAutomorphism.
Lemma exprNn_char x n : (char R).-nat n → (- x) ^+ n = - (x ^+ n).
Section Char2.
Hypothesis charR2 : 2 \in char R.
Lemma oppr_char2 x : - x = x.
Lemma subr_char2 x y : x - y = x + y.
Lemma addrK_char2 x : involutive (+%R^~ x).
Lemma addKr_char2 x : involutive (+%R x).
End Char2.
Section ClosedPredicates.
Variable S : {pred R}.
Definition smulr_closed := -1 \in S ∧ mulr_2closed S.
Definition subring_closed := [/\ 1 \in S, subr_2closed S & mulr_2closed S].
Lemma smulr_closedM : smulr_closed → mulr_closed S.
Lemma smulr_closedN : smulr_closed → oppr_closed S.
Lemma subring_closedB : subring_closed → zmod_closed S.
Lemma subring_closedM : subring_closed → smulr_closed.
Lemma subring_closed_semi : subring_closed → semiring_closed S.
End ClosedPredicates.
End RingTheory.
Section ConverseRing.
#[export]
HB.instance Definition _ (T : eqType) := Equality.on T^c.
#[export]
HB.instance Definition _ (T : choiceType) := Choice.on T^c.
#[export]
HB.instance Definition _ (U : nmodType) := Nmodule.on U^c.
#[export]
HB.instance Definition _ (U : zmodType) := Zmodule.on U^c.
#[export]
HB.instance Definition _ (R : semiRingType) :=
let mul' (x y : R) := y × x in
let mulrA' x y z := esym (mulrA z y x) in
let mulrDl' x y z := mulrDr z x y in
let mulrDr' x y z := mulrDl y z x in
Nmodule_isSemiRing.Build R^c
mulrA' mulr1 mul1r mulrDl' mulrDr' mulr0 mul0r oner_neq0.
#[export]
HB.instance Definition _ (R : ringType) := SemiRing.on R^c.
End ConverseRing.
Lemma rev_prodr (R : semiRingType)
(I : Type) (r : seq I) (P : pred I) (E : I → R) :
\prod_(i <- r | P i) (E i : R^c) = \prod_(i <- rev r | P i) E i.
Section SemiRightRegular.
Variable R : semiRingType.
Implicit Types x y : R.
Lemma mulIr_eq0 x y : rreg x → (y × x == 0) = (y == 0).
Lemma rreg_neq0 x : rreg x → x != 0.
Lemma rreg1 : rreg (1 : R).
Lemma rregM x y : rreg x → rreg y → rreg (x × y).
Lemma revrX x n : (x : R^c) ^+ n = (x : R) ^+ n.
Lemma rregX x n : rreg x → rreg (x ^+ n).
End SemiRightRegular.
Section RightRegular.
Variable R : ringType.
Implicit Types x y : R.
Lemma mulIr0_rreg x : (∀ y, y × x = 0 → y = 0) → rreg x.
Lemma rregN x : rreg x → rreg (- x).
End RightRegular.
#[short(type="lmodType")]
HB.structure Definition Lmodule (R : ringType) :=
{M of Zmodule M & Zmodule_isLmodule R M}.
Local Notation "R ^c" := (converse R) (at level 2, format "R ^c") : type_scope.
Section SemiRingTheory.
Variable R : semiRingType.
Implicit Types x y : R.
Lemma oner_eq0 : (1 == 0 :> R) = false.
#[export]
HB.instance Definition _ := Monoid.isLaw.Build R 1 *%R mulrA mul1r mulr1.
#[export]
HB.instance Definition _ := Monoid.isMulLaw.Build R 0 *%R mul0r mulr0.
#[export]
HB.instance Definition _ := Monoid.isAddLaw.Build R *%R +%R mulrDl mulrDr.
Lemma mulr_suml I r P (F : I → R) x :
(\sum_(i <- r | P i) F i) × x = \sum_(i <- r | P i) F i × x.
Lemma mulr_sumr I r P (F : I → R) x :
x × (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) x × F i.
Lemma mulrnAl x y n : (x *+ n) × y = (x × y) *+ n.
Lemma mulrnAr x y n : x × (y *+ n) = (x × y) *+ n.
Lemma mulr_natl x n : n%:R × x = x *+ n.
Lemma mulr_natr x n : x × n%:R = x *+ n.
Lemma natrD m n : (m + n)%:R = m%:R + n%:R :> R.
Lemma natr1 n : n%:R + 1 = n.+1%:R :> R.
Lemma nat1r n : 1 + n%:R = n.+1%:R :> R.
Definition natr_sum := big_morph (natmul 1) natrD (mulr0n 1).
Lemma natrM m n : (m × n)%:R = m%:R × n%:R :> R.
Lemma expr0 x : x ^+ 0 = 1.
Lemma expr1 x : x ^+ 1 = x.
Lemma expr2 x : x ^+ 2 = x × x.
Lemma exprS x n : x ^+ n.+1 = x × x ^+ n.
Lemma expr0n n : 0 ^+ n = (n == 0%N)%:R :> R.
Lemma expr1n n : 1 ^+ n = 1 :> R.
Lemma exprD x m n : x ^+ (m + n) = x ^+ m × x ^+ n.
Lemma exprSr x n : x ^+ n.+1 = x ^+ n × x.
Lemma expr_sum x (I : Type) (s : seq I) (P : pred I) F :
x ^+ (\sum_(i <- s | P i) F i) = \prod_(i <- s | P i) x ^+ F i :> R.
Lemma commr_sym x y : comm x y → comm y x.
Lemma commr_refl x : comm x x.
Lemma commr0 x : comm x 0.
Lemma commr1 x : comm x 1.
Lemma commrD x y z : comm x y → comm x z → comm x (y + z).
Lemma commr_sum (I : Type) (s : seq I) (P : pred I) (F : I → R) x :
(∀ i, P i → comm x (F i)) → comm x (\sum_(i <- s | P i) F i).
Lemma commrMn x y n : comm x y → comm x (y *+ n).
Lemma commrM x y z : comm x y → comm x z → comm x (y × z).
Lemma commr_prod (I : Type) (s : seq I) (P : pred I) (F : I → R) x :
(∀ i, P i → comm x (F i)) → comm x (\prod_(i <- s | P i) F i).
Lemma commr_nat x n : comm x n%:R.
Lemma commrX x y n : comm x y → comm x (y ^+ n).
Lemma exprMn_comm x y n : comm x y → (x × y) ^+ n = x ^+ n × y ^+ n.
Lemma exprMn_n x m n : (x *+ m) ^+ n = x ^+ n *+ (m ^ n) :> R.
Lemma exprM x m n : x ^+ (m × n) = x ^+ m ^+ n.
Lemma exprAC x m n : (x ^+ m) ^+ n = (x ^+ n) ^+ m.
Lemma expr_mod n x i : x ^+ n = 1 → x ^+ (i %% n) = x ^+ i.
Lemma expr_dvd n x i : x ^+ n = 1 → n %| i → x ^+ i = 1.
Lemma natrX n k : (n ^ k)%:R = n%:R ^+ k :> R.
Lemma lastr_eq0 (s : seq R) x : x != 0 → (last x s == 0) = (last 1 s == 0).
Lemma mulrI_eq0 x y : lreg x → (x × y == 0) = (y == 0).
Lemma lreg_neq0 x : lreg x → x != 0.
Lemma lreg1 : lreg (1 : R).
Lemma lregM x y : lreg x → lreg y → lreg (x × y).
Lemma lregMl (a b: R) : lreg (a × b) → lreg b.
Lemma rregMr (a b: R) : rreg (a × b) → rreg a.
Lemma lregX x n : lreg x → lreg (x ^+ n).
Lemma iter_mulr n x y : iter n ( *%R x) y = x ^+ n × y.
Lemma iter_mulr_1 n x : iter n ( *%R x) 1 = x ^+ n.
Lemma prodr_const (I : finType) (A : pred I) x : \prod_(i in A) x = x ^+ #|A|.
Lemma prodr_const_nat n m x : \prod_(n ≤ i < m) x = x ^+ (m - n).
Lemma prodrXr x I r P (F : I → nat) :
\prod_(i <- r | P i) x ^+ F i = x ^+ (\sum_(i <- r | P i) F i).
Lemma prodrM_comm {I : eqType} r (P : pred I) (F G : I → R) :
(∀ i j, P i → P j → comm (F i) (G j)) →
\prod_(i <- r | P i) (F i × G i) =
\prod_(i <- r | P i) F i × \prod_(i <- r | P i) G i.
Lemma prodrMl_comm {I : finType} (A : pred I) (x : R) F :
(∀ i, A i → comm x (F i)) →
\prod_(i in A) (x × F i) = x ^+ #|A| × \prod_(i in A) F i.
Lemma prodrMr_comm {I : finType} (A : pred I) (x : R) F :
(∀ i, A i → comm x (F i)) →
\prod_(i in A) (F i × x) = \prod_(i in A) F i × x ^+ #|A|.
Lemma prodrMn (I : Type) (s : seq I) (P : pred I) (F : I → R) (g : I → nat) :
\prod_(i <- s | P i) (F i *+ g i) =
\prod_(i <- s | P i) (F i) *+ \prod_(i <- s | P i) g i.
Lemma prodrMn_const n (I : finType) (A : pred I) (F : I → R) :
\prod_(i in A) (F i *+ n) = \prod_(i in A) F i *+ n ^ #|A|.
Lemma natr_prod I r P (F : I → nat) :
(\prod_(i <- r | P i) F i)%:R = \prod_(i <- r | P i) (F i)%:R :> R.
Lemma exprDn_comm x y n (cxy : comm x y) :
(x + y) ^+ n = \sum_(i < n.+1) (x ^+ (n - i) × y ^+ i) *+ 'C(n, i).
Lemma exprD1n x n : (x + 1) ^+ n = \sum_(i < n.+1) x ^+ i *+ 'C(n, i).
Lemma sqrrD1 x : (x + 1) ^+ 2 = x ^+ 2 + x *+ 2 + 1.
Definition Frobenius_aut p of p \in char R := fun x ⇒ x ^+ p.
Section FrobeniusAutomorphism.
Variable p : nat.
Hypothesis charFp : p \in char R.
Lemma charf0 : p%:R = 0 :> R.
Lemma charf_prime : prime p.
Hint Resolve charf_prime : core.
Lemma mulrn_char x : x *+ p = 0.
Lemma natr_mod_char n : (n %% p)%:R = n%:R :> R.
Lemma dvdn_charf n : (p %| n)%N = (n%:R == 0 :> R).
Lemma charf_eq : char R =i (p : nat_pred).
Lemma bin_lt_charf_0 k : 0 < k < p → 'C(p, k)%:R = 0 :> R.
Local Notation "x ^f" := (Frobenius_aut charFp x).
Lemma Frobenius_autE x : x^f = x ^+ p.
Local Notation fE := Frobenius_autE.
Lemma Frobenius_aut0 : 0^f = 0.
Lemma Frobenius_aut1 : 1^f = 1.
Lemma Frobenius_autD_comm x y (cxy : comm x y) : (x + y)^f = x^f + y^f.
Lemma Frobenius_autMn x n : (x *+ n)^f = x^f *+ n.
Lemma Frobenius_aut_nat n : (n%:R)^f = n%:R.
Lemma Frobenius_autM_comm x y : comm x y → (x × y)^f = x^f × y^f.
Lemma Frobenius_autX x n : (x ^+ n)^f = x^f ^+ n.
End FrobeniusAutomorphism.
Section Char2.
Hypothesis charR2 : 2 \in char R.
Lemma addrr_char2 x : x + x = 0.
End Char2.
Section ClosedPredicates.
Variable S : {pred R}.
Definition mulr_2closed := {in S &, ∀ u v, u × v \in S}.
Definition mulr_closed := 1 \in S ∧ mulr_2closed.
Definition semiring_closed := addr_closed S ∧ mulr_closed.
Lemma semiring_closedD : semiring_closed → addr_closed S.
Lemma semiring_closedM : semiring_closed → mulr_closed.
End ClosedPredicates.
End SemiRingTheory.
#[short(type="ringType")]
HB.structure Definition Ring := { R of SemiRing R & Zmodule R }.
Local Notation "1" := one.
Local Notation "x * y" := (mul x y).
Lemma mul0r : @left_zero R R 0 mul.
Lemma mulr0 : @right_zero R R 0 mul.
Module RingExports.
Bind Scope ring_scope with Ring.sort.
End RingExports.
Notation sign R b := (exp (- @one R) (nat_of_bool b)) (only parsing).
Local Notation "- 1" := (- (1)) : ring_scope.
Section RingTheory.
Variable R : ringType.
Implicit Types x y : R.
Lemma mulrN x y : x × (- y) = - (x × y).
Lemma mulNr x y : (- x) × y = - (x × y).
Lemma mulrNN x y : (- x) × (- y) = x × y.
Lemma mulN1r x : -1 × x = - x.
Lemma mulrN1 x : x × -1 = - x.
Lemma mulrBl x y z : (y - z) × x = y × x - z × x.
Lemma mulrBr x y z : x × (y - z) = x × y - x × z.
Lemma natrB m n : n ≤ m → (m - n)%:R = m%:R - n%:R :> R.
Lemma commrN x y : comm x y → comm x (- y).
Lemma commrN1 x : comm x (-1).
Lemma commrB x y z : comm x y → comm x z → comm x (y - z).
Lemma commr_sign x n : comm x ((-1) ^+ n).
Lemma signr_odd n : (-1) ^+ (odd n) = (-1) ^+ n :> R.
Lemma signr_eq0 n : ((-1) ^+ n == 0 :> R) = false.
Lemma mulr_sign (b : bool) x : (-1) ^+ b × x = (if b then - x else x).
Lemma signr_addb b1 b2 : (-1) ^+ (b1 (+) b2) = (-1) ^+ b1 × (-1) ^+ b2 :> R.
Lemma signrE (b : bool) : (-1) ^+ b = 1 - b.*2%:R :> R.
Lemma signrN b : (-1) ^+ (~~ b) = - (-1) ^+ b :> R.
Lemma mulr_signM (b1 b2 : bool) x1 x2 :
((-1) ^+ b1 × x1) × ((-1) ^+ b2 × x2) = (-1) ^+ (b1 (+) b2) × (x1 × x2).
Lemma exprNn x n : (- x) ^+ n = (-1) ^+ n × x ^+ n :> R.
Lemma sqrrN x : (- x) ^+ 2 = x ^+ 2.
Lemma sqrr_sign n : ((-1) ^+ n) ^+ 2 = 1 :> R.
Lemma signrMK n : @involutive R ( *%R ((-1) ^+ n)).
Lemma mulrI0_lreg x : (∀ y, x × y = 0 → y = 0) → lreg x.
Lemma lregN x : lreg x → lreg (- x).
Lemma lreg_sign n : lreg ((-1) ^+ n : R).
Lemma prodrN (I : finType) (A : pred I) (F : I → R) :
\prod_(i in A) - F i = (- 1) ^+ #|A| × \prod_(i in A) F i.
Lemma exprBn_comm x y n (cxy : comm x y) :
(x - y) ^+ n =
\sum_(i < n.+1) ((-1) ^+ i × x ^+ (n - i) × y ^+ i) *+ 'C(n, i).
Lemma subrXX_comm x y n (cxy : comm x y) :
x ^+ n - y ^+ n = (x - y) × (\sum_(i < n) x ^+ (n.-1 - i) × y ^+ i).
Lemma subrX1 x n : x ^+ n - 1 = (x - 1) × (\sum_(i < n) x ^+ i).
Lemma sqrrB1 x : (x - 1) ^+ 2 = x ^+ 2 - x *+ 2 + 1.
Lemma subr_sqr_1 x : x ^+ 2 - 1 = (x - 1) × (x + 1).
Section FrobeniusAutomorphism.
Variable p : nat.
Hypothesis charFp : p \in char R.
Hint Resolve charf_prime : core.
Local Notation "x ^f" := (Frobenius_aut charFp x).
Lemma Frobenius_autN x : (- x)^f = - x^f.
Lemma Frobenius_autB_comm x y : comm x y → (x - y)^f = x^f - y^f.
End FrobeniusAutomorphism.
Lemma exprNn_char x n : (char R).-nat n → (- x) ^+ n = - (x ^+ n).
Section Char2.
Hypothesis charR2 : 2 \in char R.
Lemma oppr_char2 x : - x = x.
Lemma subr_char2 x y : x - y = x + y.
Lemma addrK_char2 x : involutive (+%R^~ x).
Lemma addKr_char2 x : involutive (+%R x).
End Char2.
Section ClosedPredicates.
Variable S : {pred R}.
Definition smulr_closed := -1 \in S ∧ mulr_2closed S.
Definition subring_closed := [/\ 1 \in S, subr_2closed S & mulr_2closed S].
Lemma smulr_closedM : smulr_closed → mulr_closed S.
Lemma smulr_closedN : smulr_closed → oppr_closed S.
Lemma subring_closedB : subring_closed → zmod_closed S.
Lemma subring_closedM : subring_closed → smulr_closed.
Lemma subring_closed_semi : subring_closed → semiring_closed S.
End ClosedPredicates.
End RingTheory.
Section ConverseRing.
#[export]
HB.instance Definition _ (T : eqType) := Equality.on T^c.
#[export]
HB.instance Definition _ (T : choiceType) := Choice.on T^c.
#[export]
HB.instance Definition _ (U : nmodType) := Nmodule.on U^c.
#[export]
HB.instance Definition _ (U : zmodType) := Zmodule.on U^c.
#[export]
HB.instance Definition _ (R : semiRingType) :=
let mul' (x y : R) := y × x in
let mulrA' x y z := esym (mulrA z y x) in
let mulrDl' x y z := mulrDr z x y in
let mulrDr' x y z := mulrDl y z x in
Nmodule_isSemiRing.Build R^c
mulrA' mulr1 mul1r mulrDl' mulrDr' mulr0 mul0r oner_neq0.
#[export]
HB.instance Definition _ (R : ringType) := SemiRing.on R^c.
End ConverseRing.
Lemma rev_prodr (R : semiRingType)
(I : Type) (r : seq I) (P : pred I) (E : I → R) :
\prod_(i <- r | P i) (E i : R^c) = \prod_(i <- rev r | P i) E i.
Section SemiRightRegular.
Variable R : semiRingType.
Implicit Types x y : R.
Lemma mulIr_eq0 x y : rreg x → (y × x == 0) = (y == 0).
Lemma rreg_neq0 x : rreg x → x != 0.
Lemma rreg1 : rreg (1 : R).
Lemma rregM x y : rreg x → rreg y → rreg (x × y).
Lemma revrX x n : (x : R^c) ^+ n = (x : R) ^+ n.
Lemma rregX x n : rreg x → rreg (x ^+ n).
End SemiRightRegular.
Section RightRegular.
Variable R : ringType.
Implicit Types x y : R.
Lemma mulIr0_rreg x : (∀ y, y × x = 0 → y = 0) → rreg x.
Lemma rregN x : rreg x → rreg (- x).
End RightRegular.
#[short(type="lmodType")]
HB.structure Definition Lmodule (R : ringType) :=
{M of Zmodule M & Zmodule_isLmodule R M}.
FIXME: see #1126 and #1127
Arguments scalerA [R s] (a b)%ring_scope v.
Module LmodExports.
Bind Scope ring_scope with Lmodule.sort.
End LmodExports.
Local Notation "*:%R" := (@scale _ _) : function_scope.
Local Notation "a *: v" := (scale a v) : ring_scope.
Section LmoduleTheory.
Variables (R : ringType) (V : lmodType R).
Implicit Types (a b c : R) (u v : V).
Lemma scale0r v : 0 *: v = 0.
Lemma scaler0 a : a *: 0 = 0 :> V.
Lemma scaleNr a v : - a *: v = - (a *: v).
Lemma scaleN1r v : (- 1) *: v = - v.
Lemma scalerN a v : a *: (- v) = - (a *: v).
Lemma scalerBl a b v : (a - b) *: v = a *: v - b *: v.
Lemma scalerBr a u v : a *: (u - v) = a *: u - a *: v.
Lemma scaler_nat n v : n%:R *: v = v *+ n.
Lemma scaler_sign (b : bool) v: (-1) ^+ b *: v = (if b then - v else v).
Lemma signrZK n : @involutive V ( *:%R ((-1) ^+ n)).
Lemma scalerMnl a v n : a *: v *+ n = (a *+ n) *: v.
Lemma scalerMnr a v n : a *: v *+ n = a *: (v *+ n).
Lemma scaler_suml v I r (P : pred I) F :
(\sum_(i <- r | P i) F i) *: v = \sum_(i <- r | P i) F i *: v.
Lemma scaler_sumr a I r (P : pred I) (F : I → V) :
a *: (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) a *: F i.
Section ClosedPredicates.
Variable S : {pred V}.
Definition scaler_closed := ∀ a, {in S, ∀ v, a *: v \in S}.
Definition linear_closed := ∀ a, {in S &, ∀ u v, a *: u + v \in S}.
Definition submod_closed := 0 \in S ∧ linear_closed.
Lemma linear_closedB : linear_closed → subr_2closed S.
Lemma submod_closedB : submod_closed → zmod_closed S.
Lemma submod_closedZ : submod_closed → scaler_closed.
End ClosedPredicates.
End LmoduleTheory.
#[short(type="lalgType")]
HB.structure Definition Lalgebra R :=
{A of Lmodule_isLalgebra R A & Ring A & Lmodule R A}.
Module LalgExports.
Bind Scope ring_scope with Lalgebra.sort.
End LalgExports.
Module LmodExports.
Bind Scope ring_scope with Lmodule.sort.
End LmodExports.
Local Notation "*:%R" := (@scale _ _) : function_scope.
Local Notation "a *: v" := (scale a v) : ring_scope.
Section LmoduleTheory.
Variables (R : ringType) (V : lmodType R).
Implicit Types (a b c : R) (u v : V).
Lemma scale0r v : 0 *: v = 0.
Lemma scaler0 a : a *: 0 = 0 :> V.
Lemma scaleNr a v : - a *: v = - (a *: v).
Lemma scaleN1r v : (- 1) *: v = - v.
Lemma scalerN a v : a *: (- v) = - (a *: v).
Lemma scalerBl a b v : (a - b) *: v = a *: v - b *: v.
Lemma scalerBr a u v : a *: (u - v) = a *: u - a *: v.
Lemma scaler_nat n v : n%:R *: v = v *+ n.
Lemma scaler_sign (b : bool) v: (-1) ^+ b *: v = (if b then - v else v).
Lemma signrZK n : @involutive V ( *:%R ((-1) ^+ n)).
Lemma scalerMnl a v n : a *: v *+ n = (a *+ n) *: v.
Lemma scalerMnr a v n : a *: v *+ n = a *: (v *+ n).
Lemma scaler_suml v I r (P : pred I) F :
(\sum_(i <- r | P i) F i) *: v = \sum_(i <- r | P i) F i *: v.
Lemma scaler_sumr a I r (P : pred I) (F : I → V) :
a *: (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) a *: F i.
Section ClosedPredicates.
Variable S : {pred V}.
Definition scaler_closed := ∀ a, {in S, ∀ v, a *: v \in S}.
Definition linear_closed := ∀ a, {in S &, ∀ u v, a *: u + v \in S}.
Definition submod_closed := 0 \in S ∧ linear_closed.
Lemma linear_closedB : linear_closed → subr_2closed S.
Lemma submod_closedB : submod_closed → zmod_closed S.
Lemma submod_closedZ : submod_closed → scaler_closed.
End ClosedPredicates.
End LmoduleTheory.
#[short(type="lalgType")]
HB.structure Definition Lalgebra R :=
{A of Lmodule_isLalgebra R A & Ring A & Lmodule R A}.
Module LalgExports.
Bind Scope ring_scope with Lalgebra.sort.
End LalgExports.
Scalar injection (see the definition of in_alg A below).
Regular ring algebra tag.
Definition regular R : Type := R.
Local Notation "R ^o" := (regular R) (at level 2, format "R ^o") : type_scope.
Section RegularAlgebra.
#[export]
HB.instance Definition _ (V : nmodType) := Nmodule.on V^o.
#[export]
HB.instance Definition _ (V : zmodType) := Zmodule.on V^o.
#[export]
HB.instance Definition _ (R : semiRingType) := SemiRing.on R^o.
#[export]
HB.instance Definition _ (R : ringType) := Ring.on R^o.
#[export]
HB.instance Definition _ (R : ringType) :=
@Zmodule_isLmodule.Build R R^o
(@mul R) (@mulrA R) (@mul1r R) (@mulrDr R) (fun v a b ⇒ mulrDl a b v).
#[export]
HB.instance Definition _ (R : ringType) :=
Lmodule_isLalgebra.Build R R^o mulrA.
End RegularAlgebra.
Section LalgebraTheory.
Variables (R : ringType) (A : lalgType R).
Implicit Types x y : A.
Lemma mulr_algl a x : (a *: 1) × x = a *: x.
Section ClosedPredicates.
Variable S : {pred A}.
Definition subalg_closed := [/\ 1 \in S, linear_closed S & mulr_2closed S].
Lemma subalg_closedZ : subalg_closed → submod_closed S.
Lemma subalg_closedBM : subalg_closed → subring_closed S.
End ClosedPredicates.
End LalgebraTheory.
Local Notation "R ^o" := (regular R) (at level 2, format "R ^o") : type_scope.
Section RegularAlgebra.
#[export]
HB.instance Definition _ (V : nmodType) := Nmodule.on V^o.
#[export]
HB.instance Definition _ (V : zmodType) := Zmodule.on V^o.
#[export]
HB.instance Definition _ (R : semiRingType) := SemiRing.on R^o.
#[export]
HB.instance Definition _ (R : ringType) := Ring.on R^o.
#[export]
HB.instance Definition _ (R : ringType) :=
@Zmodule_isLmodule.Build R R^o
(@mul R) (@mulrA R) (@mul1r R) (@mulrDr R) (fun v a b ⇒ mulrDl a b v).
#[export]
HB.instance Definition _ (R : ringType) :=
Lmodule_isLalgebra.Build R R^o mulrA.
End RegularAlgebra.
Section LalgebraTheory.
Variables (R : ringType) (A : lalgType R).
Implicit Types x y : A.
Lemma mulr_algl a x : (a *: 1) × x = a *: x.
Section ClosedPredicates.
Variable S : {pred A}.
Definition subalg_closed := [/\ 1 \in S, linear_closed S & mulr_2closed S].
Lemma subalg_closedZ : subalg_closed → submod_closed S.
Lemma subalg_closedBM : subalg_closed → subring_closed S.
End ClosedPredicates.
End LalgebraTheory.
Morphism hierarchy.
Definition semi_additive (U V : nmodType) (f : U → V) : Prop :=
(f 0 = 0) × {morph f : x y / x + y}.
#[mathcomp(axiom="semi_additive")]
HB.structure Definition Additive (U V : nmodType) :=
{f of isSemiAdditive U V f}.
Definition additive (U V : zmodType) (f : U → V) := {morph f : x y / x - y}.
Local Lemma raddf0 : apply 0 = 0.
Local Lemma raddfD : {morph apply : x y / x + y}.
Module AdditiveExports.
Notation "{ 'additive' U -> V }" := (Additive.type U%type V%type) : type_scope.
End AdditiveExports.
Lifted additive operations.
Section LiftedNmod.
Variables (U : Type) (V : nmodType).
Definition null_fun of U : V := 0.
Definition add_fun (f g : U → V) x := f x + g x.
End LiftedNmod.
Section LiftedZmod.
Variables (U : Type) (V : zmodType).
Definition sub_fun (f g : U → V) x := f x - g x.
Definition opp_fun (f : U → V) x := - f x.
End LiftedZmod.
Variables (U : Type) (V : nmodType).
Definition null_fun of U : V := 0.
Definition add_fun (f g : U → V) x := f x + g x.
End LiftedNmod.
Section LiftedZmod.
Variables (U : Type) (V : zmodType).
Definition sub_fun (f g : U → V) x := f x - g x.
Definition opp_fun (f : U → V) x := - f x.
End LiftedZmod.
Lifted multiplication.
Section LiftedSemiRing.
Variables (R : semiRingType) (T : Type).
Implicit Type f : T → R.
Definition mull_fun a f x := a × f x.
Definition mulr_fun a f x := f x × a.
Definition mul_fun f g x := f x × g x.
End LiftedSemiRing.
Variables (R : semiRingType) (T : Type).
Implicit Type f : T → R.
Definition mull_fun a f x := a × f x.
Definition mulr_fun a f x := f x × a.
Definition mul_fun f g x := f x × g x.
End LiftedSemiRing.
Lifted linear operations.
Section LiftedScale.
Variables (R : ringType) (U : Type) (V : lmodType R) (A : lalgType R).
Definition scale_fun a (f : U → V) x := a *: f x.
Definition in_alg k : A := k%:A.
End LiftedScale.
Local Notation "\0" := (null_fun _) : function_scope.
Local Notation "f \+ g" := (add_fun f g) : function_scope.
Local Notation "f \- g" := (sub_fun f g) : function_scope.
Local Notation "\- f" := (opp_fun f) : function_scope.
Local Notation "a \*: f" := (scale_fun a f) : function_scope.
Local Notation "x \*o f" := (mull_fun x f) : function_scope.
Local Notation "x \o* f" := (mulr_fun x f) : function_scope.
Local Notation "f \* g" := (mul_fun f g) : function_scope.
Arguments null_fun {_} V _ /.
Arguments in_alg {_} A _ /.
Arguments add_fun {_ _} f g _ /.
Arguments sub_fun {_ _} f g _ /.
Arguments opp_fun {_ _} f _ /.
Arguments mull_fun {_ _} a f _ /.
Arguments mulr_fun {_ _} a f _ /.
Arguments scale_fun {_ _ _} a f _ /.
Arguments mul_fun {_ _} f g _ /.
Section AdditiveTheory.
Section Properties.
Variables (U V : nmodType) (f : {additive U → V}).
Lemma raddf0 : f 0 = 0.
Lemma raddf_eq0 x : injective f → (f x == 0) = (x == 0).
Lemma raddfD : {morph f : x y / x + y}.
Lemma raddfMn n : {morph f
Variables (R : ringType) (U : Type) (V : lmodType R) (A : lalgType R).
Definition scale_fun a (f : U → V) x := a *: f x.
Definition in_alg k : A := k%:A.
End LiftedScale.
Local Notation "\0" := (null_fun _) : function_scope.
Local Notation "f \+ g" := (add_fun f g) : function_scope.
Local Notation "f \- g" := (sub_fun f g) : function_scope.
Local Notation "\- f" := (opp_fun f) : function_scope.
Local Notation "a \*: f" := (scale_fun a f) : function_scope.
Local Notation "x \*o f" := (mull_fun x f) : function_scope.
Local Notation "x \o* f" := (mulr_fun x f) : function_scope.
Local Notation "f \* g" := (mul_fun f g) : function_scope.
Arguments null_fun {_} V _ /.
Arguments in_alg {_} A _ /.
Arguments add_fun {_ _} f g _ /.
Arguments sub_fun {_ _} f g _ /.
Arguments opp_fun {_ _} f _ /.
Arguments mull_fun {_ _} a f _ /.
Arguments mulr_fun {_ _} a f _ /.
Arguments scale_fun {_ _ _} a f _ /.
Arguments mul_fun {_ _} f g _ /.
Section AdditiveTheory.
Section Properties.
Variables (U V : nmodType) (f : {additive U → V}).
Lemma raddf0 : f 0 = 0.
Lemma raddf_eq0 x : injective f → (f x == 0) = (x == 0).
Lemma raddfD : {morph f : x y / x + y}.
Lemma raddfMn n : {morph f