mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-12 23:28:39 +00:00
Lots of speedups (#116)
This commit is contained in:
@@ -7,13 +7,13 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.Cyclic.Definition {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) where
|
||||
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
open import Groups.Lemmas G
|
||||
|
||||
positiveEltPower : (x : A) (i : ℕ) → A
|
||||
positiveEltPower x 0 = Group.0G G
|
||||
@@ -40,9 +40,10 @@ elementPowerWellDefinedZ' i j i=j {g} = identityOfIndiscernablesRight _∼_ refl
|
||||
|
||||
elementPowerWellDefinedG : (g h : A) → (g ∼ h) → {n : ℤ} → (elementPower g n) ∼ (elementPower h n)
|
||||
elementPowerWellDefinedG g h g=h {nonneg n} = positiveEltPowerWellDefinedG g h g=h n
|
||||
elementPowerWellDefinedG g h g=h {negSucc x} = inverseWellDefined G (+WellDefined g=h (positiveEltPowerWellDefinedG g h g=h x))
|
||||
elementPowerWellDefinedG g h g=h {negSucc x} = inverseWellDefined (+WellDefined g=h (positiveEltPowerWellDefinedG g h g=h x))
|
||||
|
||||
record CyclicGroup : Set (m ⊔ n) where
|
||||
field
|
||||
generator : A
|
||||
cyclic : {a : A} → (Sg ℤ (λ i → Setoid._∼_ S (elementPower generator i) a))
|
||||
|
||||
|
@@ -15,15 +15,18 @@ module Groups.FreeProduct.Addition {i : _} {I : Set i} (decidableIndex : (x y :
|
||||
open import Groups.FreeProduct.Definition decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
||||
|
||||
prependEqual : (i : I) (g : A i) .(nonZero : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → (j : I) → i ≡ j → (w : ReducedSequenceBeginningWith j) → ReducedSequence
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) | inl sumZero = empty
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) | inr sumNotZero = nonempty i (ofEmpty i (_+_ i g h) sumNotZero)
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero x x₁) with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero {j} x x₁) | inl sumZero = nonempty j x
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero x pr) | inr sumNotZero = nonempty i (prependLetter i (_+_ i g h) sumNotZero x pr)
|
||||
|
||||
prepend : (i : I) → (g : A i) → .(nonZero : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → ReducedSequence → ReducedSequence
|
||||
prepend i g nonzero empty = nonempty i (ofEmpty i g nonzero)
|
||||
prepend i g nonzero (nonempty j x) with decidableIndex i j
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl | inl sumZero = empty
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl | inr sumNotZero = nonempty i (ofEmpty i (_+_ i g h) sumNotZero)
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero x x₁)) | inl refl with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero {j} x x₁)) | inl refl | inl sumZero = nonempty j x
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero x pr)) | inl refl | inr sumNotZero = nonempty i (prependLetter i (_+_ i g h) sumNotZero x pr)
|
||||
prepend i g nonzero (nonempty j w) | inl i=j = prependEqual i g nonzero j i=j w
|
||||
prepend i g nonzero (nonempty j x) | inr i!=j = nonempty i (prependLetter i g nonzero x i!=j)
|
||||
|
||||
plus' : {j : _} → (ReducedSequenceBeginningWith j) → ReducedSequence → ReducedSequence
|
||||
|
@@ -1,5 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Setoids
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Groups.Definition
|
||||
@@ -12,6 +13,8 @@ record GroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A →
|
||||
field
|
||||
groupHom : {x y : A} → f (x ·A y) ∼ (f x) ·B (f y)
|
||||
wellDefined : {x y : A} → Setoid._∼_ S x y → f x ∼ f y
|
||||
groupHom' : {x y : A} → (f x) ·B (f y) ∼ f (x ·A y)
|
||||
groupHom' = Equivalence.symmetric eq groupHom
|
||||
|
||||
record InjectiveGroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underf : A → B} (f : GroupHom G H underf) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
|
@@ -42,3 +42,9 @@ zeroImpliesInverseZero {x} fx=0 = transitive homRespectsInverse (transitive (inv
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
homRespectsInverse' : {a b : A} → Setoid._∼_ T (f (Group.inverse G a) +B f (Group.inverse G b)) (Group.inverse H (f (b +A a)))
|
||||
homRespectsInverse' {a} {b} = transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Equivalence.symmetric (Setoid.eq S) (invContravariant G))) (homRespectsInverse))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
20
Groups/Orders/Archimedean.agda
Normal file
20
Groups/Orders/Archimedean.agda
Normal file
@@ -0,0 +1,20 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Numbers.Naturals.Definition
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Orders.Archimedean {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {G : Group S _+_} {c : _} {_<_ : A → A → Set c} {pOrder : SetoidPartialOrder S _<_} (p : PartiallyOrderedGroup G pOrder) where
|
||||
|
||||
open Setoid S
|
||||
open import Groups.Cyclic.Definition G
|
||||
open Group G
|
||||
|
||||
Archimedean : Set (a ⊔ c)
|
||||
Archimedean = (x y : A) → (0G < x) → (0G < y) → Sg ℕ (λ n → y < (positiveEltPower x n))
|
17
Groups/Orders/Partial/Definition.agda
Normal file
17
Groups/Orders/Partial/Definition.agda
Normal file
@@ -0,0 +1,17 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Orders.Partial.Definition {n m : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
open Group G
|
||||
open Setoid S
|
||||
|
||||
record PartiallyOrderedGroup {p : _} {_<_ : Rel {_} {p} A} (pOrder : SetoidPartialOrder S _<_) : Set (lsuc n ⊔ m ⊔ p) where
|
||||
field
|
||||
orderRespectsAddition : {a b : A} → (a < b) → (c : A) → (a + c) < (b + c)
|
Reference in New Issue
Block a user