Skip to content

Commit 9cbbffd

Browse files
committed
finite sums of ring elements
Signed-off-by: Ali Caglayan <[email protected]>
1 parent 5a54a54 commit 9cbbffd

File tree

2 files changed

+242
-3
lines changed

2 files changed

+242
-3
lines changed

theories/Algebra/Groups/Group.v

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -486,6 +486,16 @@ Proof.
486486
exact (ap (fun m => f g + m) IHn).
487487
Defined.
488488

489+
(** All powers of the unit are the unit. *)
490+
Definition grp_pow_unit {G : Group} (n : nat)
491+
: grp_pow (G:=G) mon_unit n = mon_unit.
492+
Proof.
493+
induction n.
494+
1: reflexivity.
495+
lhs rapply left_identity.
496+
exact IHn.
497+
Defined.
498+
489499
(** ** The category of Groups *)
490500

491501
(** ** Groups together with homomorphisms form a 1-category whose equivalences are the group isomorphisms. *)

theories/Algebra/Rings/Ring.v

Lines changed: 232 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,8 @@
11
Require Import WildCat.
2-
Require Import Spaces.Nat.Core.
2+
Require Import Spaces.Nat.Core Spaces.Nat.Arithmetic.
33
(* Some of the material in abstract_algebra and canonical names could be selectively exported to the user, as is done in Groups/Group.v. *)
44
Require Import Classes.interfaces.abstract_algebra.
5-
Require Import Algebra.AbGroups.
5+
Require Export Algebra.AbGroups.
66
Require Export Classes.theory.rings.
77
Require Import Modalities.ReflectiveSubuniverse.
88

@@ -454,7 +454,7 @@ Proof.
454454
- intros R S T f g; cbn; reflexivity.
455455
Defined.
456456

457-
(** ** More Ring laws *)
457+
(** ** Powers *)
458458

459459
(** Powers of ring elements *)
460460
Definition rng_power {R : Ring} (x : R) (n : nat) : R := nat_iter n (x *.) ring_one.
@@ -468,3 +468,232 @@ Proof.
468468
refine ((rng_mult_assoc _ _ _)^ @ _).
469469
exact (ap (x *.) IHn).
470470
Defined.
471+
472+
(** ** Finite Sums *)
473+
474+
(** Indexed finite sum of ring elements. *)
475+
Definition rng_sum {R : Ring} (n : nat) (f : forall k, (k < n)%nat -> R) : R.
476+
Proof.
477+
induction n as [|n IHn].
478+
- exact ring_zero.
479+
- refine (f n _ + IHn _).
480+
intros k Hk.
481+
refine (f k _).
482+
apply leq_S.
483+
exact Hk.
484+
Defined.
485+
486+
(** If the function is constant in the range of a finite sum then the sum is equal to the constant times [n]. This is a group power in the underlying abelian group. *)
487+
Definition rng_sum_const {R : Ring} (n : nat) (r : R)
488+
(f : forall k, (k < n)%nat -> R) (p : forall k Hk, f k Hk = r)
489+
: rng_sum n f = grp_pow r n.
490+
Proof.
491+
induction n as [|n IHn] in f, p |- *.
492+
1: reflexivity.
493+
simpl; f_ap.
494+
apply IHn.
495+
intros k Hk.
496+
apply p.
497+
Defined.
498+
499+
(** If the function is zero in the range of a finite sum then the sum is zero. *)
500+
Definition rng_sum_zero {R : Ring} (n : nat)
501+
(f : forall k, (k < n)%nat -> R) (p : forall k Hk, f k Hk = 0)
502+
: rng_sum n f = 0.
503+
Proof.
504+
lhs nrapply (rng_sum_const _ 0 f p).
505+
apply grp_pow_unit.
506+
Defined.
507+
508+
(** Finite sums distribute over addition. *)
509+
Definition rng_sum_plus {R : Ring} (n : nat) (f g : forall k, (k < n)%nat -> R)
510+
: rng_sum n (fun k Hk => f k Hk + g k Hk)
511+
= rng_sum n (fun k Hk => f k Hk) + rng_sum n (fun k Hk => g k Hk).
512+
Proof.
513+
induction n as [|n IHn].
514+
1: by rewrite rng_plus_zero_l.
515+
simpl.
516+
rhs_V nrapply rng_plus_assoc.
517+
rhs nrapply (ap (_ +)).
518+
2: lhs nrapply rng_plus_assoc.
519+
2: nrapply (ap (+ _)).
520+
2: apply rng_plus_comm.
521+
lhs_V nrapply rng_plus_assoc; f_ap.
522+
rhs_V nrapply rng_plus_assoc; f_ap.
523+
Defined.
524+
525+
(** Double finite sums commute. *)
526+
Definition rng_sum_sum {R : Ring} (m n : nat)
527+
(f : forall i j, (i < m)%nat -> (j < n)%nat -> R)
528+
: rng_sum m (fun i Hi => rng_sum n (fun j Hj => f i j Hi Hj))
529+
= rng_sum n (fun j Hj => rng_sum m (fun i Hi => f i j Hi Hj)).
530+
Proof.
531+
induction n as [|n IHn] in m, f |- *.
532+
1: by nrapply rng_sum_zero.
533+
lhs nrapply rng_sum_plus; cbn; f_ap.
534+
Defined.
535+
536+
(** Finite sums are equal if the functions are equal in the range. *)
537+
Definition path_rng_sum {R : Ring} {n : nat} {f g : forall k, (k < n)%nat -> R}
538+
(p : forall k Hk, f k Hk = g k Hk)
539+
: rng_sum n f = rng_sum n g.
540+
Proof.
541+
induction n as [|n IHn].
542+
1: reflexivity.
543+
cbn; f_ap.
544+
by apply IHn.
545+
Defined.
546+
547+
(** Ring multiplication distributes over finite sums on the left. *)
548+
Definition rng_sum_dist_l {R : Ring} (n : nat) (f : forall k, (k < n)%nat -> R) (r : R)
549+
: r * rng_sum n f = rng_sum n (fun k Hk => r * f k Hk).
550+
Proof.
551+
induction n as [|n IHn].
552+
1: apply rng_mult_zero_r.
553+
lhs nrapply rng_dist_l; simpl; f_ap.
554+
Defined.
555+
556+
(** Ring multiplication distributes over finite sums on the right. *)
557+
Definition rng_sum_dist_r {R : Ring} (n : nat) (f : forall k, (k < n)%nat -> R) (r : R)
558+
: rng_sum n f * r = rng_sum n (fun k Hk => f k Hk * r).
559+
Proof.
560+
induction n as [|n IHn].
561+
1: apply rng_mult_zero_l.
562+
lhs nrapply rng_dist_r; simpl; f_ap.
563+
Defined.
564+
565+
(** ** Kronecker Delta *)
566+
567+
(** The Kronecker delta function is a function of two natural numbers that is 1 when the two numbers are equal and 0 otherwise. It is useful for working with finite sums of ring elements. *)
568+
Definition kronecker_delta {R : Ring} (i j : nat) : R
569+
:= if dec (i = j) then 1 else 0.
570+
571+
(** Kronecker delta with the same index is 1. This holds definitionally for a given natural number, but not definitionally for an arbitrary one. *)
572+
Definition kronecker_delta_refl {R : Ring} (i : nat)
573+
: kronecker_delta (R:=R) i i = 1.
574+
Proof.
575+
unfold kronecker_delta.
576+
generalize (dec (i = i)).
577+
by rapply decidable_paths_refl.
578+
Defined.
579+
580+
(** Kronecker delta with different indices is 0. *)
581+
Definition kronecker_delta_neq {R : Ring} {i j : nat} (p : i <> j)
582+
: kronecker_delta (R:=R) i j = 0.
583+
Proof.
584+
unfold kronecker_delta.
585+
by decidable_false (dec (i = j)) p.
586+
Defined.
587+
588+
(** Kronecker delta is symmetric in its arguments. *)
589+
Definition kronecker_delta_symm {R : Ring} (i j : nat)
590+
: kronecker_delta (R:=R) i j = kronecker_delta j i.
591+
Proof.
592+
unfold kronecker_delta.
593+
destruct (dec (i = j)) as [p|q].
594+
- by decidable_true (dec (j = i)) p^.
595+
- by decidable_false (dec (j = i)) (symmetric_neq q).
596+
Defined.
597+
598+
(** Kronecker delta where the first index is strictly less than the second is 0. *)
599+
Definition kronecker_delta_lt {R : Ring} {i j : nat} (p : (i < j)%nat)
600+
: kronecker_delta (R:=R) i j = 0.
601+
Proof.
602+
apply kronecker_delta_neq.
603+
intros q; destruct q.
604+
by apply not_lt_n_n in p.
605+
Defined.
606+
607+
(** Kronecker delta where the first index is strictly greater than the second is 0. *)
608+
Definition kronecker_delta_gt {R : Ring} {i j : nat} (p : (j < i)%nat)
609+
: kronecker_delta (R:=R) i j = 0.
610+
Proof.
611+
apply kronecker_delta_neq.
612+
intros q; destruct q.
613+
by apply not_lt_n_n in p.
614+
Defined.
615+
616+
(** An injective endofunction on nat preserves the Kronecker delta. *)
617+
Definition kronecker_delta_map_inj {R : Ring} (i j : nat) (f : nat -> nat)
618+
`{!IsInjective f}
619+
: kronecker_delta (R:=R) (f i) (f j) = kronecker_delta i j.
620+
Proof.
621+
unfold kronecker_delta.
622+
destruct (dec (i = j)) as [p|p].
623+
- by decidable_true (dec (f i = f j)) (ap f p).
624+
- destruct (dec (f i = f j)) as [q|q].
625+
+ apply (injective f) in q.
626+
contradiction.
627+
+ reflexivity.
628+
Defined.
629+
630+
(** Kronecker delta commutes with any ring element. *)
631+
Definition kronecker_delta_comm {R : Ring} (i j : nat) (r : R)
632+
: r * kronecker_delta i j = kronecker_delta i j * r.
633+
Proof.
634+
unfold kronecker_delta.
635+
destruct (dec (i = j)).
636+
- exact (rng_mult_one_r _ @ (rng_mult_one_l _)^).
637+
- exact (rng_mult_zero_r _ @ (rng_mult_zero_l _)^).
638+
Defined.
639+
640+
(** Kronecker delta can be used to extract a single term from a finite sum. *)
641+
Definition rng_sum_kronecker_delta_l {R : Ring} (n i : nat) (Hi : (i < n)%nat)
642+
(f : forall k, (k < n)%nat -> R)
643+
: rng_sum n (fun j Hj => kronecker_delta i j * f j Hj) = f i Hi.
644+
Proof.
645+
induction n as [|n IHn] in i, Hi, f |- *.
646+
1: destruct (not_leq_Sn_0 _ Hi).
647+
destruct (dec (i = n)) as [p|p].
648+
- destruct p; simpl.
649+
rewrite kronecker_delta_refl.
650+
rewrite rng_mult_one_l.
651+
rewrite <- rng_plus_zero_r.
652+
f_ap; [f_ap; rapply path_ishprop|].
653+
nrapply rng_sum_zero.
654+
intros k Hk.
655+
rewrite (kronecker_delta_gt Hk).
656+
apply rng_mult_zero_l.
657+
- simpl; lhs nrapply ap.
658+
+ nrapply IHn.
659+
apply diseq_implies_lt in p.
660+
destruct p; [assumption|].
661+
contradiction (lt_implies_not_geq Hi).
662+
+ rewrite (kronecker_delta_neq p).
663+
rewrite rng_mult_zero_l.
664+
rewrite rng_plus_zero_l.
665+
f_ap; apply path_ishprop.
666+
Defined.
667+
668+
(** Variant of [rng_sum_kronecker_delta_l] where the indexing is swapped. *)
669+
Definition rng_sum_kronecker_delta_l' {R : Ring} (n i : nat) (Hi : (i < n)%nat)
670+
(f : forall k, (k < n)%nat -> R)
671+
: rng_sum n (fun j Hj => kronecker_delta j i * f j Hj) = f i Hi.
672+
Proof.
673+
lhs nrapply path_rng_sum.
674+
2: nrapply rng_sum_kronecker_delta_l.
675+
intros k Hk.
676+
cbn; f_ap; apply kronecker_delta_symm.
677+
Defined.
678+
679+
(** Variant of [rng_sum_kronecker_delta_l] where the Kronecker delta appears on the right. *)
680+
Definition rng_sum_kronecker_delta_r {R : Ring} (n i : nat) (Hi : (i < n)%nat)
681+
(f : forall k, (k < n)%nat -> R)
682+
: rng_sum n (fun j Hj => f j Hj * kronecker_delta i j) = f i Hi.
683+
Proof.
684+
lhs nrapply path_rng_sum.
685+
2: nrapply rng_sum_kronecker_delta_l.
686+
intros k Hk.
687+
apply kronecker_delta_comm.
688+
Defined.
689+
690+
(** Variant of [rng_sum_kronecker_delta_r] where the indexing is swapped. *)
691+
Definition rng_sum_kronecker_delta_r' {R : Ring} (n i : nat) (Hi : (i < n)%nat)
692+
(f : forall k, (k < n)%nat -> R)
693+
: rng_sum n (fun j Hj => f j Hj * kronecker_delta j i) = f i Hi.
694+
Proof.
695+
lhs nrapply path_rng_sum.
696+
2: nrapply rng_sum_kronecker_delta_l'.
697+
intros k Hk.
698+
apply kronecker_delta_comm.
699+
Defined.

0 commit comments

Comments
 (0)