Skip to content

Commit 6bb1c3f

Browse files
authored
Merge pull request #2000 from Alizter/ps/rr/switch_to_int_in_cring_z
switch to Int in cring_Z
2 parents 2763488 + 1e60610 commit 6bb1c3f

File tree

12 files changed

+377
-396
lines changed

12 files changed

+377
-396
lines changed

test/Algebra/Rings/Matrix.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
From HoTT Require Import Basics.
22
From HoTT Require Import Algebra.Rings.Matrix.
33
From HoTT Require Import Spaces.Nat.Core Spaces.List.Core.
4-
From HoTT Require Import Algebra.Rings.Z Spaces.BinInt.Core Algebra.Rings.CRing.
4+
From HoTT Require Import Algebra.Rings.Z Spaces.Int Algebra.Rings.CRing.
55
From HoTT Require Import Classes.interfaces.canonical_names.
66

77
Local Open Scope mc_scope.
@@ -38,7 +38,7 @@ Definition test2_B := Build_Matrix' nat 4 4
3838

3939
Definition test2 : entries test2_A = entries test2_B := idpath.
4040

41-
Local Open Scope binint_scope.
41+
Local Open Scope int_scope.
4242

4343
(** Matrices with ring entries can be multiplied *)
4444

theories/Algebra/AbGroups/AbelianGroup.v

Lines changed: 28 additions & 29 deletions
Original file line numberDiff line numberDiff line change
@@ -228,6 +228,20 @@ Proof.
228228
1-2: exact negate_involutive.
229229
Defined.
230230

231+
(** This goal comes up twice in the proof of [ab_mul], so we factor it out. *)
232+
Local Definition ab_mul_helper {A : AbGroup} (a b x y z : A)
233+
(p : x = y + z)
234+
: a + b + x = a + y + (b + z).
235+
Proof.
236+
lhs_V srapply grp_assoc.
237+
rhs_V srapply grp_assoc.
238+
apply grp_cancelL.
239+
rhs srapply grp_assoc.
240+
rhs nrapply (ap (fun g => g + z) (ab_comm y b)).
241+
rhs_V srapply grp_assoc.
242+
apply grp_cancelL, p.
243+
Defined.
244+
231245
(** Multiplication by [n : Int] defines an endomorphism of any abelian group [A]. *)
232246
Definition ab_mul {A : AbGroup} (n : Int) : GroupHomomorphism A A.
233247
Proof.
@@ -236,32 +250,18 @@ Proof.
236250
intros a b.
237251
induction n; cbn.
238252
- exact (grp_unit_l _)^.
239-
- destruct n.
240-
+ reflexivity.
241-
+ rhs_V srapply (associativity a).
242-
rhs srapply (ap _ (associativity _ b _)).
243-
rewrite (commutativity (nat_iter _ _ _) b).
244-
rhs_V srapply (ap _ (associativity b _ _)).
245-
rhs srapply (associativity a).
246-
apply grp_cancelL.
247-
exact IHn.
248-
- destruct n.
249-
+ rewrite (commutativity (-a)).
250-
exact (grp_inv_op a b).
251-
+ rhs_V srapply (associativity (-a)).
252-
rhs srapply (ap _ (associativity _ (-b) _)).
253-
rewrite (commutativity (nat_iter _ _ _) (-b)).
254-
rhs_V srapply (ap _ (associativity (-b) _ _)).
255-
rhs srapply (associativity (-a)).
256-
rewrite (commutativity (-a) (-b)), <- (grp_inv_op a b).
257-
apply grp_cancelL.
258-
exact IHn.
253+
- rewrite 3 grp_pow_succ.
254+
by apply ab_mul_helper.
255+
- rewrite 3 grp_pow_pred.
256+
rewrite (grp_inv_op a b), (commutativity (-b) (-a)).
257+
by apply ab_mul_helper.
259258
Defined.
260259

261-
Definition ab_mul_homo {A B : AbGroup}
260+
(** [ab_mul n] is natural. *)
261+
Definition ab_mul_natural {A B : AbGroup}
262262
(f : GroupHomomorphism A B) (n : Int)
263263
: f o ab_mul n == ab_mul n o f
264-
:= grp_pow_homo f n.
264+
:= grp_pow_natural f n.
265265

266266
(** The image of an inclusion is a normal subgroup. *)
267267
Definition ab_image_embedding {A B : AbGroup} (f : A $-> B) `{IsEmbedding f} : NormalSubgroup B
@@ -303,18 +303,17 @@ Proof.
303303
Defined.
304304

305305
(** 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 group. *)
306-
Definition ab_sum_const {A : AbGroup} (n : nat) (r : A)
307-
(f : forall k, (k < n)%nat -> A) (p : forall k Hk, f k Hk = r)
308-
: ab_sum n f = grp_pow r n.
306+
Definition ab_sum_const {A : AbGroup} (n : nat) (a : A)
307+
(f : forall k, (k < n)%nat -> A) (p : forall k Hk, f k Hk = a)
308+
: ab_sum n f = ab_mul n a.
309309
Proof.
310310
induction n as [|n IHn] in f, p |- *.
311311
- reflexivity.
312312
- rhs nrapply (ap@{Set _} _ (int_of_nat_succ_commute n)).
313-
rhs nrapply grp_pow_int_add_1.
313+
rhs nrapply grp_pow_succ.
314314
simpl. f_ap.
315-
rewrite IHn.
316-
+ reflexivity.
317-
+ intros. apply p.
315+
apply IHn.
316+
intros. apply p.
318317
Defined.
319318

320319
(** If the function is zero in the range of a finite sum then the sum is zero. *)

theories/Algebra/AbGroups/Cyclic.v

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,6 @@
11
Require Import Basics Types WildCat.Core Truncations.Core
22
AbelianGroup AbHom Centralizer AbProjective
3-
Groups.FreeGroup AbGroups.Z BinInt.Core Spaces.Int.
3+
Groups.FreeGroup AbGroups.Z Spaces.Int.
44

55
(** * Cyclic groups *)
66

@@ -48,7 +48,7 @@ Lemma Z1_mul_nat_beta {A : AbGroup} (a : A) (n : nat)
4848
Proof.
4949
induction n as [|n H].
5050
1: easy.
51-
refine (grp_pow_homo _ _ _ @ _); simpl.
51+
refine (grp_pow_natural _ _ _ @ _); simpl.
5252
by rewrite grp_unit_r.
5353
Defined.
5454

@@ -71,7 +71,7 @@ Defined.
7171

7272
(** The map sending the generator to [1 : Int]. *)
7373
Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z
74-
:= Z1_rec (G:=abgroup_Z) 1%binint.
74+
:= Z1_rec (G:=abgroup_Z) 1%int.
7575

7676
(** TODO: Prove that [Z1_to_Z] is a group isomorphism. *)
7777

theories/Algebra/AbGroups/Z.v

Lines changed: 17 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,5 @@
11
Require Import Basics.
2-
Require Import Spaces.Pos.Core Spaces.BinInt.
2+
Require Import Spaces.Pos.Core Spaces.Int.
33
Require Import Algebra.AbGroups.AbelianGroup.
44

55
Local Set Universe Minimization ToSet.
@@ -8,17 +8,25 @@ Local Set Universe Minimization ToSet.
88

99
(** See also Cyclic.v for a definition of the integers as the free group on one generator. *)
1010

11-
Local Open Scope binint_scope.
11+
Local Open Scope int_scope.
1212

13-
(** TODO: switch to [Int] *)
1413
Definition abgroup_Z@{} : AbGroup@{Set}.
1514
Proof.
1615
snrapply Build_AbGroup.
17-
- refine (Build_Group BinInt binint_add 0 binint_negation _); repeat split.
16+
- refine (Build_Group Int int_add 0 int_neg _); repeat split.
1817
+ exact _.
19-
+ exact binint_add_assoc.
20-
+ exact binint_add_0_r.
21-
+ exact binint_add_negation_l.
22-
+ exact binint_add_negation_r.
23-
- exact binint_add_comm.
18+
+ exact int_add_assoc.
19+
+ exact int_add_0_r.
20+
+ exact int_add_neg_l.
21+
+ exact int_add_neg_r.
22+
- exact int_add_comm.
23+
Defined.
24+
25+
(** For every group [G] and element [g : G], the map sending an integer to that power of [g] is a homomorphism. *)
26+
Definition grp_pow_homo {G : Group} (g : G)
27+
: GroupHomomorphism abgroup_Z G.
28+
Proof.
29+
snrapply Build_GroupHomomorphism.
30+
1: exact (grp_pow g).
31+
intros m n; apply grp_pow_add.
2432
Defined.

theories/Algebra/AbSES/SixTerm.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -211,7 +211,7 @@ Proof.
211211
apply moveR_equiv_V; symmetry.
212212
refine (ap f _ @ _).
213213
1: apply Z1_rec_beta.
214-
exact (ab_mul_homo f n Z1_gen).
214+
exact (ab_mul_natural f n Z1_gen).
215215
- (* we get rid of [equiv_Z1_hom] *)
216216
apply isexact_equiv_fiber.
217217
apply isexact_ext_cyclic_ab_iii.

theories/Algebra/Groups/Group.v

Lines changed: 34 additions & 71 deletions
Original file line numberDiff line numberDiff line change
@@ -476,100 +476,63 @@ End GroupMovement.
476476

477477
(** ** Power operation *)
478478

479-
(** For a given [n : nat] we can define the [n]th power of a group element. *)
479+
(** For a given [g : G] we can define the function [Int -> G] sending an integer to that power of [g]. *)
480480
Definition grp_pow {G : Group} (g : G) (n : Int) : G
481-
:= match n with
482-
| posS n => nat_iter n (g *.) g
483-
| zero => mon_unit
484-
| negS n => nat_iter n ((- g) *.) (- g)
485-
end.
486-
487-
(** Any homomorphism respects [grp_pow]. *)
488-
Lemma grp_pow_homo {G H : Group} (f : GroupHomomorphism G H) (n : Int) (g : G)
481+
:= int_iter (g *.) n mon_unit.
482+
483+
(** Any homomorphism respects [grp_pow]. In other words, [fun g => grp_pow g n] is natural. *)
484+
Lemma grp_pow_natural {G H : Group} (f : GroupHomomorphism G H) (n : Int) (g : G)
489485
: f (grp_pow g n) = grp_pow (f g) n.
490486
Proof.
491-
induction n.
492-
- apply grp_homo_unit.
493-
- destruct n; simpl in IHn |- *.
494-
+ reflexivity.
495-
+ lhs snrapply grp_homo_op.
496-
apply ap.
497-
exact IHn.
498-
- destruct n; simpl in IHn |- *.
499-
+ apply grp_homo_inv.
500-
+ lhs snrapply grp_homo_op.
501-
apply ap011.
502-
* apply grp_homo_inv.
503-
* exact IHn.
487+
lhs snrapply (int_iter_commute_map _ ((f g) *.)).
488+
1: nrapply grp_homo_op.
489+
apply (ap (int_iter _ n)), grp_homo_unit.
504490
Defined.
505491

506492
(** All powers of the unit are the unit. *)
507493
Definition grp_pow_unit {G : Group} (n : Int)
508494
: grp_pow (G:=G) mon_unit n = mon_unit.
509495
Proof.
510-
induction n.
496+
snrapply (int_iter_invariant n _ (fun g => g = mon_unit)); cbn.
497+
1, 2: apply paths_ind_r.
498+
- apply grp_unit_r.
499+
- lhs nrapply grp_unit_r. apply grp_inv_unit.
511500
- reflexivity.
512-
- destruct n; simpl.
513-
+ reflexivity.
514-
+ by lhs nrapply grp_unit_l.
515-
- destruct n; simpl in IHn |- *.
516-
+ apply grp_inv_unit.
517-
+ destruct IHn^.
518-
rhs_V apply (@grp_inv_unit G).
519-
apply grp_unit_r.
520501
Defined.
521502

522503
(** Note that powers don't preserve the group operation as it is not commutative. This does hold in an abelian group so such a result will appear later. *)
523504

524-
(** Helper functions for [grp_pow_int_add]. This is how we can unfold [grp_pow] once. *)
505+
(** The next two results tell us how [grp_pow] unfolds. *)
506+
Definition grp_pow_succ {G : Group} (n : Int) (g : G)
507+
: grp_pow g (n.+1)%int = g * grp_pow g n
508+
:= int_iter_succ_l _ _ _.
525509

526-
Definition grp_pow_int_add_1 {G : Group} (n : Int) (g : G)
527-
: grp_pow g (n.+1)%int = g * grp_pow g n.
528-
Proof.
529-
induction n; simpl.
530-
- exact (grp_unit_r g)^.
531-
- destruct n; reflexivity.
532-
- destruct n.
533-
+ apply (grp_inv_r g)^.
534-
+ rhs srapply associativity.
535-
rewrite grp_inv_r.
536-
apply (grp_unit_l _)^.
537-
Defined.
510+
Definition grp_pow_pred {G : Group} (n : Int) (g : G)
511+
: grp_pow g (n.-1)%int = (- g) * grp_pow g n
512+
:= int_iter_pred_l _ _ _.
538513

539-
Definition grp_pow_int_sub_1 {G : Group} (n : Int) (g : G)
540-
: grp_pow g (n.-1)%int = (- g) * grp_pow g n.
514+
(** [grp_pow] satisfies an additive law of exponents. *)
515+
Definition grp_pow_add {G : Group} (m n : Int) (g : G)
516+
: grp_pow g (n + m)%int = grp_pow g n * grp_pow g m.
541517
Proof.
542-
induction n; simpl.
543-
- exact (grp_unit_r (-g))^.
544-
- destruct n.
545-
+ apply (grp_inv_l g)^.
546-
+ rhs srapply associativity.
547-
rewrite grp_inv_l.
548-
apply (grp_unit_l _)^.
549-
- destruct n; reflexivity.
518+
lhs nrapply int_iter_add.
519+
induction n; cbn.
520+
1: exact (grp_unit_l _)^.
521+
1: rewrite int_iter_succ_l, grp_pow_succ.
522+
2: rewrite int_iter_pred_l, grp_pow_pred; cbn.
523+
1,2 : rhs_V srapply associativity;
524+
apply ap, IHn.
550525
Defined.
551526

552527
(** [grp_pow] commutes negative exponents to powers of the inverse *)
553-
Definition grp_pow_int_sign_commute {G : Group} (n : Int) (g : G)
528+
Definition grp_pow_neg {G : Group} (n : Int) (g : G)
554529
: grp_pow g (int_neg n) = grp_pow (- g) n.
555530
Proof.
556-
induction n.
531+
lhs nrapply int_iter_neg.
532+
destruct n.
533+
- cbn. by rewrite grp_inv_inv.
534+
- reflexivity.
557535
- reflexivity.
558-
- destruct n; reflexivity.
559-
- destruct n; simpl; rewrite grp_inv_inv; reflexivity.
560-
Defined.
561-
562-
(** [grp_pow] satisfies a law of exponents. *)
563-
Definition grp_pow_int_add {G : Group} (m n : Int) (g : G)
564-
: grp_pow g (n + m)%int = grp_pow g n * grp_pow g m.
565-
Proof.
566-
induction n; cbn.
567-
1: exact (grp_unit_l _)^.
568-
1: rewrite int_add_succ_l, 2 grp_pow_int_add_1.
569-
2: rewrite int_add_pred_l, 2 grp_pow_int_sub_1.
570-
1,2: rhs_V srapply associativity;
571-
snrapply (ap (_ *.));
572-
exact IHn.
573536
Defined.
574537

575538
(** ** The category of Groups *)

0 commit comments

Comments
 (0)