Skip to content

Commit bb7b51b

Browse files
jamesmckinnaandreasabel
authored andcommitted
Nagata's "idealization of a module" construction (#2244)
* Nagata's construction * removed redundant `zero` * first round of Jacques' review comments * proved the additional properties * some of Matthew's suggestions, plus more vertical whitespace, less horizontal; more comments * Matthew's suggestion: using `private` modules * Matthew's suggestion: lifting out left-/right- sublemmas * standardised names, as far as possible * Matthew's suggestion: lifting out left-/right- sublemmas * fixed constraint problem with ambiguous symbol; renamed ideal lemmas * renamed module * renamed module in `CHANGELOG` * added generalised annihilation lemma * typos * use correct rexported names * now as a paramterised module instead * or did you intend this? * fix whitespace * aligned one step of reasoning * more re-alignment of reasoning steps * more re-alignment of reasoning steps * Matthew's review comments * blanklines
1 parent 3146ae5 commit bb7b51b

File tree

2 files changed

+199
-0
lines changed

2 files changed

+199
-0
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -48,6 +48,11 @@ New modules
4848

4949
* `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures
5050

51+
* Nagata's construction of the "idealization of a module":
52+
```agda
53+
Algebra.Module.Construct.Idealization
54+
```
55+
5156
Additions to existing modules
5257
-----------------------------
5358

Original file line numberDiff line numberDiff line change
@@ -0,0 +1,194 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The non-commutative analogue of Nagata's construction of
5+
-- the "idealization of a module", (Local Rings, 1962; Wiley)
6+
-- defined here on R-R-*bi*modules M over a ring R, as used in
7+
-- "Forward- or reverse-mode automatic differentiation: What's the difference?"
8+
-- (Van den Berg, Schrijvers, McKinna, Vandenbroucke;
9+
-- Science of Computer Programming, Vol. 234, January 2024
10+
-- https://doi.org/10.1016/j.scico.2023.103010)
11+
--
12+
-- The construction N =def R ⋉ M , for which there is unfortunately
13+
-- no consistent notation in the literature, consists of:
14+
-- * carrier: pairs |R| × |M|
15+
-- * with additive structure that of the direct sum R ⊕ M _of modules_
16+
-- * but with multiplication _*_ such that M forms an _ideal_ of N
17+
-- * moreover satisfying 'm * m ≈ 0' for every m ∈ M ⊆ N
18+
--
19+
-- The fundamental lemma (proved here) is that N, in fact, defines a Ring:
20+
-- this ring is essentially the 'ring of dual numbers' construction R[M]
21+
-- (Clifford, 1874; generalised!) for an ideal M, and thus the synthetic/algebraic
22+
-- analogue of the tangent space of M (considered as a 'vector space' over R)
23+
-- in differential geometry, hence its application to Automatic Differentiation.
24+
--
25+
-- Nagata's more fundamental insight (not yet shown here) is that
26+
-- the lattice of R-submodules of M is in order-isomorphism with
27+
-- the lattice of _ideals_ of R ⋉ M , and hence that the study of
28+
-- modules can be reduced to that of ideals of a ring, and vice versa.
29+
--
30+
------------------------------------------------------------------------
31+
32+
{-# OPTIONS --cubical-compatible --safe #-}
33+
34+
open import Algebra.Bundles using (AbelianGroup; Ring)
35+
open import Algebra.Module.Bundles using (Bimodule)
36+
37+
module Algebra.Module.Construct.Idealization
38+
{r ℓr m ℓm} (ring : Ring r ℓr) (bimodule : Bimodule ring ring m ℓm) where
39+
40+
open import Algebra.Core
41+
import Algebra.Consequences.Setoid as Consequences
42+
import Algebra.Definitions as Definitions
43+
import Algebra.Module.Construct.DirectProduct as DirectProduct
44+
import Algebra.Module.Construct.TensorUnit as TensorUnit
45+
open import Algebra.Structures using (IsAbelianGroup; IsRing)
46+
open import Data.Product using (_,_; ∃-syntax)
47+
open import Level using (Level; _⊔_)
48+
open import Relation.Binary using (Rel; Setoid; IsEquivalence)
49+
import Relation.Binary.Reasoning.Setoid as ≈-Reasoning
50+
51+
------------------------------------------------------------------------
52+
-- Definitions
53+
54+
private
55+
open module R = Ring ring
56+
using ()
57+
renaming (Carrier to R)
58+
59+
open module M = Bimodule bimodule
60+
renaming (Carrierᴹ to M)
61+
62+
+ᴹ-middleFour = Consequences.comm∧assoc⇒middleFour ≈ᴹ-setoid +ᴹ-cong +ᴹ-comm +ᴹ-assoc
63+
64+
open module N = Bimodule (DirectProduct.bimodule TensorUnit.bimodule bimodule)
65+
using ()
66+
renaming ( Carrierᴹ to N
67+
; _≈ᴹ_ to _≈_
68+
; _+ᴹ_ to _+_
69+
; 0ᴹ to 0#
70+
; -ᴹ_ to -_
71+
; +ᴹ-isAbelianGroup to +-isAbelianGroup
72+
)
73+
74+
open AbelianGroup M.+ᴹ-abelianGroup hiding (_≈_)
75+
open ≈-Reasoning ≈ᴹ-setoid
76+
open Definitions _≈_
77+
78+
-- Injections ι from the components of the direct sum
79+
-- ιᴹ in fact exhibits M as an _ideal_ of R ⋉ M (see below)
80+
ιᴿ : R N
81+
ιᴿ r = r , 0ᴹ
82+
83+
ιᴹ : M N
84+
ιᴹ m = R.0# , m
85+
86+
-- Multiplicative unit
87+
88+
1# : N
89+
1# = ιᴿ R.1#
90+
91+
-- Multiplication
92+
93+
infixl 7 _*_
94+
95+
_*_ : Op₂ N
96+
(r₁ , m₁) * (r₂ , m₂) = r₁ R.* r₂ , r₁ *ₗ m₂ +ᴹ m₁ *ᵣ r₂
97+
98+
-- Properties: because we work in the direct sum, every proof has
99+
-- * an 'R'-component, which inherits directly from R, and
100+
-- * an 'M'-component, where the work happens
101+
102+
*-cong : Congruent₂ _*_
103+
*-cong (r₁ , m₁) (r₂ , m₂) = R.*-cong r₁ r₂ , +ᴹ-cong (*ₗ-cong r₁ m₂) (*ᵣ-cong m₁ r₂)
104+
105+
*-identityˡ : LeftIdentity 1# _*_
106+
*-identityˡ (r , m) = R.*-identityˡ r , (begin
107+
R.1# *ₗ m +ᴹ 0ᴹ *ᵣ r ≈⟨ +ᴹ-cong (*ₗ-identityˡ m) (*ᵣ-zeroˡ r) ⟩
108+
m +ᴹ 0ᴹ ≈⟨ +ᴹ-identityʳ m ⟩
109+
m ∎)
110+
111+
*-identityʳ : RightIdentity 1# _*_
112+
*-identityʳ (r , m) = R.*-identityʳ r , (begin
113+
r *ₗ 0ᴹ +ᴹ m *ᵣ R.1# ≈⟨ +ᴹ-cong (*ₗ-zeroʳ r) (*ᵣ-identityʳ m) ⟩
114+
0ᴹ +ᴹ m ≈⟨ +ᴹ-identityˡ m ⟩
115+
m ∎)
116+
117+
*-identity : Identity 1# _*_
118+
*-identity = *-identityˡ , *-identityʳ
119+
120+
*-assoc : Associative _*_
121+
*-assoc (r₁ , m₁) (r₂ , m₂) (r₃ , m₃) = R.*-assoc r₁ r₂ r₃ , (begin
122+
(r₁ R.* r₂) *ₗ m₃ +ᴹ (r₁ *ₗ m₂ +ᴹ m₁ *ᵣ r₂) *ᵣ r₃
123+
≈⟨ +ᴹ-cong (*ₗ-assoc r₁ r₂ m₃) (*ᵣ-distribʳ r₃ (r₁ *ₗ m₂) (m₁ *ᵣ r₂)) ⟩
124+
r₁ *ₗ (r₂ *ₗ m₃) +ᴹ ((r₁ *ₗ m₂) *ᵣ r₃ +ᴹ (m₁ *ᵣ r₂) *ᵣ r₃)
125+
≈⟨ +ᴹ-congˡ (+ᴹ-congʳ (*ₗ-*ᵣ-assoc r₁ m₂ r₃)) ⟩
126+
r₁ *ₗ (r₂ *ₗ m₃) +ᴹ (r₁ *ₗ (m₂ *ᵣ r₃) +ᴹ (m₁ *ᵣ r₂) *ᵣ r₃)
127+
≈⟨ +ᴹ-assoc (r₁ *ₗ (r₂ *ₗ m₃)) (r₁ *ₗ (m₂ *ᵣ r₃)) ((m₁ *ᵣ r₂) *ᵣ r₃) ⟨
128+
(r₁ *ₗ (r₂ *ₗ m₃) +ᴹ r₁ *ₗ (m₂ *ᵣ r₃)) +ᴹ (m₁ *ᵣ r₂) *ᵣ r₃
129+
≈⟨ +ᴹ-cong (≈ᴹ-sym (*ₗ-distribˡ r₁ (r₂ *ₗ m₃) (m₂ *ᵣ r₃))) (*ᵣ-assoc m₁ r₂ r₃) ⟩
130+
r₁ *ₗ (r₂ *ₗ m₃ +ᴹ m₂ *ᵣ r₃) +ᴹ m₁ *ᵣ (r₂ R.* r₃) ∎)
131+
132+
distribˡ : _*_ DistributesOverˡ _+_
133+
distribˡ (r₁ , m₁) (r₂ , m₂) (r₃ , m₃) = R.distribˡ r₁ r₂ r₃ , (begin
134+
r₁ *ₗ (m₂ +ᴹ m₃) +ᴹ m₁ *ᵣ (r₂ R.+ r₃)
135+
≈⟨ +ᴹ-cong (*ₗ-distribˡ r₁ m₂ m₃) (*ᵣ-distribˡ m₁ r₂ r₃) ⟩
136+
(r₁ *ₗ m₂ +ᴹ r₁ *ₗ m₃) +ᴹ (m₁ *ᵣ r₂ +ᴹ m₁ *ᵣ r₃)
137+
≈⟨ +ᴹ-middleFour (r₁ *ₗ m₂) (r₁ *ₗ m₃) (m₁ *ᵣ r₂) (m₁ *ᵣ r₃) ⟩
138+
(r₁ *ₗ m₂ +ᴹ m₁ *ᵣ r₂) +ᴹ (r₁ *ₗ m₃ +ᴹ m₁ *ᵣ r₃) ∎)
139+
140+
141+
distribʳ : _*_ DistributesOverʳ _+_
142+
distribʳ (r₁ , m₁) (r₂ , m₂) (r₃ , m₃) = R.distribʳ r₁ r₂ r₃ , (begin
143+
(r₂ R.+ r₃) *ₗ m₁ +ᴹ (m₂ +ᴹ m₃) *ᵣ r₁
144+
≈⟨ +ᴹ-cong (*ₗ-distribʳ m₁ r₂ r₃) (*ᵣ-distribʳ r₁ m₂ m₃) ⟩
145+
(r₂ *ₗ m₁ +ᴹ r₃ *ₗ m₁) +ᴹ (m₂ *ᵣ r₁ +ᴹ m₃ *ᵣ r₁)
146+
≈⟨ +ᴹ-middleFour (r₂ *ₗ m₁) (r₃ *ₗ m₁) (m₂ *ᵣ r₁) (m₃ *ᵣ r₁) ⟩
147+
(r₂ *ₗ m₁ +ᴹ m₂ *ᵣ r₁) +ᴹ (r₃ *ₗ m₁ +ᴹ m₃ *ᵣ r₁) ∎)
148+
149+
distrib : _*_ DistributesOver _+_
150+
distrib = distribˡ , distribʳ
151+
152+
153+
------------------------------------------------------------------------
154+
-- The Fundamental Lemma
155+
156+
-- Structure
157+
158+
isRingᴺ : IsRing _≈_ _+_ _*_ -_ 0# 1#
159+
isRingᴺ = record
160+
{ +-isAbelianGroup = +-isAbelianGroup
161+
; *-cong = *-cong
162+
; *-assoc = *-assoc
163+
; *-identity = *-identity
164+
; distrib = distrib
165+
}
166+
167+
-- Bundle
168+
169+
ringᴺ : Ring (r ⊔ m) (ℓr ⊔ ℓm)
170+
ringᴺ = record { isRing = isRingᴺ }
171+
172+
------------------------------------------------------------------------
173+
-- M is an ideal of R ⋉ M satisfying m₁ * m₂ ≈ 0#
174+
175+
ιᴹ-idealˡ : (n : N) (m : M) ∃[ n*m ] n * ιᴹ m ≈ ιᴹ n*m
176+
ιᴹ-idealˡ n@(r , _) m = _ , R.zeroʳ r , ≈ᴹ-refl
177+
178+
ιᴹ-idealʳ : (m : M) (n : N) ∃[ m*n ] ιᴹ m * n ≈ ιᴹ m*n
179+
ιᴹ-idealʳ m n@(r , _) = _ , R.zeroˡ r , ≈ᴹ-refl
180+
181+
*-annihilates-ιᴹ : (m₁ m₂ : M) ιᴹ m₁ * ιᴹ m₂ ≈ 0#
182+
*-annihilates-ιᴹ m₁ m₂ = R.zeroˡ R.0# , (begin
183+
R.0# *ₗ m₂ +ᴹ m₁ *ᵣ R.0# ≈⟨ +ᴹ-cong (*ₗ-zeroˡ m₂) (*ᵣ-zeroʳ m₁) ⟩
184+
0ᴹ +ᴹ 0ᴹ ≈⟨ +ᴹ-identityˡ 0ᴹ ⟩
185+
0ᴹ ∎)
186+
187+
m*m≈0 : (m : M) ιᴹ m * ιᴹ m ≈ 0#
188+
m*m≈0 m = *-annihilates-ιᴹ m m
189+
190+
------------------------------------------------------------------------
191+
-- Infix notation for when opening the module unparameterised
192+
193+
infixl 4 _⋉_
194+
_⋉_ = ringᴺ

0 commit comments

Comments
 (0)