mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 07:38:40 +00:00
Define the free product (#104)
This commit is contained in:
31
Groups/FreeGroup/Lemmas.agda
Normal file
31
Groups/FreeGroup/Lemmas.agda
Normal file
@@ -0,0 +1,31 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Setoids
|
||||
open import Groups.FreeGroup.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Decidable.Sets
|
||||
open import Numbers.Naturals.Order
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Functions
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.FreeGroup.Word
|
||||
open import Groups.FreeGroup.Group
|
||||
open import Groups.FreeGroup.UniversalProperty
|
||||
|
||||
module Groups.FreeGroup.Lemmas {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
freeGroupFunctorWellDefined : {b : _} {B : Set b} (decB : DecidableSet B) → {f : A → B} → Bijection f → GroupsIsomorphic (freeGroup decA) (freeGroup decB)
|
||||
GroupsIsomorphic.isomorphism (freeGroupFunctorWellDefined decB {f} bij) = universalPropertyFunction decA (freeGroup decB) λ a → freeEmbedding decB (f a)
|
||||
GroupIso.groupHom (GroupsIsomorphic.proof (freeGroupFunctorWellDefined decB {f} bij)) = universalPropertyHom decA (freeGroup decB) λ a → freeEmbedding decB (f a)
|
||||
GroupIso.bij (GroupsIsomorphic.proof (freeGroupFunctorWellDefined decB {f} bij)) = {!!}
|
||||
|
||||
private
|
||||
subgroupGeneratedBySquares : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} (G : Group S _+_) → Set
|
||||
sub
|
||||
|
||||
|
||||
freeGroupFunctorInjective : {b : _} {B : Set b} (decB : DecidableSet B) → GroupsIsomorphic (freeGroup decA) (freeGroup decB) → Sg (A → B) (λ f → Bijection f)
|
||||
freeGroupFunctorInjective decB iso = {!!}
|
163
Groups/FreeGroup/Parity.agda
Normal file
163
Groups/FreeGroup/Parity.agda
Normal file
@@ -0,0 +1,163 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Setoids
|
||||
open import Groups.FreeGroup.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Decidable.Sets
|
||||
open import Numbers.Naturals.Order
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Functions
|
||||
open import Groups.Isomorphisms.Definition
|
||||
|
||||
module Groups.FreeGroup.Parity {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
open import Groups.FreeGroup.Word decA
|
||||
open import Groups.FreeGroup.Group decA
|
||||
|
||||
xor : Bool → Bool → Bool
|
||||
xor BoolTrue BoolTrue = BoolFalse
|
||||
xor BoolTrue BoolFalse = BoolTrue
|
||||
xor BoolFalse b = b
|
||||
|
||||
C2 : Group (reflSetoid Bool) xor
|
||||
Group.+WellDefined C2 refl refl = refl
|
||||
Group.0G C2 = BoolFalse
|
||||
Group.inverse C2 = id
|
||||
Group.+Associative C2 {BoolTrue} {BoolTrue} {BoolTrue} = refl
|
||||
Group.+Associative C2 {BoolTrue} {BoolTrue} {BoolFalse} = refl
|
||||
Group.+Associative C2 {BoolTrue} {BoolFalse} {z} = refl
|
||||
Group.+Associative C2 {BoolFalse} {y} {z} = refl
|
||||
Group.identRight C2 {BoolTrue} = refl
|
||||
Group.identRight C2 {BoolFalse} = refl
|
||||
Group.identLeft C2 {x} = refl
|
||||
Group.invLeft C2 {BoolTrue} = refl
|
||||
Group.invLeft C2 {BoolFalse} = refl
|
||||
Group.invRight C2 {BoolTrue} = refl
|
||||
Group.invRight C2 {BoolFalse} = refl
|
||||
|
||||
notNot : (x : Bool) → not (not x) ≡ x
|
||||
notNot BoolTrue = refl
|
||||
notNot BoolFalse = refl
|
||||
|
||||
notWellDefined : {x y : Bool} → (x ≡ y) → not x ≡ not y
|
||||
notWellDefined {BoolTrue} {BoolTrue} x=y = refl
|
||||
notWellDefined {BoolFalse} {BoolFalse} x=y = refl
|
||||
|
||||
parity : (x : A) → ReducedWord → Bool
|
||||
parity x empty = BoolFalse
|
||||
parity x (prependLetter (ofLetter y) w _) with DecidableSet.eq decA x y
|
||||
parity x (prependLetter (ofLetter y) w _) | inl _ = not (parity x w)
|
||||
parity x (prependLetter (ofLetter y) w _) | inr _ = parity x w
|
||||
parity x (prependLetter (ofInv y) w _) with DecidableSet.eq decA x y
|
||||
parity x (prependLetter (ofInv y) w _) | inl _ = not (parity x w)
|
||||
parity x (prependLetter (ofInv y) w _) | inr _ = parity x w
|
||||
|
||||
parityPrepend : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofLetter l)) ≡ parity a w
|
||||
parityPrepend a empty l notEq with DecidableSet.eq decA a l
|
||||
parityPrepend a empty l notEq | inl x = exFalso (notEq x)
|
||||
parityPrepend a empty l notEq | inr x = refl
|
||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq with DecidableSet.eq decA a l
|
||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inl m = exFalso (notEq m)
|
||||
parityPrepend a (prependLetter (ofLetter r) w x) l notEq | inr _ = refl
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq with DecidableSet.eq decA a r
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r with DecidableSet.eq decA l r
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with DecidableSet.eq decA a l
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with DecidableSet.eq decA a r
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r with DecidableSet.eq decA l r
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl x₁ = refl
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with DecidableSet.eq decA a l
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with DecidableSet.eq decA a r
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
||||
parityPrepend a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
||||
|
||||
parityPrepend' : (a : A) (w : ReducedWord) (l : A) → ((a ≡ l) → False) → parity a (prepend w (ofInv l)) ≡ parity a w
|
||||
parityPrepend' a empty l notEq with DecidableSet.eq decA a l
|
||||
parityPrepend' a empty l notEq | inl x = exFalso (notEq x)
|
||||
parityPrepend' a empty l notEq | inr x = refl
|
||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq with DecidableSet.eq decA l r
|
||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inl m with DecidableSet.eq decA a r
|
||||
... | inl a=r = exFalso (notEq (transitivity a=r (equalityCommutative m)))
|
||||
... | inr a!=r = refl
|
||||
parityPrepend' a (prependLetter (ofLetter r) w x) l notEq | inr l!=r with DecidableSet.eq decA a l
|
||||
... | inl a=l = exFalso (notEq a=l)
|
||||
... | inr a!=l = refl
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq with DecidableSet.eq decA a r
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r with DecidableSet.eq decA l r
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inl l=r = exFalso (notEq (transitivity a=r (equalityCommutative l=r)))
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r with DecidableSet.eq decA a l
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l with DecidableSet.eq decA a r
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inl x₁ = refl
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inl a=r | inr l!=r | inr a!=l | inr bad = exFalso (bad a=r)
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r with DecidableSet.eq decA l r
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r with DecidableSet.eq decA a l
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inl a=l = exFalso (notEq a=l)
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inl l=r | inr a!=l with DecidableSet.eq decA a r
|
||||
... | inl a=r = exFalso (a!=r a=r)
|
||||
... | inr _ = refl
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r with DecidableSet.eq decA a l
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inl a=l = exFalso (notEq a=l)
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l with DecidableSet.eq decA a r
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inl a=r = exFalso (a!=r a=r)
|
||||
parityPrepend' a (prependLetter (ofInv r) w x) l notEq | inr a!=r | inr l!=r | inr a!=l | inr x₁ = refl
|
||||
|
||||
parityPrepend'' : (a : A) (w : ReducedWord) → parity a (prepend w (ofLetter a)) ≡ not (parity a w)
|
||||
parityPrepend'' a empty with DecidableSet.eq decA a a
|
||||
... | inl _ = refl
|
||||
... | inr bad = exFalso (bad refl)
|
||||
parityPrepend'' a (prependLetter (ofLetter l) w x) with DecidableSet.eq decA a a
|
||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ with DecidableSet.eq decA a l
|
||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inl a=l = refl
|
||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inl _ | inr a!=l = refl
|
||||
parityPrepend'' a (prependLetter (ofLetter l) w x) | inr bad = exFalso (bad refl)
|
||||
parityPrepend'' a (prependLetter (ofInv l) w x) with DecidableSet.eq decA a l
|
||||
... | inl a=l = equalityCommutative (notNot (parity a w))
|
||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l with DecidableSet.eq decA a a
|
||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inl _ with DecidableSet.eq decA a l
|
||||
... | inl a=l = exFalso (a!=l a=l)
|
||||
... | inr _ = refl
|
||||
parityPrepend'' a (prependLetter (ofInv l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
||||
|
||||
parityPrepend''' : (a : A) (w : ReducedWord) → parity a (prepend w (ofInv a)) ≡ not (parity a w)
|
||||
parityPrepend''' a empty with DecidableSet.eq decA a a
|
||||
... | inl _ = refl
|
||||
... | inr bad = exFalso (bad refl)
|
||||
parityPrepend''' a (prependLetter (ofLetter l) w x) with DecidableSet.eq decA a l
|
||||
... | inl a=l = equalityCommutative (notNot _)
|
||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l with DecidableSet.eq decA a a
|
||||
... | inl _ with DecidableSet.eq decA a l
|
||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inl a=l = exFalso (a!=l a=l)
|
||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inl _ | inr _ = refl
|
||||
parityPrepend''' a (prependLetter (ofLetter l) w x) | inr a!=l | inr bad = exFalso (bad refl)
|
||||
parityPrepend''' a (prependLetter (ofInv l) w x) with DecidableSet.eq decA a a
|
||||
parityPrepend''' a (prependLetter (ofInv l) w x) | inl _ with DecidableSet.eq decA a l
|
||||
... | inl a=l = refl
|
||||
... | inr a!=l = refl
|
||||
parityPrepend''' a (prependLetter (ofInv l) w x) | inr bad = exFalso (bad refl)
|
||||
|
||||
notXor : (x y : Bool) → not (xor x y) ≡ xor (not x) y
|
||||
notXor BoolTrue BoolTrue = refl
|
||||
notXor BoolTrue BoolFalse = refl
|
||||
notXor BoolFalse BoolTrue = refl
|
||||
notXor BoolFalse BoolFalse = refl
|
||||
|
||||
parityHomIsHom : (a : A) → (x y : ReducedWord) → parity a (_+W_ x y) ≡ xor (parity a x) (parity a y)
|
||||
parityHomIsHom a empty y = refl
|
||||
parityHomIsHom a (prependLetter (ofLetter l) x _) y with DecidableSet.eq decA a l
|
||||
parityHomIsHom a (prependLetter (ofLetter .a) x _) y | inl refl = transitivity (parityPrepend'' a (x +W y)) (transitivity (notWellDefined (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
||||
parityHomIsHom a (prependLetter (ofLetter l) x _) y | inr a!=l = transitivity (parityPrepend a (_+W_ x y) l a!=l) (parityHomIsHom a x y)
|
||||
parityHomIsHom a (prependLetter (ofInv l) x _) y with DecidableSet.eq decA a l
|
||||
parityHomIsHom a (prependLetter (ofInv .a) x _) y | inl refl = transitivity (parityPrepend''' a (x +W y)) (transitivity (notWellDefined (parityHomIsHom a x y)) (notXor (parity a x) (parity a y)))
|
||||
parityHomIsHom a (prependLetter (ofInv l) x _) y | inr a!=l = transitivity (parityPrepend' a (x +W y) l a!=l) (parityHomIsHom a x y)
|
||||
|
||||
parityHom : (x : A) → GroupHom (freeGroup) C2 (parity x)
|
||||
GroupHom.groupHom (parityHom x) {y} {z} = parityHomIsHom x y z
|
||||
GroupHom.wellDefined (parityHom x) refl = refl
|
383
Groups/FreeProduct/Addition.agda
Normal file
383
Groups/FreeProduct/Addition.agda
Normal file
@@ -0,0 +1,383 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_; Setω)
|
||||
open import Setoids.Setoids
|
||||
open import Groups.Definition
|
||||
open import LogicalFormulae
|
||||
open import Orders.WellFounded.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.FreeProduct.Addition {i : _} {I : Set i} (decidableIndex : (x y : I) → ((x ≡ y) || ((x ≡ y) → False))) {a b : _} {A : I → Set a} {S : (i : I) → Setoid {a} {b} (A i)} {_+_ : (i : I) → (A i) → (A i) → A i} (decidableGroups : (i : I) → (x y : A i) → ((Setoid._∼_ (S i) x y)) || ((Setoid._∼_ (S i) x y) → False)) (G : (i : I) → Group (S i) (_+_ i)) where
|
||||
|
||||
open import Groups.FreeProduct.Definition decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
||||
|
||||
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 x) | inr i!=j = nonempty i (prependLetter i g nonzero x i!=j)
|
||||
|
||||
plus' : {j : _} → (ReducedSequenceBeginningWith j) → ReducedSequence → ReducedSequence
|
||||
plus' (ofEmpty i g nonZero) s = prepend i g nonZero s
|
||||
plus' (prependLetter i g nonZero x pr) s = prepend i g nonZero (plus' x s)
|
||||
|
||||
_+RP_ : ReducedSequence → ReducedSequence → ReducedSequence
|
||||
empty +RP b = b
|
||||
nonempty i x +RP b = plus' x b
|
||||
|
||||
abstract
|
||||
prependWD : {i : I} → (g1 g2 : A i) → .(nz1 : _) → .(nz2 : _) → (y : ReducedSequence) (eq : Setoid._∼_ (S i) g1 g2) → prepend i g1 nz1 y =RP prepend i g2 nz2 y
|
||||
prependWD {i} g1 g2 nz1 nz2 empty g1=g2 with decidableIndex i i
|
||||
prependWD {i} g1 g2 nz1 nz2 empty g1=g2 | inl refl = g1=g2
|
||||
prependWD {i} g1 g2 nz1 nz2 empty g1=g2 | inr x = exFalso (x refl)
|
||||
prependWD {i} g1 g2 nz1 nz2 (nonempty j x) g1=g2 with decidableIndex i j
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl with decidableGroups j ((j + g1) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inl x with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inl x | inl x₁ = record {}
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inl x | inr bad = exFalso (bad (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.symmetric (Setoid.eq (S j)) g1=g2) (Equivalence.reflexive (Setoid.eq (S j)))) x))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inr x with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inr bad | inl p = exFalso (bad (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) g1=g2 (Equivalence.reflexive (Setoid.eq (S j)))) p))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inr x | inr x₁ with decidableIndex j j
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inr x | inr x₁ | inl refl = Group.+WellDefined (G j) g1=g2 (Equivalence.reflexive (Setoid.eq (S j)))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (ofEmpty .j g nonZero)) g1=g2 | inl refl | inr x | inr x₁ | inr bad = exFalso (bad refl)
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl with decidableGroups j ((j + g1) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inl x with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inl x | inl x₂ = =RP'reflex w
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inl x | inr bad = exFalso (bad (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.symmetric (Setoid.eq (S j)) g1=g2) (Equivalence.reflexive (Setoid.eq (S j)))) x))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inr neq with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inr neq | inl x = exFalso (neq (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) g1=g2 (Equivalence.reflexive (Setoid.eq (S j)))) x))
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inr neq | inr x with decidableIndex j j
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inr neq | inr x | inl refl = Group.+WellDefined (G j) g1=g2 (Equivalence.reflexive (Setoid.eq (S j))) ,, =RP'reflex w
|
||||
prependWD {.j} g1 g2 nz1 nz2 (nonempty j (prependLetter .j g nonZero w x₁)) g1=g2 | inl refl | inr neq | inr x | inr bad = exFalso (bad refl)
|
||||
prependWD {i} g1 g2 nz1 nz2 (nonempty j x) g1=g2 | inr i!=j with decidableIndex i i
|
||||
prependWD {i} g1 g2 nz1 nz2 (nonempty j x) g1=g2 | inr i!=j | inl refl = g1=g2 ,, =RP'reflex x
|
||||
prependWD {i} g1 g2 nz1 nz2 (nonempty j x) g1=g2 | inr i!=j | inr bad = exFalso (bad refl)
|
||||
|
||||
abstract
|
||||
prependWD' : {i : I} → (g : A i) → .(nz : _) → (y1 y2 : ReducedSequence) (eq : y1 =RP y2) → prepend i g nz y1 =RP prepend i g nz y2
|
||||
prependWD' {i} g1 nz empty empty eq with decidableIndex i i
|
||||
prependWD' {i} g1 nz empty empty eq | inl refl = Equivalence.reflexive (Setoid.eq (S i))
|
||||
prependWD' {i} g1 nz empty empty eq | inr x = exFalso (x refl)
|
||||
prependWD' {i} g1 nz (nonempty j x) (nonempty k x₁) eq with decidableIndex i j
|
||||
prependWD' {.j} g1 nz (nonempty j w1) (nonempty k w2) eq | inl refl with decidableIndex j k
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl with decidableGroups j ((j + g1) g) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inl eq1 with decidableGroups j ((j + g1) h) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inl eq1 | inl x = record {}
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inl eq1 | inr x with decidableIndex j j
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inl eq1 | inr x | inl refl = exFalso (x (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (Equivalence.symmetric (Setoid.eq (S j)) eq)) eq1))
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inr neq1 with decidableGroups j ((j + g1) h) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inr neq1 | inl x with decidableIndex j j
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inr neq1 | inl x | inl refl = exFalso (neq1 (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) eq) x))
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inr neq1 | inr x with decidableIndex j j
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty .j (ofEmpty .j h nonZero₁)) eq | inl refl | inl refl | inr neq1 | inr x | inl refl = Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) eq
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl with decidableIndex j j
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl with decidableGroups j ((j + g1) g) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inl eq1 with decidableGroups j ((j + g1) h) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inl eq1 | inl eq2 = _&&_.snd eq
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inl eq1 | inr neq2 = exFalso (neq2 (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (Equivalence.symmetric (Setoid.eq (S j)) (_&&_.fst eq))) eq1))
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inr neq1 with decidableGroups j ((j + g1) h) (Group.0G (G j))
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inr neq1 | inl eq2 = exFalso (neq1 (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (_&&_.fst eq)) eq2))
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inr neq1 | inr _ with decidableIndex j j
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inr neq1 | inr _ | inl refl = Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (_&&_.fst eq) ,, _&&_.snd eq
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty .j (prependLetter .j h nonZero₁ w2 x₁)) eq | inl refl | inl refl | inl refl | inr neq1 | inr _ | inr bad = exFalso (bad refl)
|
||||
prependWD' {.j} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty k (ofEmpty .k g₁ nonZero₁)) eq | inl refl | inr j!=k with decidableIndex j k
|
||||
prependWD' {j} g1 nz (nonempty j (ofEmpty j g nonZero)) (nonempty k (ofEmpty .k g₁ nonZero₁)) eq | inl refl | inr j!=k | inl j=k = exFalso (j!=k j=k)
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty k (prependLetter .k g₁ nonZero₁ w2 x₁)) eq | inl refl | inr j!=k with decidableIndex j k
|
||||
prependWD' {.j} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty k (prependLetter .k g₁ nonZero₁ w2 x₁)) eq | inl refl | inr j!=k | inl j=k = exFalso (j!=k j=k)
|
||||
prependWD' {i} g1 nz (nonempty j w1) (nonempty k w2) eq | inr i!=j with decidableIndex i k
|
||||
prependWD' {.k} g1 nz (nonempty j (ofEmpty .j g nonZero)) (nonempty k (ofEmpty .k g₁ nonZero₁)) eq | inr i!=j | inl refl with decidableIndex j k
|
||||
prependWD' {.k} g1 nz (nonempty .k (ofEmpty .k g nonZero)) (nonempty k (ofEmpty .k g₁ nonZero₁)) eq | inr i!=j | inl refl | inl refl = exFalso (i!=j refl)
|
||||
prependWD' {.k} g1 nz (nonempty j (prependLetter .j g nonZero w1 x)) (nonempty k (prependLetter .k g₁ nonZero₁ w2 x₁)) eq | inr i!=j | inl refl with decidableIndex j k
|
||||
prependWD' {.k} g1 nz (nonempty .k (prependLetter .k g nonZero w1 x)) (nonempty k (prependLetter .k g₁ nonZero₁ w2 x₁)) eq | inr i!=j | inl refl | inl refl = exFalso (i!=j refl)
|
||||
prependWD' {i} g1 nz (nonempty j w1) (nonempty k w2) eq | inr i!=j | inr i!=k with decidableIndex i i
|
||||
prependWD' {i} g1 nz (nonempty j w1) (nonempty k w2) eq | inr i!=j | inr i!=k | inl refl = Equivalence.reflexive (Setoid.eq (S i)) ,, eq
|
||||
prependWD' {i} g1 nz (nonempty j w1) (nonempty k w2) eq | inr i!=j | inr i!=k | inr x = exFalso (x refl)
|
||||
|
||||
abstract
|
||||
plus'WDlemm : {i : I} (x1 x2 : ReducedSequenceBeginningWith i) (y : ReducedSequence) → (=RP' x1 x2) → plus' x1 y =RP plus' x2 y
|
||||
plus'WDlemm (ofEmpty i g nonZero) (ofEmpty .i g₁ nonZero₁) y x1=x2 with decidableIndex i i
|
||||
plus'WDlemm (ofEmpty i g nonZero) (ofEmpty .i h nonZero2) y x1=x2 | inl refl = prependWD g h nonZero nonZero2 y x1=x2
|
||||
plus'WDlemm (prependLetter i g nonZero {j} x1 x) (prependLetter .i h nonZero2 {j2} x2 pr) y x1=x2 with decidableIndex i i
|
||||
... | inl refl with decidableIndex j j2
|
||||
plus'WDlemm (prependLetter i g nonZero {j} x1 x) (prependLetter .i h nonZero2 {.j} x2 pr) y x1=x2 | inl refl | inl refl = Equivalence.transitive (Setoid.eq freeProductSetoid) {prepend i g nonZero (plus' x1 y)} {prepend i h nonZero2 (plus' x1 y)} {plus' (prependLetter i h nonZero2 x2 pr) y} (prependWD g h nonZero nonZero2 (plus' x1 y) (_&&_.fst x1=x2)) (prependWD' h nonZero2 (plus' x1 y) (plus' x2 y) (plus'WDlemm x1 x2 y (_&&_.snd x1=x2)))
|
||||
... | inr j!=j2 = exFalso (notEqualIfStartDifferent j!=j2 x1 x2 (_&&_.snd x1=x2))
|
||||
plus'WDlemm _ _ _ _ | inr bad = exFalso (bad refl)
|
||||
|
||||
abstract
|
||||
plus'WD : {i j : I} (x1 : ReducedSequenceBeginningWith i) (x2 : ReducedSequenceBeginningWith j) (y : ReducedSequence) → (=RP' x1 x2) → plus' x1 y =RP plus' x2 y
|
||||
plus'WD {i} {j} x1 x2 y x1=x2 with decidableIndex i j
|
||||
plus'WD {i} {.i} x1 x2 y x1=x2 | inl refl = plus'WDlemm x1 x2 y x1=x2
|
||||
plus'WD {i} {j} x1 x2 y x1=x2 | inr x = exFalso (notEqualIfStartDifferent x x1 x2 x1=x2)
|
||||
|
||||
private
|
||||
abstract
|
||||
prependLetterWD : {i : I} {h1 h2 : A i} (h1=h2 : Setoid._∼_ (S i) h1 h2) → .(nonZero1 : _) .(nonZero2 : _) {j1 j2 : I} (w1 : ReducedSequenceBeginningWith j1) (w2 : ReducedSequenceBeginningWith j2) → (eq : =RP' w1 w2) → (x1 : i ≡ _ → False) (x2 : i ≡ _ → False) → nonempty i (prependLetter i h1 nonZero1 w1 x1) =RP nonempty i (prependLetter i h2 nonZero2 w2 x2)
|
||||
prependLetterWD {i} h1=h2 nonZero1 nonZero2 {j1} {j2} w1 w2 eq x1 x2 with decidableIndex i i
|
||||
prependLetterWD {i} h1=h2 nonZero1 nonZero2 {j1} {j2} w1 w2 eq x1 x2 | inl refl = h1=h2 ,, eq
|
||||
prependLetterWD {i} h1=h2 nonZero1 nonZero2 {j1} {j2} w1 w2 eq x1 x2 | inr x = exFalso (x refl)
|
||||
|
||||
abstract
|
||||
plus'WD' : {i : I} (x : ReducedSequenceBeginningWith i) (y1 y2 : ReducedSequence) → (y1 =RP y2) → plus' x y1 =RP plus' x y2
|
||||
plus'WD' x empty empty eq = Equivalence.reflexive (Setoid.eq freeProductSetoid) {plus' x empty}
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 w1) (nonempty i2 w2) eq with decidableIndex j i1
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j w1) (nonempty i2 w2) eq | inl refl with decidableIndex j i2
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j g₁ nonZero₁)) (nonempty .j (ofEmpty .j g₂ nonZero₂)) eq | inl refl | inl refl with decidableIndex j j
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl with decidableGroups j ((j + g) h1) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inl x with decidableGroups j ((j + g) h2) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inl x | inl x₁ = record {}
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inl x | inr bad = exFalso (bad (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (Equivalence.symmetric (Setoid.eq (S j)) eq)) x))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inr x with decidableGroups j ((j + g) h2) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inr x | inl bad = exFalso (x (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) eq) bad))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inr x | inr x₁ with decidableIndex j j
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inr x | inr x₁ | inl refl = Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) eq
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j h1 nonZero₁)) (nonempty .j (ofEmpty .j h2 nonZero₂)) eq | inl refl | inl refl | inl refl | inr x | inr x₁ | inr bad = exFalso (bad refl)
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl with decidableGroups j ((j + g) h1) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inl eq1 with decidableGroups j ((j + g) h2) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inl eq1 | inl x₂ with decidableIndex j j
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inl eq1 | inl x₂ | inl refl = _&&_.snd eq
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inl eq1 | inr neq2 with decidableIndex j j
|
||||
... | inl refl = exFalso (neq2 (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (Equivalence.symmetric (Setoid.eq (S j)) (_&&_.fst eq))) eq1))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inr neq1 with decidableGroups j ((j + g) h2) (Group.0G (G j))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inr neq1 | inl eq2 with decidableIndex j j
|
||||
... | inl refl = exFalso (neq1 (Equivalence.transitive (Setoid.eq (S j)) (Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (_&&_.fst eq)) eq2))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j h1 nonZero₁ w1 x)) (nonempty .j (prependLetter .j h2 nonZero₂ w2 x₁)) eq | inl refl | inl refl | inr neq1 | inr x₂ with decidableIndex j j
|
||||
... | inl refl = Group.+WellDefined (G j) (Equivalence.reflexive (Setoid.eq (S j))) (_&&_.fst eq) ,, _&&_.snd eq
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j g₁ nonZero₁)) (nonempty i2 (ofEmpty .i2 g₂ nonZero₂)) eq | inl refl | inr x with decidableIndex j i2
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (ofEmpty .j g₁ nonZero₁)) (nonempty .j (ofEmpty .j g₂ nonZero₂)) eq | inl refl | inr x | inl refl = exFalso (x refl)
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty .j (prependLetter .j g₁ nonZero₁ w1 x₁)) (nonempty i2 (prependLetter .i2 g₂ nonZero₂ w2 x₂)) eq | inl refl | inr x with decidableIndex j i2
|
||||
plus'WD' {.i2} (ofEmpty .i2 g nonZero) (nonempty .i2 (prependLetter .i2 g₁ nonZero₁ w1 x₁)) (nonempty i2 (prependLetter .i2 g₂ nonZero₂ w2 x₂)) eq | inl refl | inr x | inl refl = exFalso (x refl)
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 w1) (nonempty i2 w2) eq | inr j!=i1 with decidableIndex j i2
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 (ofEmpty .i1 g₁ nonZero₁)) (nonempty .j (ofEmpty .j g₂ nonZero₂)) eq | inr j!=i1 | inl refl with decidableIndex i1 j
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 (ofEmpty .i1 g₁ nonZero₁)) (nonempty .j (ofEmpty .j g₂ nonZero₂)) eq | inr j!=i1 | inl refl | inl x = exFalso (j!=i1 (equalityCommutative x))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 (prependLetter .i1 g₁ nonZero₁ w1 x)) (nonempty .j (prependLetter .j g₂ nonZero₂ w2 x₁)) eq | inr j!=i1 | inl refl with decidableIndex i1 j
|
||||
... | inl bad = exFalso (j!=i1 (equalityCommutative bad))
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 w1) (nonempty i2 w2) eq | inr j!=i1 | inr j!=i2 with decidableIndex j j
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 w1) (nonempty i2 w2) eq | inr j!=i1 | inr j!=i2 | inl refl = Equivalence.reflexive (Setoid.eq (S j)) ,, eq
|
||||
plus'WD' {j} (ofEmpty j g nonZero) (nonempty i1 w1) (nonempty i2 w2) eq | inr j!=i1 | inr j!=i2 | inr x = exFalso (x refl)
|
||||
plus'WD' {j} (prependLetter j g nonZero l x) (nonempty i (ofEmpty .i h1 nonZero1)) (nonempty i2 (ofEmpty .i2 h2 nonZero2)) eq with decidableIndex i i2
|
||||
... | inl refl = prependWD' g nonZero (plus' l (nonempty i (ofEmpty i h1 nonZero1))) (plus' l (nonempty i2 (ofEmpty i2 h2 nonZero2))) (plus'WD' l (nonempty i2 (ofEmpty i2 h1 nonZero1)) (nonempty i2 (ofEmpty i2 h2 nonZero2)) t)
|
||||
where
|
||||
t : =RP' (ofEmpty i2 h1 nonZero1) (ofEmpty i2 h2 nonZero2)
|
||||
t with decidableIndex i2 i2
|
||||
t | inl refl = eq
|
||||
t | inr x = exFalso (x refl)
|
||||
... | inr i!=i2 = exFalso' eq
|
||||
plus'WD' {j} (prependLetter j g nonZero l x) (nonempty i (prependLetter .i g₁ nonZero₁ w1 x₁)) (nonempty i2 (prependLetter .i2 g₂ nonZero₂ w2 x₂)) eq with decidableIndex i i2
|
||||
plus'WD' {j} (prependLetter j g nonZero l x) (nonempty i (prependLetter .i h1 nonZero1 w1 x1)) (nonempty .i (prependLetter .i h2 nonZero2 w2 x2)) eq | inl refl = prependWD' g nonZero (plus' l (nonempty i (prependLetter i h1 nonZero1 w1 x1))) (plus' l (nonempty i (prependLetter i h2 nonZero2 w2 x2))) (plus'WD' l (nonempty i (prependLetter i h1 nonZero1 w1 x1)) (nonempty i (prependLetter i h2 nonZero2 w2 x2)) (prependLetterWD (_&&_.fst eq) nonZero1 nonZero2 w1 w2 (_&&_.snd eq) x1 x2))
|
||||
|
||||
abstract
|
||||
plusEmptyRight : {i : I} (x : ReducedSequenceBeginningWith i) → _=RP_ (plus' x empty) (nonempty i x)
|
||||
plusEmptyRight (ofEmpty i g nonZero) with decidableIndex i i
|
||||
plusEmptyRight (ofEmpty i g nonZero) | inl refl = Equivalence.reflexive (Setoid.eq (S i))
|
||||
plusEmptyRight (ofEmpty i g nonZero) | inr x = exFalso (x refl)
|
||||
plusEmptyRight (prependLetter i g nonZero {j} x i!=j) with plusEmptyRight x
|
||||
... | b with plus' x empty
|
||||
plusEmptyRight (prependLetter i g nonZero {j} x i!=j) | b | nonempty k w with decidableIndex i k
|
||||
plusEmptyRight (prependLetter .k g nonZero {j} x i!=j) | b | nonempty k (ofEmpty .k h nonZero₁) | inl refl with decidableGroups k ((k + g) h) (Group.0G (G k))
|
||||
plusEmptyRight (prependLetter .k g nonZero {.i} (ofEmpty i g₁ nonZero₂) i!=j) | b | nonempty k (ofEmpty .k h nonZero₁) | inl refl | inl g+k=hinv with decidableIndex k i
|
||||
plusEmptyRight (prependLetter .i g nonZero {.i} (ofEmpty i g₁ nonZero₂) i!=j) | b | nonempty .i (ofEmpty .i h nonZero₁) | inl refl | inl g+k=hinv | inl refl = exFalso (i!=j refl)
|
||||
plusEmptyRight (prependLetter .k g nonZero {.i} (ofEmpty i g₁ nonZero₂) i!=j) | b | nonempty k (ofEmpty .k h nonZero₁) | inl refl | inr x₁ with decidableIndex k i
|
||||
plusEmptyRight (prependLetter .i g nonZero {.i} (ofEmpty i g₁ nonZero₂) i!=j) | b | nonempty .i (ofEmpty .i h nonZero₁) | inl refl | inr x₁ | inl refl = exFalso (i!=j refl)
|
||||
plusEmptyRight (prependLetter .k g nonZero {j} w1 i!=j) | b | nonempty k (prependLetter .k h nonZero₁ w x₁) | inl refl with decidableGroups k ((k + g) h) (Group.0G (G k))
|
||||
plusEmptyRight (prependLetter .k g nonZero {.i} (prependLetter i g₁ nonZero₂ w1 x) i!=j) | b | nonempty k (prependLetter .k h nonZero₁ w x₁) | inl refl | inl p with decidableIndex k i
|
||||
plusEmptyRight (prependLetter .i g nonZero {.i} (prependLetter i g₁ nonZero₂ w1 x) i!=j) | b | nonempty .i (prependLetter .i h nonZero₁ w x₁) | inl refl | inl p | inl refl = exFalso (i!=j refl)
|
||||
plusEmptyRight (prependLetter .k g nonZero {j} w1 i!=j) | b | nonempty k (prependLetter .k h nonZero₁ w x₁) | inl refl | inr x₂ with decidableIndex k k
|
||||
plusEmptyRight (prependLetter .k g nonZero {.i} (prependLetter i g₁ nonZero₂ w1 x) i!=j) | b | nonempty k (prependLetter .k h nonZero₁ w x₁) | inl refl | inr x₂ | inl refl with decidableIndex k i
|
||||
plusEmptyRight (prependLetter .i g nonZero {.i} (prependLetter i g₁ nonZero₂ w1 x) i!=j) | b | nonempty .i (prependLetter .i h nonZero₁ w x₁) | inl refl | inr x₂ | inl refl | inl refl = exFalso (i!=j refl)
|
||||
plusEmptyRight (prependLetter .k g nonZero {j} w1 i!=j) | b | nonempty k (prependLetter .k h nonZero₁ w x₁) | inl refl | inr x₂ | inr x = exFalso (x refl)
|
||||
plusEmptyRight (prependLetter i g nonZero {j} w1 i!=j) | b | nonempty k w | inr i!=k with decidableIndex i i
|
||||
plusEmptyRight (prependLetter i g nonZero {j} w1 i!=j) | b | nonempty k w | inr i!=k | inl refl = Equivalence.reflexive (Setoid.eq (S i)) ,, b
|
||||
plusEmptyRight (prependLetter i g nonZero {j} w1 i!=j) | b | nonempty k w | inr i!=k | inr x = exFalso (x refl)
|
||||
|
||||
abstract
|
||||
plusWD : (m n o p : ReducedSequence) → (m =RP o) → (n =RP p) → (m +RP n) =RP (o +RP p)
|
||||
plusWD empty empty empty empty m=o n=p = record {}
|
||||
plusWD empty (nonempty i x) empty (nonempty i₁ x₁) m=o n=p = n=p
|
||||
plusWD (nonempty i x) empty (nonempty j y) empty m=o record {} = Equivalence.transitive (Setoid.eq freeProductSetoid) {plus' x empty} {nonempty i x} {plus' y empty} (plusEmptyRight x) (Equivalence.transitive (Setoid.eq freeProductSetoid) {nonempty i x} {nonempty j y} {plus' y empty} m=o (Equivalence.symmetric (Setoid.eq freeProductSetoid) {plus' y empty} {nonempty j y} (plusEmptyRight y)))
|
||||
plusWD (nonempty i1 x1) (nonempty i2 x2) (nonempty i3 x3) (nonempty i4 x4) m=o n=p = (Equivalence.transitive (Setoid.eq freeProductSetoid)) {plus' x1 (nonempty i2 x2)} {plus' x3 (nonempty i2 x2)} {plus' x3 (nonempty i4 x4)} (plus'WD x1 x3 (nonempty i2 x2) m=o) (plus'WD' x3 (nonempty i2 x2) (nonempty i4 x4) n=p)
|
||||
|
||||
abstract
|
||||
prependFrom : {i : I} (g1 g2 : A i) (pr : (Setoid._∼_ (S i) (_+_ i g1 g2) (Group.0G (G i))) → False) (c : ReducedSequence) .(p2 : _) .(p3 : _) → prepend i (_+_ i g1 g2) pr c =RP prepend i g1 p2 (prepend i g2 p3 c)
|
||||
prependFrom {i} g1 g2 pr empty _ _ with decidableIndex i i
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inl refl with decidableGroups i ((i + g1) g2) (Group.0G (G i))
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inl refl | inl bad = exFalso (pr bad)
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inl refl | inr _ with decidableIndex i i
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inl refl | inr _ | inl refl = Equivalence.reflexive (Setoid.eq (S i))
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inl refl | inr _ | inr bad = exFalso (bad refl)
|
||||
prependFrom {i} g1 g2 pr empty _ _ | inr bad = exFalso (bad refl)
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ with decidableIndex i j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl with decidableGroups j ((j + (j + g1) g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) g1Nonzero _ | inl refl | inl eq3 | inl eq2 = exFalso (g1Nonzero t)
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Group (G j)
|
||||
open Equivalence eq
|
||||
t : g1 ∼ 0G
|
||||
t = transitive (symmetric identRight) (transitive (+WellDefined reflexive (symmetric eq2)) (transitive +Associative eq3))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq | inr neq with decidableIndex j j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq | inr neq | inl refl with decidableGroups j ((j + g1) ((j + g2) g)) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq | inr neq | inl refl | inl eq2 = record {}
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq | inr neq | inl refl | inr neq2 = exFalso (neq2 (Equivalence.transitive (Setoid.eq (S j)) (Group.+Associative (G j)) eq))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inl eq | inr neq | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inl eq2 with decidableIndex j j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inl eq2 | inl refl = transitive (symmetric +Associative) (transitive (+WellDefined reflexive eq2) identRight)
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Group (G j)
|
||||
open Equivalence eq
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inl eq2 | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inr neq2 with decidableIndex j j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inr neq2 | inl refl with decidableGroups j ((j + g1) ((j + g2) g)) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inr neq2 | inl refl | inl eq2 = exFalso (neq (Equivalence.transitive (Setoid.eq (S j)) (Equivalence.symmetric (Setoid.eq (S j)) (Group.+Associative (G j))) eq2))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inr neq2 | inl refl | inr neq3 with decidableIndex j j
|
||||
... | inl refl = Equivalence.symmetric (Setoid.eq (S j)) (Group.+Associative (G j))
|
||||
... | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (ofEmpty j g nonZero)) _ _ | inl refl | inr neq | inr neq2 | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl with decidableGroups j ((j + (j + g1) g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero {k} x x₁)) p2 _ | inl refl | inl eq1 | inl eq2 = exFalso (p2 (transitive (transitive (symmetric identRight) (transitive (+WellDefined reflexive (symmetric eq2)) +Associative)) eq1))
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Group (G j)
|
||||
open Equivalence eq
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 | inr neq with decidableIndex j j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 | inr neq | inl refl with decidableGroups j ((j + g1) ((j + g2) g)) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 | inr neq | inl refl | inl eq2 = =RP'reflex x
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 | inr neq | inl refl | inr neq2 = exFalso (neq2 (Equivalence.transitive (Setoid.eq (S j)) (Group.+Associative (G j)) eq1))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inl eq1 | inr neq | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 with decidableGroups j ((j + g2) g) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl eq1 with decidableIndex j k
|
||||
prependFrom {j} g1 g2 pr (nonempty _ (prependLetter _ g nonZero {_} (ofEmpty _ h nonZero₁) bad)) _ _ | inl refl | inr neq1 | inl eq1 | inl refl = exFalso (bad refl)
|
||||
prependFrom {_} g1 g2 pr (nonempty _ (prependLetter _ g nonZero {_} (prependLetter _ g₁ nonZero₁ x x₂) bad)) _ _ | inl refl | inr neq1 | inl eq1 | inl refl = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl eq1 | inr j!=k with decidableIndex j j
|
||||
... | inl refl = transitive (symmetric +Associative) (transitive (+WellDefined reflexive eq1) identRight) ,, =RP'reflex x
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Group (G j)
|
||||
open Equivalence eq
|
||||
... | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 | inr neq2 with decidableIndex j j
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 | inr neq2 | inl refl with decidableGroups j ((j + g1) ((j + g2) g)) (Group.0G (G j))
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 | inr neq2 | inl refl | inl eq1 = exFalso (neq1 (transitive (symmetric +Associative) eq1))
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Group (G j)
|
||||
open Equivalence eq
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 | inr neq2 | inl refl | inr neq3 with decidableIndex j j
|
||||
... | inl refl = Equivalence.symmetric (Setoid.eq (S j)) (Group.+Associative (G j)) ,, =RP'reflex x
|
||||
... | inr bad = exFalso (bad refl)
|
||||
prependFrom {.j} g1 g2 pr (nonempty _ (prependLetter j g nonZero x x₁)) _ _ | inl refl | inr neq1 | inr neq2 | inr bad = exFalso (bad refl)
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j with decidableIndex i i
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inl refl with decidableGroups i ((i + g1) g2) (Group.0G (G i))
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inl refl | inl eq1 = exFalso (pr eq1)
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inl refl | inr neq1 with decidableIndex i i
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inl refl | inr neq1 | inl refl = Equivalence.reflexive (Setoid.eq (S i)) ,, =RP'reflex x
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inl refl | inr neq1 | inr bad = exFalso (bad refl)
|
||||
prependFrom {i} g1 g2 pr (nonempty j x) _ _ | inr i!=j | inr bad = exFalso (bad refl)
|
||||
|
||||
prependFrom' : {i : I} (g h : A i) (pr : (Setoid._∼_ (S i) (_+_ i g h) (Group.0G (G i)))) (c : ReducedSequence) .(p2 : _) .(p3 : _) → prepend i g p2 (prepend i h p3 c) =RP c
|
||||
prependFrom' {i} g h pr empty _ _ with decidableIndex i i
|
||||
prependFrom' {i} g h pr empty _ _ | inl refl with decidableGroups i ((i + g) h) (Group.0G (G i))
|
||||
... | inl _ = record {}
|
||||
... | inr bad = exFalso (bad pr)
|
||||
prependFrom' {i} g h pr empty _ _ | inr bad = exFalso (bad refl)
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ with decidableIndex i j
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl with decidableGroups j ((j + h) g1) (Group.0G (G j))
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inl eq1 with decidableIndex j j
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inl eq1 | inl refl = transitive t (symmetric u)
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
t : Setoid._∼_ (S j) g (Group.inverse (G j) h)
|
||||
t = rightInversesAreUnique (G j) pr
|
||||
u : Setoid._∼_ (S j) g1 (Group.inverse (G j) h)
|
||||
u = leftInversesAreUnique (G j) eq1
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inl eq1 | inr bad = exFalso (bad refl)
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inr neq with decidableIndex j j
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) p2 _ | inl refl | inr neq | inl refl with decidableGroups j ((j + g) ((j + h) g1)) (Group.0G (G j))
|
||||
... | inl eq1 = exFalso (nonZero t)
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
open Group (G j)
|
||||
t : g1 ∼ 0G
|
||||
t = transitive (transitive (symmetric identLeft) (transitive (+WellDefined (symmetric pr) reflexive) (symmetric +Associative))) eq1
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inr neq | inl refl | inr neq2 with decidableIndex j j
|
||||
... | inl refl = transitive (transitive +Associative (+WellDefined pr reflexive)) identLeft
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
open Group (G j)
|
||||
... | inr bad = exFalso (bad refl)
|
||||
prependFrom' {.j} g h pr (nonempty j (ofEmpty .j g1 nonZero)) _ _ | inl refl | inr neq | inr bad = exFalso (bad refl)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero x x₁)) _ _ | inl refl with decidableGroups j ((j + h) g1) (Group.0G (G j))
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x bad)) _ _ | inl refl | inl eq1 with decidableIndex j k
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x bad)) _ _ | inl refl | inl eq1 | inl j=k = exFalso (bad j=k)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x bad)) _ _ | inl refl | inl eq1 | inr j!=k with decidableIndex j j
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x bad)) _ _ | inl refl | inl eq1 | inr j!=k | inl refl = transitive {g} {inverse h} (rightInversesAreUnique (G j) pr) (symmetric (leftInversesAreUnique (G j) eq1)) ,, =RP'reflex x
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
open Group (G j)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x _)) _ _ | inl refl | inl eq1 | inr j!=k | inr bad = exFalso (bad refl)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 with decidableIndex j j
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl refl with decidableGroups j ((j + g) ((j + h) g1)) (Group.0G (G j))
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl refl | inl eq1 = exFalso (nonZero (transitive (transitive (symmetric identLeft) (transitive (+WellDefined (symmetric pr) reflexive) (symmetric +Associative))) eq1))
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
open Group (G j)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl refl | inr neq2 with decidableIndex j j
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl refl | inr neq2 | inl refl = transitive +Associative (transitive (+WellDefined pr reflexive) identLeft) ,, =RP'reflex x
|
||||
where
|
||||
open Setoid (S j)
|
||||
open Equivalence eq
|
||||
open Group (G j)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inl refl | inr neq2 | inr bad = exFalso (bad refl)
|
||||
prependFrom' {.j} g h pr (nonempty j (prependLetter .j g1 nonZero {k} x x₁)) _ _ | inl refl | inr neq1 | inr bad = exFalso (bad refl)
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ | inr i!=j with decidableIndex i i
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ | inr i!=j | inl refl with decidableGroups i ((i + g) h) (Group.0G (G i))
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ | inr i!=j | inl refl | inl eq1 = =RP'reflex x
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ | inr i!=j | inl refl | inr neq1 = exFalso (neq1 pr)
|
||||
prependFrom' {i} g h pr (nonempty j x) _ _ | inr i!=j | inr bad = exFalso (bad refl)
|
||||
|
||||
prependAssocLemma' : {i : I} {g : A i} .(nz : Setoid._∼_ (S i) g (Group.0G (G i)) → False) (b c : ReducedSequence) → (prepend i g nz b +RP c) =RP prepend i g nz (b +RP c)
|
||||
prependAssocLemma' {i} {g} nz empty c = Equivalence.reflexive (Setoid.eq freeProductSetoid) {prepend i g _ c}
|
||||
prependAssocLemma' {i} nz (nonempty k x) c with decidableIndex i k
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (ofEmpty .k g nonZero)) c | inl refl with decidableGroups k ((k + g1) g) (Group.0G (G k))
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (ofEmpty .k g nonZero)) c | inl refl | inl eq1 = Equivalence.symmetric (Setoid.eq freeProductSetoid) {prepend k g1 _ (prepend k g _ c)} {c} (prependFrom' g1 g eq1 c nz nonZero)
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (ofEmpty .k g nonZero)) c | inl refl | inr neq1 = (prependFrom g1 g neq1 c nz nonZero)
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (prependLetter .k g nonZero x x₁)) c | inl refl with decidableGroups k ((k + g1) g) (Group.0G (G k))
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (prependLetter .k g nonZero x x₁)) c | inl refl | inl eq1 = Equivalence.symmetric (Setoid.eq freeProductSetoid) {prepend k g1 _ (prepend k g _ (plus' x c))} {plus' x c} (prependFrom' g1 g eq1 (plus' x c) nz nonZero)
|
||||
prependAssocLemma' {.k} {g1} nz (nonempty k (prependLetter .k g nonZero x x₁)) c | inl refl | inr neq1 = prependFrom g1 g neq1 (plus' x c) nz nonZero
|
||||
prependAssocLemma' {i} {g} nz (nonempty k x) c | inr i!=k = Equivalence.reflexive (Setoid.eq freeProductSetoid) {prepend i g _ (plus' x c)}
|
||||
|
||||
abstract
|
||||
plusAssocLemma : {i : I} (x : ReducedSequenceBeginningWith i) (b c : ReducedSequence) → (plus' x b +RP c) =RP plus' x (b +RP c)
|
||||
plusAssocLemma (ofEmpty i g nonZero) empty c = Equivalence.reflexive (Setoid.eq freeProductSetoid) {prepend i g _ c}
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty j b) c with decidableIndex i j
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (ofEmpty .i g₁ nonZero₁)) c | inl refl with decidableIndex i i
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (ofEmpty .i h nonZero₁)) c | inl refl | inl refl with decidableGroups i ((i + g) h) (Group.0G (G i))
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (ofEmpty .i h nonZero₁)) c | inl refl | inl refl | inl t = Equivalence.symmetric (Setoid.eq freeProductSetoid) {prepend i g _ (prepend i h _ c)} {c} (prependFrom' g h t c _ _)
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (ofEmpty .i h nonZero₁)) c | inl refl | inl refl | inr t = prependFrom g h t c _ _
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (ofEmpty .i g₁ nonZero₁)) c | inl refl | inr bad = exFalso (bad refl)
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (prependLetter .i h nonZero₁ b x)) c | inl refl with decidableGroups i ((i + g) h) (Group.0G (G i))
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (prependLetter .i h nonZero₁ b x)) c | inl refl | inl eq1 = Equivalence.symmetric (Setoid.eq freeProductSetoid) {prepend i g _ (prepend i h _ (plus' b c))} {plus' b c} (prependFrom' g h eq1 (plus' b c) _ _)
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty .i (prependLetter .i h nonZero₁ b x)) c | inl refl | inr neq1 = prependFrom g h neq1 (plus' b c) nonZero _
|
||||
plusAssocLemma (ofEmpty i g nonZero) (nonempty j b) c | inr i!=j = Equivalence.reflexive (Setoid.eq freeProductSetoid) {prepend i g _ (plus' b c)}
|
||||
plusAssocLemma (prependLetter i g nonZero {j} w i!=j) b c = Equivalence.transitive (Setoid.eq freeProductSetoid) {prepend i g _ (plus' w b) +RP c} {prepend i g _ (plus' w b +RP c)} {prepend i g _ (plus' w (b +RP c))} (prependAssocLemma' nonZero (plus' w b) c) (prependWD' g nonZero (plus' w b +RP c) (plus' w (b +RP c)) (plusAssocLemma w b c))
|
||||
|
||||
plusAssoc : (a b c : ReducedSequence) → ((a +RP b) +RP c) =RP (a +RP (b +RP c))
|
||||
plusAssoc empty b c = Equivalence.reflexive (Setoid.eq freeProductSetoid) {b +RP c}
|
||||
plusAssoc (nonempty i x) b c = plusAssocLemma x b c
|
24
Groups/FreeProduct/Definition.agda
Normal file
24
Groups/FreeProduct/Definition.agda
Normal file
@@ -0,0 +1,24 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_; Setω)
|
||||
open import Setoids.Setoids
|
||||
open import Groups.Definition
|
||||
open import LogicalFormulae
|
||||
open import Orders.WellFounded.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.FreeProduct.Definition {i : _} {I : Set i} (decidableIndex : (x y : I) → ((x ≡ y) || ((x ≡ y) → False))) {a b : _} {A : I → Set a} {S : (i : I) → Setoid {a} {b} (A i)} {_+_ : (i : I) → (A i) → (A i) → A i} (decidableGroups : (i : I) → (x y : A i) → ((Setoid._∼_ (S i) x y)) || ((Setoid._∼_ (S i) x y) → False)) (G : (i : I) → Group (S i) (_+_ i)) where
|
||||
|
||||
data ReducedSequenceBeginningWith : I → Set (a ⊔ b ⊔ i) where
|
||||
ofEmpty : (i : I) → (g : A i) → .(nonZero : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → ReducedSequenceBeginningWith i
|
||||
prependLetter : (i : I) → (g : A i) → .(nonZero : Setoid._∼_ (S i) g (Group.0G (G i)) → False) → {j : I} → ReducedSequenceBeginningWith j → ((i ≡ j) → False) → ReducedSequenceBeginningWith i
|
||||
|
||||
data ReducedSequence : Set (a ⊔ b ⊔ i) where
|
||||
empty : ReducedSequence
|
||||
nonempty : (i : I) → ReducedSequenceBeginningWith i → ReducedSequence
|
||||
|
||||
injection : {i : I} (x : A i) .(nonzero : (Setoid._∼_ (S i) x (Group.0G (G i))) → False) → ReducedSequence
|
||||
injection {i} x nonzero = nonempty i (ofEmpty i x nonzero)
|
118
Groups/FreeProduct/Group.agda
Normal file
118
Groups/FreeProduct/Group.agda
Normal file
@@ -0,0 +1,118 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_; Setω)
|
||||
open import Setoids.Setoids
|
||||
open import Groups.Definition
|
||||
open import LogicalFormulae
|
||||
open import Orders.WellFounded.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.FreeProduct.Group {i : _} {I : Set i} (decidableIndex : (x y : I) → ((x ≡ y) || ((x ≡ y) → False))) {a b : _} {A : I → Set a} {S : (i : I) → Setoid {a} {b} (A i)} {_+_ : (i : I) → (A i) → (A i) → A i} (decidableGroups : (i : I) → (x y : A i) → ((Setoid._∼_ (S i) x y)) || ((Setoid._∼_ (S i) x y) → False)) (G : (i : I) → Group (S i) (_+_ i)) where
|
||||
|
||||
open import Groups.FreeProduct.Definition decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Addition decidableIndex decidableGroups G
|
||||
|
||||
private
|
||||
inv' : {i : I} (x : ReducedSequenceBeginningWith i) → ReducedSequence
|
||||
inv' (ofEmpty i g nonZero) = nonempty i (ofEmpty i (Group.inverse (G i) g) (λ pr → nonZero (Equivalence.transitive (Setoid.eq (S i)) (swapInv (G i) pr) (invIdent (G i)))))
|
||||
inv' (prependLetter i g nonZero {j} w i!=j) = (inv' w) +RP injection (Group.inverse (G i) g) (λ pr → nonZero (Equivalence.transitive (Setoid.eq (S i)) (swapInv (G i) pr) (invIdent (G i))))
|
||||
|
||||
inv : (x : ReducedSequence) → ReducedSequence
|
||||
inv empty = empty
|
||||
inv (nonempty i w) = inv' w
|
||||
|
||||
private
|
||||
abstract
|
||||
lemma1 : {i j k : I} (i!=j : (i ≡ j) → False) (g : A i) (h : A j) (w : ReducedSequenceBeginningWith k) (j!=k : (j ≡ k) → False) → .(pr : _) .(pr2 : _) .(pr3 : _) → (prepend i (Group.inverse (G i) g) pr (nonempty i (prependLetter i g pr2 (prependLetter j h pr3 w j!=k) i!=j))) =RP nonempty j (prependLetter j h pr3 w j!=k)
|
||||
lemma1 {i} {j} {k} i!=j g h r j!=k pr pr2 pr3 with decidableIndex i i
|
||||
... | inr x = exFalso (x refl)
|
||||
lemma1 {i} {j} {k} i!=j g h r j!=k pr pr2 pr3 | inl refl with decidableGroups i ((i + Group.inverse (G i) g) g) (Group.0G (G i))
|
||||
... | inr bad = exFalso (bad (Group.invLeft (G i)))
|
||||
lemma1 {i} {j} {k} i!=j g h r j!=k pr pr2 pr3 | inl refl | inl eq1 with decidableIndex j j
|
||||
... | inr bad = exFalso (bad refl)
|
||||
... | inl refl = Equivalence.reflexive (Setoid.eq (S j)) ,, =RP'reflex r
|
||||
|
||||
abstract
|
||||
lemma2 : {j k : I} (j!=k : (j ≡ k) → False) (h : A j) (w : ReducedSequenceBeginningWith k) .(pr : _) .(pr2 : _) → (prepend j (Group.inverse (G j) h) pr (nonempty j (prependLetter j h pr2 w j!=k))) =RP (nonempty k w)
|
||||
lemma2 {j} {k} j!=k h r pr pr2 with decidableIndex j j
|
||||
... | inr bad = exFalso (bad refl)
|
||||
lemma2 {j} {k} j!=k h r pr pr2 | inl refl with decidableGroups j ((j + Group.inverse (G j) h) h) (Group.0G (G j))
|
||||
... | inr bad = exFalso (bad (Group.invLeft (G j)))
|
||||
... | inl x = =RP'reflex r
|
||||
|
||||
abstract
|
||||
unpeel : (k : ReducedSequence) {j : I} (g : A j) .(pr : _) .(pr2 : _) → k =RP empty → (prepend j g pr (k +RP (nonempty j (ofEmpty j (Group.inverse (G j) g) pr2)))) =RP empty
|
||||
unpeel empty {j} g pr pr2 x with decidableIndex j j
|
||||
... | inr bad = exFalso (bad refl)
|
||||
unpeel empty {j} g pr pr2 x | inl refl with decidableGroups j ((j + g) (Group.inverse (G j) g)) (Group.0G (G j))
|
||||
... | inl _ = record {}
|
||||
... | inr bad = exFalso (bad (Group.invRight (G j)))
|
||||
|
||||
abstract
|
||||
invRight' : {i : I} (x : ReducedSequenceBeginningWith i) → ((nonempty i x) +RP inv (nonempty i x)) =RP empty
|
||||
invRight' {i} (ofEmpty _ g nonZero) with decidableIndex i i
|
||||
... | inr x = exFalso (x refl)
|
||||
invRight' {i} (ofEmpty _ g nonZero) | inl refl with decidableGroups i ((i + g) (Group.inverse (G i) g)) (Group.0G (G i))
|
||||
... | inr x = exFalso (x (Group.invRight (G i)))
|
||||
... | inl x = record {}
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (ofEmpty .k h nonZero1) j!=k) with decidableIndex k j
|
||||
... | inl x = exFalso (j!=k (equalityCommutative x))
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (ofEmpty .k h nonZero1) j!=k) | inr _ with decidableIndex k k
|
||||
... | inr x = exFalso (x refl)
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (ofEmpty .k h nonZero1) j!=k) | inr _ | inl refl with decidableGroups k ((k + h) (Group.inverse (G k) h)) (Group.0G (G k))
|
||||
... | inr x = exFalso (x (Group.invRight (G k)))
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (ofEmpty .k h nonZero1) j!=k) | inr _ | inl refl | inl _ with decidableIndex j j
|
||||
... | inr x = exFalso (x refl)
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (ofEmpty .k h nonZero1) j!=k) | inr _ | inl refl | inl _ | inl refl with decidableGroups j ((j + g) (Group.inverse (G j) g)) (Group.0G (G j))
|
||||
... | inr bad = exFalso (bad (Group.invRight (G j)))
|
||||
... | inl r = record {}
|
||||
invRight' {j} (prependLetter _ g nonZero {k} (prependLetter .k h nonZero1 {i} x k!=i) j!=k) rewrite refl {x = 0} = Equivalence.transitive (Setoid.eq freeProductSetoid) {prepend j g _ (prepend k h _ (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _)) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) _))))} {prepend j g nonZero (prepend k h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t)))) +RP (nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G j) t))))} {empty} (prependWD' g nonZero (prepend k h nonZero1 (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _)) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G j) t))))) (prepend k h _ (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t)))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) λ i → nonZero (invZeroImpliesZero (G j) i))) (Equivalence.symmetric (Setoid.eq freeProductSetoid) {(prepend k h _ (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) _))} {prepend k h _ (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _)) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) _)))} t)) (unpeel (prepend k h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _)))) g nonZero (λ t → nonZero (invZeroImpliesZero (G j) t)) (Equivalence.transitive (Setoid.eq freeProductSetoid) {prepend k h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t))))} {prepend k h nonZero1 ((plus' x (inv' x)) +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t)))} {empty} (prependWD' h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t)))) (plus' x (inv' x) +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t))) (Equivalence.symmetric (Setoid.eq freeProductSetoid) {plus' x (inv' x) +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G k) t))} (plusAssoc (nonempty _ x) (inv' x) (nonempty k (ofEmpty k (Group.inverse (G k) h) (λ t → nonZero1 (invZeroImpliesZero (G k) t))))))) (unpeel (plus' x (inv' x)) h nonZero1 (λ t → nonZero1 (invZeroImpliesZero (G k) t)) (invRight' x))))
|
||||
where
|
||||
t : (prepend k h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) (λ t → nonZero1 (invZeroImpliesZero (G _) t))))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) (λ t → nonZero (invZeroImpliesZero (G j) t)))) =RP (prepend k h nonZero1 (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) (λ t → nonZero1 (invZeroImpliesZero (G _) t)))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) (λ t → nonZero (invZeroImpliesZero (G j) t))))))
|
||||
t = Equivalence.transitive (Setoid.eq freeProductSetoid) {(prepend k h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) (λ t → nonZero1 (invZeroImpliesZero (G _) t))))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) (λ t → nonZero (invZeroImpliesZero (G j) t))))} {prepend k h _ (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G _) t))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G _) t)))} {(prepend k h nonZero1 (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) (λ t → nonZero1 (invZeroImpliesZero (G _) t)))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) (λ t → nonZero (invZeroImpliesZero (G j) t))))))} (prependAssocLemma' {k} {h} nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) _))) (nonempty j (ofEmpty j (Group.inverse (G j) g) _))) (prependWD' h nonZero1 (plus' x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G _) t))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G _) t))) (plus' x ((inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G _) t))) +RP nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G _) t)))) (plusAssocLemma x (inv' x +RP nonempty k (ofEmpty k (Group.inverse (G k) h) λ t → nonZero1 (invZeroImpliesZero (G _) t))) (nonempty j (ofEmpty j (Group.inverse (G j) g) λ t → nonZero (invZeroImpliesZero (G _) t)))))
|
||||
|
||||
abstract
|
||||
invRight : (x : ReducedSequence) → (x +RP (inv x)) =RP empty
|
||||
invRight empty = record {}
|
||||
invRight (nonempty i w) = invRight' {i} w
|
||||
|
||||
abstract
|
||||
invLeft' : {i : I} (x : ReducedSequenceBeginningWith i) → (inv (nonempty i x) +RP (nonempty i x)) =RP empty
|
||||
invLeft' {i} (ofEmpty .i g nonZero) with decidableIndex i i
|
||||
invLeft' {i} (ofEmpty .i g nonZero) | inl refl with decidableGroups i ((i + Group.inverse (G i) g) g) (Group.0G (G i))
|
||||
... | inl good = record {}
|
||||
... | inr bad = exFalso (bad (Group.invLeft (G i) {g}))
|
||||
invLeft' {i} (ofEmpty .i g nonZero) | inr x = exFalso (x refl)
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) with decidableIndex j i
|
||||
... | inl pr = exFalso (i!=j (equalityCommutative pr))
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr with decidableIndex i i
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr | inl refl with decidableGroups i ((i + Group.inverse (G i) g) g) (Group.0G (G i))
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr | inl refl | inl k with decidableIndex j j
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j h nonZero₁) i!=j) | inr pr | inl refl | inl k | inl refl with decidableGroups j ((j + Group.inverse (G j) h) h) (Group.0G (G j))
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j h nonZero₁) i!=j) | inr pr | inl refl | inl k | inl refl | inl good = record {}
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j h nonZero₁) i!=j) | inr pr | inl refl | inl k | inl refl | inr bad = exFalso (bad (Group.invLeft (G j)))
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr | inl refl | inl k | inr bad = exFalso (bad refl)
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr | inl refl | inr k = exFalso (k (Group.invLeft (G i) {g}))
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (ofEmpty j g₁ nonZero₁) i!=j) | inr pr | inr bad = exFalso (bad refl)
|
||||
invLeft' {i} (prependLetter .i g nonZero {.j} (prependLetter j h nonZero1 {k} w x) i!=j) = Equivalence.transitive (Setoid.eq freeProductSetoid) {(((inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) +RP nonempty i (ofEmpty i (Group.inverse (G i) g) _)) +RP nonempty i (prependLetter i g _ (prependLetter j h _ w x) i!=j))} {_} {empty} (plusAssoc (inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) (nonempty i (ofEmpty i (Group.inverse (G i) g) _)) (nonempty i (prependLetter i g _ (prependLetter j h _ w x) i!=j))) (Equivalence.transitive (Setoid.eq freeProductSetoid) {((inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) +RP (prepend i (Group.inverse (G i) g) _ (nonempty i (prependLetter i g _ (prependLetter j h _ w x) i!=j))))} {(inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) +RP (nonempty j (prependLetter j h nonZero1 w x))} {empty} (plusWD (inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) (prepend i (Group.inverse (G i) g) _ (nonempty i (prependLetter i g _ (prependLetter j h _ w x) i!=j))) (inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) (nonempty j (prependLetter j h _ w x)) (Equivalence.reflexive (Setoid.eq freeProductSetoid) {inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)}) (lemma1 {i} {j} {k} i!=j g h w x (λ p → nonZero (invZeroImpliesZero (G i) p)) nonZero nonZero1)) (Equivalence.transitive (Setoid.eq freeProductSetoid) {(inv' w +RP nonempty j (ofEmpty j (Group.inverse (G j) h) _)) +RP nonempty j (prependLetter j h _ w x)} {inv' w +RP (nonempty j (ofEmpty j (Group.inverse (G j) h) λ p → nonZero1 (invZeroImpliesZero (G j) p)) +RP nonempty j (prependLetter j h nonZero1 w x))} {empty} (plusAssoc (inv' w) (nonempty j (ofEmpty j (Group.inverse (G j) h) _)) (nonempty j (prependLetter j h _ w x))) (Equivalence.transitive (Setoid.eq freeProductSetoid) {inv' w +RP (prepend j (Group.inverse (G j) h) _ (nonempty j (prependLetter j h nonZero1 w x)))} {inv' w +RP (nonempty k w)} {empty} (plusWD (inv' w) (prepend j (Group.inverse (G j) h) _ (nonempty j (prependLetter j h nonZero1 w x))) (inv' w) (nonempty k w) (Equivalence.reflexive (Setoid.eq freeProductSetoid) {inv' w}) (lemma2 {j} {k} x h w (λ p → nonZero1 (invZeroImpliesZero (G j) p)) nonZero1)) (invLeft' {k} w))))
|
||||
|
||||
abstract
|
||||
invLeft : (x : ReducedSequence) → ((inv x) +RP x) =RP empty
|
||||
invLeft empty = record {}
|
||||
invLeft (nonempty i w) = invLeft' {i} w
|
||||
|
||||
FreeProductGroup : Group freeProductSetoid _+RP_
|
||||
Group.+WellDefined FreeProductGroup {m} {n} {x} {y} m=x n=y = plusWD m n x y m=x n=y
|
||||
Group.0G FreeProductGroup = empty
|
||||
Group.inverse FreeProductGroup = inv
|
||||
Group.+Associative FreeProductGroup {a} {b} {c} = Equivalence.symmetric (Setoid.eq freeProductSetoid) {(a +RP b) +RP c} {a +RP (b +RP c)} (plusAssoc a b c)
|
||||
Group.identRight FreeProductGroup {empty} = Equivalence.reflexive (Setoid.eq freeProductSetoid) {empty}
|
||||
Group.identRight FreeProductGroup {nonempty i x} rewrite refl {x = 0} = plusEmptyRight x
|
||||
Group.identLeft FreeProductGroup {empty} = Equivalence.reflexive (Setoid.eq freeProductSetoid) {empty}
|
||||
Group.identLeft FreeProductGroup {nonempty i x} = Equivalence.reflexive (Setoid.eq freeProductSetoid) {nonempty i x}
|
||||
Group.invLeft FreeProductGroup {x} = invLeft x
|
||||
Group.invRight FreeProductGroup {x} = invRight x
|
96
Groups/FreeProduct/Setoid.agda
Normal file
96
Groups/FreeProduct/Setoid.agda
Normal file
@@ -0,0 +1,96 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_; Setω)
|
||||
open import Setoids.Setoids
|
||||
open import Groups.Definition
|
||||
open import LogicalFormulae
|
||||
open import Orders.WellFounded.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.FreeProduct.Setoid {i : _} {I : Set i} (decidableIndex : (x y : I) → ((x ≡ y) || ((x ≡ y) → False))) {a b : _} {A : I → Set a} {S : (i : I) → Setoid {a} {b} (A i)} {_+_ : (i : I) → (A i) → (A i) → A i} (decidableGroups : (i : I) → (x y : A i) → ((Setoid._∼_ (S i) x y)) || ((Setoid._∼_ (S i) x y) → False)) (G : (i : I) → Group (S i) (_+_ i)) where
|
||||
|
||||
open import Groups.FreeProduct.Definition decidableIndex decidableGroups G
|
||||
|
||||
=RP' : {m n : _} (s1 : ReducedSequenceBeginningWith m) (s2 : ReducedSequenceBeginningWith n) → Set b
|
||||
=RP' (ofEmpty i g nonZero) (ofEmpty j h nonZero₁) with decidableIndex i j
|
||||
=RP' (ofEmpty i g nonZero) (ofEmpty .i h nonZero₁) | inl refl = Setoid._∼_ (S i) g h
|
||||
=RP' (ofEmpty i g nonZero) (ofEmpty j h nonZero₁) | inr x = False'
|
||||
=RP' (ofEmpty i g nonZero) (prependLetter i₁ g₁ nonZero₁ s2 x) = False'
|
||||
=RP' (prependLetter i g nonZero s1 x) (ofEmpty i₁ g₁ nonZero₁) = False'
|
||||
=RP' (prependLetter i g nonZero s1 x) (prependLetter j h nonZero₁ s2 x₁) with decidableIndex i j
|
||||
=RP' (prependLetter i g nonZero s1 x) (prependLetter .i h nonZero₁ s2 x₁) | inl refl = Setoid._∼_ (S i) g h && =RP' s1 s2
|
||||
=RP' (prependLetter i g nonZero s1 x) (prependLetter j h nonZero₁ s2 x₁) | inr x₂ = False'
|
||||
|
||||
_=RP_ : Rel ReducedSequence
|
||||
empty =RP empty = True'
|
||||
empty =RP nonempty i x = False'
|
||||
nonempty i x =RP empty = False'
|
||||
nonempty i x =RP nonempty j y = =RP' x y
|
||||
|
||||
=RP'reflex : {i : _} (x : ReducedSequenceBeginningWith i) → =RP' x x
|
||||
=RP'reflex (ofEmpty i g nonZero) with decidableIndex i i
|
||||
=RP'reflex (ofEmpty i g nonZero) | inl refl = Equivalence.reflexive (Setoid.eq (S i))
|
||||
=RP'reflex (ofEmpty i g nonZero) | inr x = exFalso (x refl)
|
||||
=RP'reflex (prependLetter i g nonZero x x₁) with decidableIndex i i
|
||||
=RP'reflex (prependLetter i g nonZero x x₁) | inl refl = Equivalence.reflexive (Setoid.eq (S i)) ,, =RP'reflex x
|
||||
=RP'reflex (prependLetter i g nonZero x x₁) | inr bad = exFalso (bad refl)
|
||||
|
||||
private
|
||||
reflex : Reflexive _=RP_
|
||||
reflex {empty} = record {}
|
||||
reflex {nonempty i (ofEmpty .i g nonZero)} with decidableIndex i i
|
||||
reflex {nonempty i (ofEmpty .i g nonZero)} | inl refl = Equivalence.reflexive (Setoid.eq (S i))
|
||||
... | inr bad = exFalso (bad refl)
|
||||
reflex {nonempty i (prependLetter .i g nonZero x x₁)} with decidableIndex i i
|
||||
reflex {nonempty i (prependLetter .i g nonZero x x₁)} | inl refl = Equivalence.reflexive (Setoid.eq (S i)) ,, =RP'reflex x
|
||||
reflex {nonempty i (prependLetter .i g nonZero x x₁)} | inr bad = exFalso (bad refl)
|
||||
|
||||
=RP'symm : {i j : _} (x : ReducedSequenceBeginningWith i) (y : ReducedSequenceBeginningWith j) → =RP' x y → =RP' y x
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) with decidableIndex i j
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) | inl pr with decidableIndex j i
|
||||
=RP'symm (ofEmpty .j g nonZero) (ofEmpty j h nonZero2) | inl refl | inl refl = Equivalence.symmetric (Setoid.eq (S j)) {g} {h}
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) | inl pr | inr x = exFalso (x (equalityCommutative pr))
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) | inr x with decidableIndex j i
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) | inr x | inl pr = exFalso (x (equalityCommutative pr))
|
||||
=RP'symm (ofEmpty i g nonZero) (ofEmpty j h nonZero2) | inr x | inr _ = id
|
||||
=RP'symm (ofEmpty i g nonZero) (prependLetter i₁ g₁ nonZero₁ y x) = id
|
||||
=RP'symm (prependLetter i g nonZero x x₁) (ofEmpty i₁ g₁ nonZero₁) = id
|
||||
=RP'symm (prependLetter i g nonZero x x₁) (prependLetter j h nonZero2 y pr2) pr with decidableIndex i j
|
||||
=RP'symm (prependLetter .j g nonZero x x₁) (prependLetter j h nonZero2 y pr2) pr | inl refl with decidableIndex j j
|
||||
=RP'symm (prependLetter .j g nonZero x x₁) (prependLetter j h nonZero2 y pr2) (fst ,, snd) | inl refl | inl refl = Equivalence.symmetric (Setoid.eq (S j)) fst ,, =RP'symm x y snd
|
||||
=RP'symm (prependLetter .j g nonZero x x₁) (prependLetter j h nonZero2 y pr2) pr | inl refl | inr bad = exFalso (bad refl)
|
||||
=RP'symm (prependLetter i g nonZero x x₁) (prependLetter j h nonZero2 y pr2) () | inr i!=j
|
||||
|
||||
private
|
||||
symm : Symmetric _=RP_
|
||||
symm {empty} {empty} x = record {}
|
||||
symm {nonempty i m} {nonempty i₁ n} x = =RP'symm m n x
|
||||
|
||||
=RP'trans : {i j k : _} (x : ReducedSequenceBeginningWith i) (y : ReducedSequenceBeginningWith j) (z : ReducedSequenceBeginningWith k) → =RP' x y → =RP' y z → =RP' x z
|
||||
=RP'trans (ofEmpty i g nonZero) (ofEmpty j g₁ nonZero₁) (ofEmpty k g₂ nonZero₂) x=y y=z with decidableIndex i j
|
||||
=RP'trans (ofEmpty .j g nonZero) (ofEmpty j g₁ nonZero₁) (ofEmpty k g₂ nonZero₂) x=y y=z | inl refl with decidableIndex j k
|
||||
=RP'trans (ofEmpty .j g nonZero) (ofEmpty j g₁ nonZero₁) (ofEmpty .j g₂ nonZero₂) x=y y=z | inl refl | inl refl = Equivalence.transitive (Setoid.eq (S j)) x=y y=z
|
||||
=RP'trans (prependLetter i g nonZero x x₁) (prependLetter j g₁ nonZero₁ y x₂) (prependLetter k g₂ nonZero₂ z x₃) x=y y=z with decidableIndex i j
|
||||
=RP'trans (prependLetter .j g nonZero x x₁) (prependLetter j g₁ nonZero₁ y x₂) (prependLetter k g₂ nonZero₂ z x₃) x=y y=z | inl refl with decidableIndex j k
|
||||
=RP'trans (prependLetter .j g nonZero x x₁) (prependLetter j g₁ nonZero₁ y x₂) (prependLetter .j g₂ nonZero₂ z x₃) (fst1 ,, snd1) (fst2 ,, snd2) | inl refl | inl refl = Equivalence.transitive (Setoid.eq (S j)) fst1 fst2 ,, =RP'trans x y z snd1 snd2
|
||||
|
||||
private
|
||||
trans : (x y z : ReducedSequence) → x =RP y → y =RP z → x =RP z
|
||||
trans empty empty empty x=y y=z = record {}
|
||||
trans (nonempty i x) (nonempty i₁ y) (nonempty i₂ z) x=y y=z = =RP'trans x y z x=y y=z
|
||||
|
||||
notEqualIfStartDifferent : {j1 j2 : I} (neq : (j1 ≡ j2) → False) → (x1 : ReducedSequenceBeginningWith j1) (x2 : ReducedSequenceBeginningWith j2) → =RP' x1 x2 → False
|
||||
notEqualIfStartDifferent neq (ofEmpty i g nonZero) (ofEmpty j g₁ nonZero₁) eq with decidableIndex i j
|
||||
notEqualIfStartDifferent neq (ofEmpty i g nonZero) (ofEmpty j g₁ nonZero₁) eq | inl i=j = neq i=j
|
||||
notEqualIfStartDifferent neq (prependLetter i g nonZero x1 x) (prependLetter j g₁ nonZero₁ x2 x₁) eq with decidableIndex i j
|
||||
notEqualIfStartDifferent neq (prependLetter i g nonZero x1 x) (prependLetter j g₁ nonZero₁ x2 x₁) eq | inl eq' = neq eq'
|
||||
|
||||
freeProductSetoid : Setoid ReducedSequence
|
||||
Setoid._∼_ freeProductSetoid = _=RP_
|
||||
Equivalence.reflexive (Setoid.eq freeProductSetoid) {x} = reflex {x}
|
||||
Equivalence.symmetric (Setoid.eq freeProductSetoid) {a} {b} = symm {a} {b}
|
||||
Equivalence.transitive (Setoid.eq freeProductSetoid) {x} {y} {z} = trans x y z
|
||||
|
@@ -32,7 +32,7 @@ GroupHom.groupHom (groupHomsCompose {U = U} {_+C_ = _·C_} {I} {g} gHom) {x} {y}
|
||||
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom hom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
|
||||
|
||||
homRespectsInverse : {x : A} → Setoid._∼_ T (f (Group.inverse G x)) (Group.inverse H (f x))
|
||||
homRespectsInverse {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Group.invLeft G)) imageOfIdentityIsIdentity))
|
||||
homRespectsInverse {x} = rightInversesAreUnique H {f x} {f (Group.inverse G x)} (transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Group.invLeft G)) imageOfIdentityIsIdentity))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
@@ -44,8 +44,8 @@ groupsHaveRightCancellation x y z pr = transitive m identRight
|
||||
m : y ∼ z · 0G
|
||||
m = transitive l (+WellDefined ~refl invRight)
|
||||
|
||||
rightInversesAreUnique : (x : A) → (y : A) → (y · x) ∼ 0G → y ∼ (inverse x)
|
||||
rightInversesAreUnique x y f = transitive i (transitive j (transitive k (transitive l m)))
|
||||
rightInversesAreUnique : {x : A} → {y : A} → (y · x) ∼ 0G → y ∼ (inverse x)
|
||||
rightInversesAreUnique {x} {y} f = transitive i (transitive j (transitive k (transitive l m)))
|
||||
where
|
||||
_^-1 = inverse
|
||||
i : y ∼ y · 0G
|
||||
@@ -60,7 +60,7 @@ rightInversesAreUnique x y f = transitive i (transitive j (transitive k (transit
|
||||
m = identLeft
|
||||
|
||||
leftInversesAreUnique : {x : A} → {y : A} → (x · y) ∼ 0G → y ∼ (inverse x)
|
||||
leftInversesAreUnique {x} {y} f = rightInversesAreUnique x y l
|
||||
leftInversesAreUnique {x} {y} f = rightInversesAreUnique {x} {y} l
|
||||
where
|
||||
_^-1 = inverse
|
||||
i : y · (x · y) ∼ y · 0G
|
||||
@@ -76,7 +76,7 @@ leftInversesAreUnique {x} {y} f = rightInversesAreUnique x y l
|
||||
|
||||
|
||||
invTwice : (x : A) → (inverse (inverse x)) ∼ x
|
||||
invTwice x = symmetric (rightInversesAreUnique (x ^-1) x invRight)
|
||||
invTwice x = symmetric (rightInversesAreUnique {x ^-1} {x} invRight)
|
||||
where
|
||||
_^-1 = inverse
|
||||
|
||||
@@ -139,3 +139,6 @@ invContravariant {x} {y} = ans
|
||||
|
||||
equalsDoubleImpliesZero : {x : A} → (x · x) ∼ x → x ∼ 0G
|
||||
equalsDoubleImpliesZero 2x=x = transitive (symmetric identLeft) (transitive (+WellDefined (symmetric invLeft) reflexive) (transitive (symmetric +Associative) (transitive (+WellDefined reflexive 2x=x) invLeft)))
|
||||
|
||||
invZeroImpliesZero : {x : A} → (inverse x) ∼ 0G → x ∼ 0G
|
||||
invZeroImpliesZero {x} pr = transitive (symmetric (invTwice x)) (transitive (inverseWellDefined pr) invIdent)
|
||||
|
81
Groups/Vector.agda
Normal file
81
Groups/Vector.agda
Normal file
@@ -0,0 +1,81 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Vectors
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Vector {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
|
||||
vecEquiv : {n : ℕ} → Vec A n → Vec A n → Set b
|
||||
vecEquiv {zero} [] [] = True'
|
||||
vecEquiv {succ n} (x ,- v1) (y ,- v2) = (x ∼ y) && vecEquiv v1 v2
|
||||
|
||||
private
|
||||
vecEquivRefl : {n : ℕ} → (x : Vec A n) → vecEquiv x x
|
||||
vecEquivRefl [] = record {}
|
||||
vecEquivRefl (x ,- xs) = reflexive ,, vecEquivRefl xs
|
||||
|
||||
vecEquivSymm : {n : ℕ} → (x y : Vec A n) → vecEquiv x y → vecEquiv y x
|
||||
vecEquivSymm [] [] record {} = record {}
|
||||
vecEquivSymm (x ,- xs) (y ,- ys) (fst ,, snd) = symmetric fst ,, vecEquivSymm _ _ snd
|
||||
|
||||
vecEquivTrans : {n : ℕ} → (x y z : Vec A n) → vecEquiv x y → vecEquiv y z → vecEquiv x z
|
||||
vecEquivTrans [] [] [] x=y y=z = record {}
|
||||
vecEquivTrans (x ,- xs) (y ,- ys) (z ,- zs) (fst1 ,, snd1) (fst2 ,, snd2) = transitive fst1 fst2 ,, vecEquivTrans xs ys zs snd1 snd2
|
||||
|
||||
vectorSetoid : (n : ℕ) → Setoid (Vec A n)
|
||||
Setoid._∼_ (vectorSetoid n) = vecEquiv {n}
|
||||
Equivalence.reflexive (Setoid.eq (vectorSetoid n)) {x} = vecEquivRefl x
|
||||
Equivalence.symmetric (Setoid.eq (vectorSetoid n)) {x} {y} = vecEquivSymm x y
|
||||
Equivalence.transitive (Setoid.eq (vectorSetoid n)) {x} {y} {z} = vecEquivTrans x y z
|
||||
|
||||
vectorAdd : {n : ℕ} → Vec A n → Vec A n → Vec A n
|
||||
vectorAdd [] [] = []
|
||||
vectorAdd (x ,- v1) (y ,- v2) = (x + y) ,- (vectorAdd v1 v2)
|
||||
|
||||
private
|
||||
addWellDefined : {n : ℕ} → (m k x y : Vec A n) → Setoid._∼_ (vectorSetoid n) m x → Setoid._∼_ (vectorSetoid n) k y → Setoid._∼_ (vectorSetoid n) (vectorAdd m k) (vectorAdd x y)
|
||||
addWellDefined [] [] [] [] m=x k=y = record {}
|
||||
addWellDefined (m ,- ms) (k ,- ks) (x ,- xs) (y ,- ys) (m=x ,, ms=xs) (k=y ,, ks=ys) = +WellDefined m=x k=y ,, addWellDefined ms ks xs ys ms=xs ks=ys
|
||||
|
||||
addAssoc : {n : ℕ} → (x y z : Vec A n) → Setoid._∼_ (vectorSetoid n) (vectorAdd x (vectorAdd y z)) (vectorAdd (vectorAdd x y) z)
|
||||
addAssoc [] [] [] = record {}
|
||||
addAssoc (x ,- xs) (y ,- ys) (z ,- zs) = +Associative ,, addAssoc xs ys zs
|
||||
|
||||
vecIdentRight : {n : ℕ} → (a : Vec A n) → Setoid._∼_ (vectorSetoid n) (vectorAdd a (vecPure 0G)) a
|
||||
vecIdentRight [] = record {}
|
||||
vecIdentRight (x ,- a) = identRight ,, vecIdentRight a
|
||||
|
||||
vecIdentLeft : {n : ℕ} → (a : Vec A n) → Setoid._∼_ (vectorSetoid n) (vectorAdd (vecPure 0G) a) a
|
||||
vecIdentLeft [] = record {}
|
||||
vecIdentLeft (x ,- a) = identLeft ,, vecIdentLeft a
|
||||
|
||||
vecInvLeft : {n : ℕ} → (a : Vec A n) → Setoid._∼_ (vectorSetoid n) (vectorAdd (vecMap inverse a) a) (vecPure 0G)
|
||||
vecInvLeft [] = record {}
|
||||
vecInvLeft (x ,- a) = invLeft ,, vecInvLeft a
|
||||
|
||||
vecInvRight : {n : ℕ} → (a : Vec A n) → Setoid._∼_ (vectorSetoid n) (vectorAdd a (vecMap inverse a)) (vecPure 0G)
|
||||
vecInvRight [] = record {}
|
||||
vecInvRight (x ,- a) = invRight ,, vecInvRight a
|
||||
|
||||
vectorGroup : {n : ℕ} → Group (vectorSetoid n) (vectorAdd {n})
|
||||
Group.+WellDefined vectorGroup {m} {n} {x} {y} = addWellDefined m n x y
|
||||
Group.0G vectorGroup = vecPure 0G
|
||||
Group.inverse vectorGroup x = vecMap inverse x
|
||||
Group.+Associative vectorGroup {x} {y} {z} = addAssoc x y z
|
||||
Group.identRight vectorGroup {a} = vecIdentRight a
|
||||
Group.identLeft vectorGroup {a} = vecIdentLeft a
|
||||
Group.invLeft vectorGroup {a} = vecInvLeft a
|
||||
Group.invRight vectorGroup {a} = vecInvRight a
|
||||
|
||||
abelianVectorGroup : {n : ℕ} → AbelianGroup G → AbelianGroup (vectorGroup {n})
|
||||
AbelianGroup.commutative (abelianVectorGroup grp) {[]} {[]} = record {}
|
||||
AbelianGroup.commutative (abelianVectorGroup grp) {x ,- xs} {y ,- ys} = AbelianGroup.commutative grp ,, AbelianGroup.commutative (abelianVectorGroup grp) {xs} {ys}
|
Reference in New Issue
Block a user