Skip to content

Commit c57fa3c

Browse files
committed
centrosymmetric matrices
Signed-off-by: Ali Caglayan <[email protected]>
1 parent 73b396a commit c57fa3c

File tree

1 file changed

+166
-0
lines changed

1 file changed

+166
-0
lines changed

theories/Algebra/Rings/Matrix.v

Lines changed: 166 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1018,6 +1018,172 @@ Defined.
10181018

10191019
(** Skew-symmetric matrices degenerate to symmetric matrices in rings with characteristic 2. In odd characteristic the module of matrices can be decomposed into the direct sum of symmetric and skew-symmetric matrices. *)
10201020

1021+
(** ** Exchange matrix *)
1022+
1023+
(** The exchange matrix is the matrix with ones on the anti-diagonal and zeros elsewhere. It will play an important role in defining classes of matrices with certain symmetric properties. *)
1024+
Definition exchange_matrix (R : Ring) n : Matrix R n n
1025+
:= Build_Matrix R n n (fun i j Hi Hj
1026+
=> kronecker_delta (R:=R) (i + j) (pred n))%nat.
1027+
1028+
(** Multiplying a matrix by the exchange matrix on the left is equivalent to exchanging the rows. Similarly multiplying on the right exchanges the columns. *)
1029+
1030+
(** The exchange matrix is invariant under transpose. *)
1031+
Definition exchange_matrix_transpose {R n}
1032+
: matrix_transpose (exchange_matrix R n) = exchange_matrix R n.
1033+
Proof.
1034+
apply path_matrix.
1035+
intros i j Hi Hj.
1036+
rewrite 3 entry_Build_Matrix.
1037+
by rewrite nat_add_comm.
1038+
Defined.
1039+
1040+
(** The exchange matrix is symmetric. *)
1041+
Global Instance issymmetric_exchange {R : Ring} {n : nat}
1042+
: IsSymmetric (exchange_matrix R n)
1043+
:= exchange_matrix_transpose.
1044+
1045+
(** The exchange matrix has order 2. This proof is only long because of arithmetic. *)
1046+
Definition exchange_matrix_square {R : Ring} {n : nat}
1047+
: matrix_mult (exchange_matrix R n) (exchange_matrix R n) = identity_matrix R n.
1048+
Proof.
1049+
apply path_matrix.
1050+
intros i j Hi Hj.
1051+
assert (r : (i <= pred n)%nat) by auto with nat.
1052+
rewrite 2 entry_Build_Matrix.
1053+
lhs nrapply path_ab_sum.
1054+
{ intros k Hk.
1055+
rewrite entry_Build_Matrix.
1056+
rewrite <- (nat_add_sub_eq _ r).
1057+
unshelve erewrite (kronecker_delta_map_inj _ _ (fun x => i + x)).
1058+
2: reflexivity.
1059+
intros x y H; exact (isinj_nat_add_l i x y H). }
1060+
unshelve erewrite rng_sum_kronecker_delta_l'.
1061+
{ unfold Core.lt.
1062+
destruct n.
1063+
1: by apply not_lt_n_0 in Hi.
1064+
auto with nat. }
1065+
rewrite entry_Build_Matrix.
1066+
set (t := (pred n - i + j)%nat).
1067+
rewrite <- (natminuspluseq _ _ r).
1068+
unfold t; clear t.
1069+
unshelve erewrite (kronecker_delta_map_inj j i (fun x => pred n - i + x)%nat).
1070+
2: apply kronecker_delta_symm.
1071+
intros x y H.
1072+
exact (isinj_nat_add_l _ _ _ H).
1073+
Defined.
1074+
1075+
(** ** Centrosymmetric matrices *)
1076+
1077+
(** A centrosymmetric matrix is symmetric about its center. *)
1078+
Class IsCentrosymmetric {R : Ring@{i}} {n : nat} (M : Matrix@{i} R n n) : Type@{i}
1079+
:= exchange_matrix_iscentrosymmetric
1080+
: matrix_mult (exchange_matrix R n) M = matrix_mult M (exchange_matrix R n).
1081+
1082+
Global Instance ishprop_iscentrosymmetric {R : Ring@{i}} {n : nat} (M : Matrix R n n)
1083+
: IsHProp (IsCentrosymmetric M).
1084+
Proof.
1085+
unfold IsCentrosymmetric; exact _.
1086+
Defined.
1087+
1088+
(** The zero matrix is centrosymmetric. *)
1089+
Global Instance iscentrosymmetric_matrix_zero {R : Ring@{i}} {n : nat}
1090+
: IsCentrosymmetric (matrix_zero R n n)
1091+
:= rng_mult_zero_r (A:=matrix_ring R n) _ @ (rng_mult_zero_l _)^.
1092+
1093+
(** The identity matrix is centrosymmetric. *)
1094+
Global Instance iscentrosymmetric_matrix_identity {R : Ring@{i}} {n : nat}
1095+
: IsCentrosymmetric (identity_matrix R n)
1096+
:= rng_mult_one_r (A:=matrix_ring R n) _ @ (rng_mult_one_l _)^.
1097+
1098+
(** The sum of two centrosymmetric matrices is centrosymmetric. *)
1099+
Global Instance iscentrosymmetric_matrix_plus {R : Ring@{i}} {n : nat}
1100+
(M N : Matrix R n n) {H1 : IsCentrosymmetric M} {H2 : IsCentrosymmetric N}
1101+
: IsCentrosymmetric (matrix_plus M N).
1102+
Proof.
1103+
unfold IsCentrosymmetric.
1104+
lhs nrapply (rng_dist_l (A:=matrix_ring R n)).
1105+
rhs nrapply (rng_dist_r (A:=matrix_ring R n)).
1106+
cbn; by rewrite H1, H2.
1107+
Defined.
1108+
1109+
(** The negation of a centrosymmetric matrix is centrosymmetric. *)
1110+
Global Instance iscentrosymmetric_matrix_negate {R : Ring@{i}} {n : nat}
1111+
(M : Matrix R n n) {H : IsCentrosymmetric M}
1112+
: IsCentrosymmetric (matrix_negate M).
1113+
Proof.
1114+
unfold IsCentrosymmetric.
1115+
lhs nrapply (rng_mult_negate_r (A:=matrix_ring R n) _).
1116+
rhs nrapply (rng_mult_negate_l (A:=matrix_ring R n) _).
1117+
f_ap.
1118+
Defined.
1119+
1120+
(** A scalar multiple of a centrosymmetric matrix is centrosymmetric. *)
1121+
Global Instance iscentrosymmetric_matrix_scale {R : Ring@{i}} {n : nat}
1122+
(r : R) (M : Matrix R n n) {H : IsCentrosymmetric M}
1123+
: IsCentrosymmetric (matrix_lact r M).
1124+
Proof.
1125+
unfold IsCentrosymmetric.
1126+
(** TODO: Need associative algebra structure. *)
1127+
Abort.
1128+
1129+
(** A centrosymmetric matrix is also centrosymmetric when considered over the opposite ring. *)
1130+
Global Instance iscentrosymmetric_rng_op {R : Ring@{i}} {n : nat} (M : Matrix R n n)
1131+
`{!IsCentrosymmetric M}
1132+
: IsCentrosymmetric (R := rng_op R) M.
1133+
Proof.
1134+
hnf.
1135+
apply path_matrix.
1136+
intros i j Hi Hj.
1137+
rewrite 2 entry_Build_Matrix.
1138+
pose (p := exchange_matrix_iscentrosymmetric).
1139+
apply (ap (fun x => entry x i j)) in p.
1140+
rewrite 2 entry_Build_Matrix in p.
1141+
refine (_ @ p @ _).
1142+
- apply path_ab_sum.
1143+
intros k Hk.
1144+
rewrite 2 entry_Build_Matrix.
1145+
rapply kronecker_delta_comm.
1146+
- apply path_ab_sum.
1147+
intros k Hk.
1148+
rewrite 2 entry_Build_Matrix.
1149+
rapply kronecker_delta_comm.
1150+
Defined.
1151+
1152+
(** The transpose of a centrosymmetric matrix is centrosymmetric. (Over a commutative ring. TODO: can we weaken CRing to Ring? *)
1153+
Global Instance iscentrosymmetric_matrix_transpose {R : Ring@{i}} {n : nat}
1154+
(M : Matrix R n n) {H : IsCentrosymmetric M}
1155+
: IsCentrosymmetric (matrix_transpose M).
1156+
Proof.
1157+
unfold IsCentrosymmetric.
1158+
rewrite <- exchange_matrix_transpose.
1159+
lhs_V nrapply (matrix_transpose_mult (R:=rng_op R) M (exchange_matrix R n)).
1160+
rhs_V nrapply (matrix_transpose_mult (R:=rng_op R) (exchange_matrix R n) M).
1161+
pose (iscentrosymmetric_rng_op M).
1162+
f_ap.
1163+
symmetry.
1164+
rapply (exchange_matrix_iscentrosymmetric (R:=rng_op R)).
1165+
Defined.
1166+
1167+
(** The product of two centrosymmetric matrices is centrosymmetric. *)
1168+
Global Instance iscentrosymmetric_matrix_mult {R : Ring@{i}} {n : nat}
1169+
(M N : Matrix R n n) {H1 : IsCentrosymmetric M} {H2 : IsCentrosymmetric N}
1170+
: IsCentrosymmetric (matrix_mult M N).
1171+
Proof.
1172+
lhs nrapply (rng_mult_assoc (A:=matrix_ring R n)).
1173+
rhs_V nrapply (rng_mult_assoc (A:=matrix_ring R n)).
1174+
hnf in H1, H2; cbn.
1175+
rewrite H1.
1176+
lhs_V nrapply (rng_mult_assoc (A:=matrix_ring R n)); f_ap.
1177+
Defined.
1178+
1179+
(** Centrosymmetric rings form a subring of the matrix ring. *)
1180+
Definition centrosymmetric_matrix_ring (R : Ring@{i}) (n : nat)
1181+
: Subring (matrix_ring R n).
1182+
Proof.
1183+
nrapply (Build_Subring' (fun M : matrix_ring R n => IsCentrosymmetric M));
1184+
cbn beta; exact _.
1185+
Defined.
1186+
10211187
Section MatrixCat.
10221188

10231189
(** The wild category [MatrixCat R] of [R]-valued matrices. This category has natural numbers as objects and m x n matrices as the arrows between [m] and [n]. *)

0 commit comments

Comments
 (0)