mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 15:48:39 +00:00
Modules
This commit is contained in:
25
Modules/Definition.agda
Normal file
25
Modules/Definition.agda
Normal file
@@ -0,0 +1,25 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+R_ _*_) {m n : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M → M → M} {G' : Group T _+_} (G : AbelianGroup G') (_·_ : A → M → M) where
|
||||
|
||||
record Module : Set (a ⊔ b ⊔ m ⊔ n) where
|
||||
field
|
||||
dotWellDefined : {r s : A} {t u : M} → Setoid._∼_ S r s → Setoid._∼_ T t u → Setoid._∼_ T (r · t) (s · u)
|
||||
dotDistributesLeft : {r : A} {x y : M} → Setoid._∼_ T (r · (x + y)) ((r · x) + (r · y))
|
||||
dotDistributesRight : {r s : A} {x : M} → Setoid._∼_ T ((r +R s) · x) ((r · x) + (s · x))
|
||||
dotAssociative : {r s : A} {x : M} → Setoid._∼_ T ((r * s) · x) (r · (s · x))
|
||||
dotIdentity : {x : M} → Setoid._∼_ T ((Ring.1R R) · x) x
|
34
Modules/Examples.agda
Normal file
34
Modules/Examples.agda
Normal file
@@ -0,0 +1,34 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Modules.Definition
|
||||
open import Groups.Cyclic.Definition
|
||||
open import Groups.Cyclic.DefinitionLemmas
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.Examples where
|
||||
|
||||
abGrpModule : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {G' : Group S _+_} (G : AbelianGroup G') → Module ℤRing G (λ x i → elementPower G' i x)
|
||||
Module.dotWellDefined (abGrpModule {S = S} {G' = G'} G) {m} {n} {g} {h} m=n g=h = transitive (elementPowerWellDefinedG G' g h g=h {m}) (elementPowerWellDefinedZ' G' m n m=n {h})
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
Module.dotDistributesLeft (abGrpModule {G' = G'} G) {n} {x} {y} = elementPowerHomAbelian G' (AbelianGroup.commutative G) x y n
|
||||
Module.dotDistributesRight (abGrpModule {S = S} {G' = G'} G) {r} {s} {x} = symmetric (elementPowerCollapse G' x r s)
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
Module.dotAssociative (abGrpModule {G' = G'} G) {r} {s} {x} = elementPowerMultiplies G' r s x
|
||||
Module.dotIdentity (abGrpModule {G' = G'} G) = Group.identRight G'
|
39
Modules/Lemmas.agda
Normal file
39
Modules/Lemmas.agda
Normal file
@@ -0,0 +1,39 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Modules.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Modules.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+R_ _*_} {m n : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M → M → M} {G' : Group T _+_} {G : AbelianGroup G'} {_·_ : A → M → M} (mod : Module R G _·_) where
|
||||
|
||||
open Group G'
|
||||
open Ring R
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
open Module mod
|
||||
|
||||
moduleTimesZero : {x : M} → (0R · x) ∼ 0G
|
||||
moduleTimesZero {x} = equalsDoubleImpliesZero G' (symmetric x=2x)
|
||||
where
|
||||
x=2x : (0R · x) ∼ (0R · x) + (0R · x)
|
||||
x=2x = transitive (dotWellDefined (Equivalence.symmetric (Setoid.eq S) (Group.identLeft additiveGroup)) reflexive) dotDistributesRight
|
||||
|
||||
moduleTimes-1 : {x : M} → ((Group.inverse additiveGroup 1R) · x) ∼ inverse x
|
||||
moduleTimes-1 {x} = transitive (transferToRight' G' j) (inverseWellDefined G' dotIdentity)
|
||||
where
|
||||
i : ((1R · x) + ((Group.inverse additiveGroup 1R) · x)) ∼ 0G
|
||||
i = transitive (symmetric (transitive (dotWellDefined (Equivalence.symmetric (Setoid.eq S) (Group.invRight additiveGroup {1R})) reflexive) dotDistributesRight)) (moduleTimesZero)
|
||||
j : (((Group.inverse additiveGroup 1R) · x) + (1R · x)) ∼ 0G
|
||||
j = transitive (AbelianGroup.commutative G) i
|
Reference in New Issue
Block a user