mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-16 08:58:39 +00:00
Free group universal property (#102)
This commit is contained in:
@@ -17,7 +17,7 @@ open import Sets.FinSetWithK
|
||||
open import Rings.Examples.Examples
|
||||
open import Rings.Examples.Proofs
|
||||
|
||||
open import Groups.FreeGroups
|
||||
open import Groups.FreeGroup.UniversalProperty
|
||||
open import Groups.Examples.ExampleSheet1
|
||||
open import Groups.LectureNotes.Lecture1
|
||||
|
||||
|
57
Groups/FreeGroup/Definition.agda
Normal file
57
Groups/FreeGroup/Definition.agda
Normal file
@@ -0,0 +1,57 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Decidable.Sets
|
||||
|
||||
module Groups.FreeGroup.Definition where
|
||||
|
||||
data FreeCompletion {a : _} (A : Set a) : Set a where
|
||||
ofLetter : A → FreeCompletion A
|
||||
ofInv : A → FreeCompletion A
|
||||
|
||||
freeInverse : {a : _} {A : Set a} (l : FreeCompletion A) → FreeCompletion A
|
||||
freeInverse (ofLetter x) = ofInv x
|
||||
freeInverse (ofInv x) = ofLetter x
|
||||
|
||||
ofLetterInjective : {a : _} {A : Set a} {x y : A} → (ofLetter x ≡ ofLetter y) → x ≡ y
|
||||
ofLetterInjective refl = refl
|
||||
|
||||
ofInvInjective : {a : _} {A : Set a} {x y : A} → (ofInv x ≡ ofInv y) → x ≡ y
|
||||
ofInvInjective refl = refl
|
||||
|
||||
decidableFreeCompletion : {a : _} {A : Set a} → DecidableSet A → DecidableSet (FreeCompletion A)
|
||||
decidableFreeCompletion {A = A} record { eq = dec } = record { eq = pr }
|
||||
where
|
||||
pr : (a b : FreeCompletion A) → (a ≡ b) || (a ≡ b → False)
|
||||
pr (ofLetter x) (ofLetter y) with dec x y
|
||||
... | inl refl = inl refl
|
||||
... | inr x!=y = inr λ p → x!=y (ofLetterInjective p)
|
||||
pr (ofLetter x) (ofInv y) = inr λ ()
|
||||
pr (ofInv x) (ofLetter y) = inr λ ()
|
||||
pr (ofInv x) (ofInv y) with dec x y
|
||||
... | inl refl = inl refl
|
||||
... | inr x!=y = inr λ p → x!=y (ofInvInjective p)
|
||||
|
||||
freeCompletionEqual : {a : _} {A : Set a} (dec : DecidableSet A) (x y : FreeCompletion A) → Bool
|
||||
freeCompletionEqual dec x y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqual dec x y | inl x₁ = BoolTrue
|
||||
freeCompletionEqual dec x y | inr x₁ = BoolFalse
|
||||
|
||||
freeCompletionEqualFalse : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → ((x ≡ y) → False) → (freeCompletionEqual dec x y) ≡ BoolFalse
|
||||
freeCompletionEqualFalse dec {x = x} {y} x!=y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inl x=y = exFalso (x!=y x=y)
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inr _ = refl
|
||||
|
||||
freeCompletionEqualFalse' : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → .((freeCompletionEqual dec x y) ≡ BoolFalse) → (x ≡ y) → False
|
||||
freeCompletionEqualFalse' dec {x} {y} pr with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse' dec {x} {y} () | inl x₁
|
||||
freeCompletionEqualFalse' dec {x} {y} pr | inr ans = ans
|
||||
|
132
Groups/FreeGroup/Group.agda
Normal file
132
Groups/FreeGroup/Group.agda
Normal file
@@ -0,0 +1,132 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Setoids.Setoids
|
||||
open import Groups.FreeGroup.Definition
|
||||
open import Groups.Definition
|
||||
open import Decidable.Sets
|
||||
open import Numbers.Naturals.Order
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
|
||||
module Groups.FreeGroup.Group {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
open import Groups.FreeGroup.Word decA
|
||||
|
||||
prepend : ReducedWord → FreeCompletion A → ReducedWord
|
||||
prepend empty x = prependLetter x empty (wordEmpty refl)
|
||||
prepend (prependLetter (ofLetter y) w pr) (ofLetter x) = prependLetter (ofLetter x) (prependLetter (ofLetter y) w pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA {ofLetter x} {ofInv x} λ ()))
|
||||
prepend (prependLetter (ofInv y) w pr) (ofLetter x) with DecidableSet.eq decA x y
|
||||
... | inl x=y = w
|
||||
... | inr x!=y = prependLetter (ofLetter x) (prependLetter (ofInv y) w pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA {ofLetter x} {ofLetter y} λ pr → x!=y (ofLetterInjective pr)))
|
||||
prepend (prependLetter (ofLetter y) w pr) (ofInv x) with DecidableSet.eq decA x y
|
||||
... | inl x=y = w
|
||||
... | inr x!=y = prependLetter (ofInv x) (prependLetter (ofLetter y) w pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA λ pr → x!=y (ofInvInjective pr)))
|
||||
prepend (prependLetter (ofInv y) w pr) (ofInv x) = prependLetter (ofInv x) (prependLetter (ofInv y) w pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA {ofInv x} {ofLetter x} (λ ())))
|
||||
|
||||
_+W_ : ReducedWord → ReducedWord → ReducedWord
|
||||
empty +W b = b
|
||||
prependLetter letter a x +W b = prepend (a +W b) letter
|
||||
|
||||
prependValid : (w : ReducedWord) (l : A) → (x : PrependIsValid w (ofLetter l)) → prepend w (ofLetter l) ≡ prependLetter (ofLetter l) w x
|
||||
prependValid empty l (wordEmpty refl) = refl
|
||||
prependValid (prependLetter (ofLetter l2) w x) l pr = prependLetterRefl
|
||||
prependValid (prependLetter (ofInv l2) w x) l pr with DecidableSet.eq decA l l2
|
||||
prependValid (prependLetter (ofInv l2) w x) .l2 (wordEnding _ x1) | inl refl = exFalso (freeCompletionEqualFalse' decA x1 refl)
|
||||
... | inr l!=l2 = prependLetterRefl
|
||||
|
||||
prependValid' : (w : ReducedWord) (l : A) → (x : PrependIsValid w (ofInv l)) → prepend w (ofInv l) ≡ prependLetter (ofInv l) w x
|
||||
prependValid' empty l (wordEmpty refl) = refl
|
||||
prependValid' (prependLetter (ofLetter l2) w x) l pr with DecidableSet.eq decA l l2
|
||||
prependValid' (prependLetter (ofLetter l2) w x) .l2 (wordEnding _ x1) | inl refl = exFalso (freeCompletionEqualFalse' decA x1 refl)
|
||||
... | inr l!=l2 = prependLetterRefl
|
||||
prependValid' (prependLetter (ofInv l2) w x) l pr = prependLetterRefl
|
||||
|
||||
prependInv : (w : ReducedWord) (l : A) → prepend (prepend w (ofLetter l)) (ofInv l) ≡ w
|
||||
prependInv empty l with DecidableSet.eq decA l l
|
||||
... | inl l=l = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
prependInv (prependLetter (ofLetter l2) w x) l with DecidableSet.eq decA l l
|
||||
... | inl l=l = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
prependInv (prependLetter (ofInv l2) w x) l with DecidableSet.eq decA l l2
|
||||
prependInv (prependLetter (ofInv l2) w x) .l2 | inl refl = prependValid' w l2 x
|
||||
... | inr l!=l2 with DecidableSet.eq decA l l
|
||||
prependInv (prependLetter (ofInv l2) w x) l | inr l!=l2 | inl refl = refl
|
||||
prependInv (prependLetter (ofInv l2) w x) l | inr l!=l2 | inr bad = exFalso (bad refl)
|
||||
|
||||
prependInv' : (w : ReducedWord) (l : A) → prepend (prepend w (ofInv l)) (ofLetter l) ≡ w
|
||||
prependInv' empty l with DecidableSet.eq decA l l
|
||||
... | inl l=l = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
prependInv' (prependLetter (ofLetter l2) w x) l with DecidableSet.eq decA l l2
|
||||
prependInv' (prependLetter (ofLetter l2) w x) .l2 | inl refl = prependValid w l2 x
|
||||
... | inr l!=l2 with DecidableSet.eq decA l l
|
||||
... | inl refl = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
prependInv' (prependLetter (ofInv l2) w x) l with DecidableSet.eq decA l l
|
||||
prependInv' (prependLetter (ofInv l2) w x) l | inl refl = refl
|
||||
prependInv' (prependLetter (ofInv l2) w x) l | inr l!=l = exFalso (l!=l refl)
|
||||
|
||||
prependAndAdd : (a b : ReducedWord) (l : FreeCompletion A) → prepend (a +W b) l ≡ (prepend a l) +W b
|
||||
prependAndAdd empty b l = refl
|
||||
prependAndAdd (prependLetter (ofLetter x) w pr) b (ofLetter y) = refl
|
||||
prependAndAdd (prependLetter (ofLetter x) w pr) b (ofInv y) with DecidableSet.eq decA y x
|
||||
prependAndAdd (prependLetter (ofLetter x) w pr) b (ofInv .x) | inl refl = prependInv _ _
|
||||
... | inr y!=x = refl
|
||||
prependAndAdd (prependLetter (ofInv x) w pr) b (ofLetter y) with DecidableSet.eq decA y x
|
||||
prependAndAdd (prependLetter (ofInv x) w pr) b (ofLetter .x) | inl refl = prependInv' _ _
|
||||
... | inr y!=x = refl
|
||||
prependAndAdd (prependLetter (ofInv x) w pr) b (ofInv y) = refl
|
||||
|
||||
+WAssoc : (a b c : ReducedWord) → (a +W (b +W c)) ≡ ((a +W b) +W c)
|
||||
+WAssoc empty b c = refl
|
||||
+WAssoc (prependLetter letter a x) b c rewrite equalityCommutative (prependAndAdd (a +W b) c letter) | +WAssoc a b c = refl
|
||||
|
||||
inverseW : ReducedWord → ReducedWord
|
||||
inverseW empty = empty
|
||||
inverseW (prependLetter letter w x) = (inverseW w) +W prependLetter (freeInverse letter) empty (wordEmpty refl)
|
||||
|
||||
identRightW : (a : ReducedWord) → a +W empty ≡ a
|
||||
identRightW empty = refl
|
||||
identRightW (prependLetter (ofLetter l) a x) rewrite identRightW a = prependValid a l x
|
||||
identRightW (prependLetter (ofInv l) a x) rewrite identRightW a = prependValid' a l x
|
||||
|
||||
invLeftW : (a : ReducedWord) → (inverseW a) +W a ≡ empty
|
||||
invLeftW empty = refl
|
||||
invLeftW (prependLetter (ofLetter l) a x) rewrite equalityCommutative (+WAssoc (inverseW a) (prependLetter (ofInv l) empty (wordEmpty refl)) (prependLetter (ofLetter l) a x)) = t
|
||||
where
|
||||
t : (inverseW a +W (prepend (prependLetter (ofLetter l) a x) (ofInv l))) ≡ empty
|
||||
t with DecidableSet.eq decA l l
|
||||
... | inl refl = invLeftW a
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
invLeftW (prependLetter (ofInv l) a x) rewrite equalityCommutative (+WAssoc (inverseW a) (prependLetter (ofLetter l) empty (wordEmpty refl)) (prependLetter (ofInv l) a x)) = t
|
||||
where
|
||||
t : (inverseW a +W (prepend (prependLetter (ofInv l) a x) (ofLetter l))) ≡ empty
|
||||
t with DecidableSet.eq decA l l
|
||||
... | inl refl = invLeftW a
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
|
||||
invRightW : (a : ReducedWord) → a +W (inverseW a) ≡ empty
|
||||
invRightW empty = refl
|
||||
invRightW (prependLetter (ofLetter l) a x) rewrite +WAssoc a (inverseW a) (prependLetter (ofInv l) empty (wordEmpty refl)) | invRightW a = t
|
||||
where
|
||||
t : prepend (prependLetter (ofInv l) empty (wordEmpty refl)) (ofLetter l) ≡ empty
|
||||
t with DecidableSet.eq decA l l
|
||||
... | inl refl = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
invRightW (prependLetter (ofInv l) a x) rewrite +WAssoc a (inverseW a) (prependLetter (ofLetter l) empty (wordEmpty refl)) | invRightW a = t
|
||||
where
|
||||
t : prepend (prependLetter (ofLetter l) empty (wordEmpty refl)) (ofInv l) ≡ empty
|
||||
t with DecidableSet.eq decA l l
|
||||
... | inl refl = refl
|
||||
... | inr l!=l = exFalso (l!=l refl)
|
||||
|
||||
freeGroup : Group (reflSetoid ReducedWord) _+W_
|
||||
Group.+WellDefined freeGroup refl refl = refl
|
||||
Group.0G freeGroup = empty
|
||||
Group.inverse freeGroup = inverseW
|
||||
Group.+Associative freeGroup {a} {b} {c} = +WAssoc a b c
|
||||
Group.identRight freeGroup {a} = identRightW a
|
||||
Group.identLeft freeGroup {a} = refl
|
||||
Group.invLeft freeGroup {a} = invLeftW a
|
||||
Group.invRight freeGroup {a} = invRightW a
|
111
Groups/FreeGroup/UniversalProperty.agda
Normal file
111
Groups/FreeGroup/UniversalProperty.agda
Normal file
@@ -0,0 +1,111 @@
|
||||
{-# 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.Homomorphisms.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Decidable.Sets
|
||||
open import Numbers.Naturals.Order
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.FreeGroup.UniversalProperty {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
open import Groups.FreeGroup.Word decA
|
||||
open import Groups.FreeGroup.Group decA
|
||||
|
||||
universalPropertyFunction : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} (G : Group S _+_) → (f : A → C) → ReducedWord → C
|
||||
universalPropertyFunction G f empty = Group.0G G
|
||||
universalPropertyFunction {_+_ = _+_} G f (prependLetter (ofLetter l) w x) = (f l) + universalPropertyFunction G f w
|
||||
universalPropertyFunction {_+_ = _+_} G f (prependLetter (ofInv l) w x) = (Group.inverse G (f l)) + universalPropertyFunction G f w
|
||||
|
||||
private
|
||||
prepLemma : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} → (G : Group S _+_) → (f : A → C) → (x : ReducedWord) (l : _) → Setoid._∼_ S (universalPropertyFunction G f (prepend x (ofLetter l))) ((f l) + universalPropertyFunction G f x)
|
||||
prepLemma {S = S} G f empty l = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
prepLemma {S = S} G f (prependLetter (ofLetter x) w pr) l = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
prepLemma {S = S} G f (prependLetter (ofInv x) w pr) l with DecidableSet.eq decA l x
|
||||
... | inl refl = transitive (symmetric identLeft) (transitive (+WellDefined (symmetric invRight) reflexive) (symmetric +Associative))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
... | inr l!=x = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
prepLemma' : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} → (G : Group S _+_) → (f : A → C) → (x : ReducedWord) (l : _) → Setoid._∼_ S (universalPropertyFunction G f (prepend x (ofInv l))) (Group.inverse G (f l) + universalPropertyFunction G f x)
|
||||
prepLemma' {S = S} G f empty l = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
prepLemma' {S = S} G f (prependLetter (ofLetter x) w pr) l with DecidableSet.eq decA l x
|
||||
... | inl refl = symmetric (transitive +Associative (transitive (+WellDefined invLeft reflexive) identLeft))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
... | inr l!=x = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
prepLemma' {S = S} G f (prependLetter (ofInv x) w pr) l = reflexive
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
homLemma : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} → (G : Group S _+_) → (f : A → C) → (x y : ReducedWord) → Setoid._∼_ S (universalPropertyFunction G f (x +W y)) (universalPropertyFunction G f x + universalPropertyFunction G f y)
|
||||
homLemma {S = S} G f empty y = symmetric identLeft
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
homLemma {S = S} G f (prependLetter (ofLetter l) x pr) y = transitive (transitive (prepLemma G f (x +W y) l) (+WellDefined reflexive (homLemma G f x y))) (+Associative {f l})
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
homLemma {S = S} G f (prependLetter (ofInv l) x pr) y = transitive (transitive (prepLemma' G f (x +W y) l) (+WellDefined reflexive (homLemma G f x y))) +Associative
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
|
||||
universalPropertyHom : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} (G : Group S _+_) → (f : A → C) → GroupHom freeGroup G (universalPropertyFunction G f)
|
||||
GroupHom.groupHom (universalPropertyHom {S = S} G f) {x} {y} = homLemma G f x y
|
||||
GroupHom.wellDefined (universalPropertyHom {S = S} G f) refl = Equivalence.reflexive (Setoid.eq S)
|
||||
|
||||
freeEmbedding : A → ReducedWord
|
||||
freeEmbedding a = prependLetter (ofLetter a) empty (wordEmpty refl)
|
||||
|
||||
freeEmbeddingInjective : {a b : A} → freeEmbedding a ≡ freeEmbedding b → a ≡ b
|
||||
freeEmbeddingInjective {a} {b} pr = ofLetterInjective (prependLetterInjective' pr)
|
||||
|
||||
universalPropertyDiagram : {c d : _} {C : Set c} {S : Setoid {c} {d} C} → {_+_ : C → C → C} → (G : Group S _+_) → (f : A → C) → (x : A) → Setoid._∼_ S (f x) (universalPropertyFunction G f (freeEmbedding x))
|
||||
universalPropertyDiagram {S = S} G f x = symmetric (Group.identRight G)
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
universalPropertyUniqueness : {c d : _} {C : Set c} {S : Setoid {c} {d} C} → {_+_ : C → C → C} → (G : Group S _+_) → (f : A → C) → {f' : ReducedWord → C} → (GroupHom freeGroup G f') → ((x : A) → Setoid._∼_ S (f x) (f' (freeEmbedding x))) → (x : ReducedWord) → Setoid._∼_ S (f' x) (universalPropertyFunction G f x)
|
||||
universalPropertyUniqueness {S = S} G f {f'} f'IsHom commutes empty = imageOfIdentityIsIdentity f'IsHom
|
||||
universalPropertyUniqueness {S = S} G f {f'} f'IsHom commutes (prependLetter (ofLetter a) w pr) = transitive (transitive (GroupHom.wellDefined f'IsHom (equalityCommutative (prependValid w a pr))) (GroupHom.groupHom f'IsHom)) (+WellDefined (symmetric (commutes a)) (universalPropertyUniqueness G f f'IsHom commutes w))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
universalPropertyUniqueness {S = S} G f {f'} f'IsHom commutes (prependLetter (ofInv a) w pr) = transitive (transitive (GroupHom.wellDefined f'IsHom (equalityCommutative (prependValid' w a pr))) (transitive (GroupHom.groupHom f'IsHom) (+WellDefined (homRespectsInverse f'IsHom) reflexive))) (+WellDefined (symmetric (inverseWellDefined G (commutes a))) (universalPropertyUniqueness G f f'IsHom commutes w))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Group G
|
||||
|
@@ -1,114 +1,68 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import WithK
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Groups.FreeGroup.Definition
|
||||
open import Decidable.Sets
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Decidable.Sets
|
||||
open import LogicalFormulae
|
||||
|
||||
module Groups.FreeGroups where
|
||||
module Groups.FreeGroup.Word {a : _} {A : Set a} (decA : DecidableSet A) where
|
||||
|
||||
data FreeCompletion {a : _} (A : Set a) : Set a where
|
||||
ofLetter : A → FreeCompletion A
|
||||
ofInv : A → FreeCompletion A
|
||||
data ReducedWord : Set a
|
||||
wordLength : ReducedWord → ℕ
|
||||
firstLetter : (w : ReducedWord) → .(0 <N wordLength w) → FreeCompletion A
|
||||
|
||||
freeInverse : {a : _} {A : Set a} (l : FreeCompletion A) → FreeCompletion A
|
||||
freeInverse (ofLetter x) = ofInv x
|
||||
freeInverse (ofInv x) = ofLetter x
|
||||
data PrependIsValid (w : ReducedWord) (l : FreeCompletion A) : Set a where
|
||||
wordEmpty : wordLength w ≡ 0 → PrependIsValid w l
|
||||
wordEnding : .(pr : 0 <N wordLength w) → .(freeCompletionEqual decA l (freeInverse (firstLetter w pr)) ≡ BoolFalse) → PrependIsValid w l
|
||||
|
||||
ofLetterInjective : {a : _} {A : Set a} {x y : A} → (ofLetter x ≡ ofLetter y) → x ≡ y
|
||||
ofLetterInjective refl = refl
|
||||
data ReducedWord where
|
||||
empty : ReducedWord
|
||||
prependLetter : (letter : FreeCompletion A) → (w : ReducedWord) → PrependIsValid w letter → ReducedWord
|
||||
|
||||
ofInvInjective : {a : _} {A : Set a} {x y : A} → (ofInv x ≡ ofInv y) → x ≡ y
|
||||
ofInvInjective refl = refl
|
||||
firstLetter empty ()
|
||||
firstLetter (prependLetter letter w x) pr = letter
|
||||
|
||||
decidableFreeCompletion : {a : _} {A : Set a} → DecidableSet A → DecidableSet (FreeCompletion A)
|
||||
decidableFreeCompletion {A = A} record { eq = dec } = record { eq = pr }
|
||||
where
|
||||
pr : (a b : FreeCompletion A) → (a ≡ b) || (a ≡ b → False)
|
||||
pr (ofLetter x) (ofLetter y) with dec x y
|
||||
... | inl refl = inl refl
|
||||
... | inr x!=y = inr λ p → x!=y (ofLetterInjective p)
|
||||
pr (ofLetter x) (ofInv y) = inr λ ()
|
||||
pr (ofInv x) (ofLetter y) = inr λ ()
|
||||
pr (ofInv x) (ofInv y) with dec x y
|
||||
... | inl refl = inl refl
|
||||
... | inr x!=y = inr λ p → x!=y (ofInvInjective p)
|
||||
wordLength empty = 0
|
||||
wordLength (prependLetter letter w pr) = succ (wordLength w)
|
||||
|
||||
freeCompletionEqual : {a : _} {A : Set a} (dec : DecidableSet A) (x y : FreeCompletion A) → Bool
|
||||
freeCompletionEqual dec x y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqual dec x y | inl x₁ = BoolTrue
|
||||
freeCompletionEqual dec x y | inr x₁ = BoolFalse
|
||||
prependLetterRefl : {x : FreeCompletion A} {w : ReducedWord} → {pr1 pr2 : PrependIsValid w x} → prependLetter x w pr1 ≡ prependLetter x w pr2
|
||||
prependLetterRefl {x} {empty} {wordEmpty refl} {wordEmpty refl} = refl
|
||||
prependLetterRefl {x} {empty} {wordEmpty refl} {wordEnding () x₂}
|
||||
prependLetterRefl {x} {empty} {wordEnding () x₁} {pr2}
|
||||
prependLetterRefl {x} {prependLetter letter w x₁} {wordEmpty ()} {pr2}
|
||||
prependLetterRefl {x} {prependLetter letter w x₁} {wordEnding pr x₂} {wordEmpty ()}
|
||||
prependLetterRefl {x} {prependLetter letter w x₁} {wordEnding pr2 r2} {wordEnding pr1 r1} = refl
|
||||
|
||||
freeCompletionEqualFalse : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → ((x ≡ y) → False) → (freeCompletionEqual dec x y) ≡ BoolFalse
|
||||
freeCompletionEqualFalse dec {x = x} {y} x!=y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inl x=y = exFalso (x!=y x=y)
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inr _ = refl
|
||||
prependLetterInjective : {x : FreeCompletion A} {w1 w2 : ReducedWord} {pr1 : PrependIsValid w1 x} {pr2 : PrependIsValid w2 x} → prependLetter x w1 pr1 ≡ prependLetter x w2 pr2 → w1 ≡ w2
|
||||
prependLetterInjective {x = x} {empty} {empty} {pr1} {pr2} pr = refl
|
||||
prependLetterInjective {x = x} {prependLetter l1 w1 x1} {prependLetter .l1 .w1 .x1} {wordEnding pr₁ x₁} {wordEnding pr₂ x₂} refl = refl
|
||||
|
||||
freeCompletionEqualFalse' : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → (freeCompletionEqual dec x y) ≡ BoolFalse → (x ≡ y) → False
|
||||
freeCompletionEqualFalse' dec {x} {y} pr with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse' dec {x} {y} () | inl x₁
|
||||
freeCompletionEqualFalse' dec {x} {y} pr | inr ans = ans
|
||||
|
||||
data ReducedWord {a : _} {A : Set a} (decA : DecidableSet A) : Set a
|
||||
wordLength : {a : _} {A : Set a} {decA : DecidableSet A} → ReducedWord decA → ℕ
|
||||
firstLetter : {a : _} {A : Set a} {decA : DecidableSet A} → (w : ReducedWord decA) → (0 <N wordLength w) → FreeCompletion A
|
||||
|
||||
data PrependIsValid {a : _} {A : Set a} (decA : DecidableSet A) (w : ReducedWord decA) (l : FreeCompletion A) : Set a where
|
||||
wordEmpty : wordLength w ≡ 0 → PrependIsValid decA w l
|
||||
wordEnding : (pr : 0 <N wordLength w) → (freeCompletionEqual decA l (freeInverse (firstLetter w pr)) ≡ BoolFalse) → PrependIsValid decA w l
|
||||
|
||||
data ReducedWord {a} {A} decA where
|
||||
empty : ReducedWord decA
|
||||
prependLetter : (letter : FreeCompletion A) → (w : ReducedWord decA) → PrependIsValid decA w letter → ReducedWord decA
|
||||
|
||||
firstLetter {a} {A} empty (le x ())
|
||||
firstLetter {a} {A} (prependLetter letter w x) pr = letter
|
||||
|
||||
wordLength {a} {A} empty = 0
|
||||
wordLength {a} {A} (prependLetter letter w pr) = succ (wordLength w)
|
||||
|
||||
data FreeGroupGenerators {a : _} (A : Set a) : Set a where
|
||||
χ : A → FreeGroupGenerators A
|
||||
|
||||
prependLetterInjective : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 x} → prependLetter x w1 pr1 ≡ prependLetter x w2 pr2 → w1 ≡ w2
|
||||
prependLetterInjective {x = x} {w1} {.w1} {pr1} {.pr1} refl = refl
|
||||
|
||||
prependLetterInjective' : {a : _} {A : Set a} {decA : DecidableSet A} {x y : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 y} → prependLetter x w1 pr1 ≡ prependLetter y w2 pr2 → x ≡ y
|
||||
prependLetterInjective' : {x y : FreeCompletion A} {w1 w2 : ReducedWord} {pr1 : PrependIsValid w1 x} {pr2 : PrependIsValid w2 y} → prependLetter x w1 pr1 ≡ prependLetter y w2 pr2 → x ≡ y
|
||||
prependLetterInjective' refl = refl
|
||||
|
||||
prependLetterRefl : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w : ReducedWord decA} → {pr1 pr2 : PrependIsValid decA w x} → prependLetter x w pr1 ≡ prependLetter x w pr2
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEmpty refl} = refl
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEnding (le x₁ ()) x₂}
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEnding (le x₂ ()) x₁} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEmpty ()} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr x₂} {wordEmpty ()}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr2 r2} {wordEnding pr1 r1} rewrite <NRefl pr1 (succIsPositive (wordLength w)) | <NRefl pr2 (succIsPositive (wordLength w)) | reflRefl r1 r2 = refl
|
||||
|
||||
badPrepend : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofInv x)} → (PrependIsValid decA (prependLetter (ofInv x) w pr) (ofLetter x)) → False
|
||||
badPrepend : {x : A} {w : ReducedWord} {pr : PrependIsValid w (ofInv x)} → (PrependIsValid (prependLetter (ofInv x) w pr) (ofLetter x)) → False
|
||||
badPrepend (wordEmpty ())
|
||||
badPrepend {decA = decA} {x = x} (wordEnding pr bad) with DecidableSet.eq (decidableFreeCompletion decA) (ofLetter x) (ofLetter x)
|
||||
badPrepend {decA = decA} {x} (wordEnding pr ()) | inl x₁
|
||||
badPrepend {decA = decA} {x} (wordEnding pr bad) | inr pr2 = pr2 refl
|
||||
badPrepend {x = x} (wordEnding pr bad) with DecidableSet.eq (decidableFreeCompletion decA) (ofLetter x) (ofLetter x)
|
||||
badPrepend {x} (wordEnding pr ()) | inl x₁
|
||||
badPrepend {x} (wordEnding pr bad) | inr pr2 = pr2 refl
|
||||
|
||||
badPrepend' : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofLetter x)} → (PrependIsValid decA (prependLetter (ofLetter x) w pr) (ofInv x)) → False
|
||||
badPrepend' : {x : A} {w : ReducedWord} {pr : PrependIsValid w (ofLetter x)} → (PrependIsValid (prependLetter (ofLetter x) w pr) (ofInv x)) → False
|
||||
badPrepend' (wordEmpty ())
|
||||
badPrepend' {decA = decA} {x = x} (wordEnding pr x₁) with DecidableSet.eq (decidableFreeCompletion decA) (ofInv x) (ofInv x)
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr ()) | inl x₂
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr x₁) | inr pr2 = pr2 refl
|
||||
badPrepend' {x = x} (wordEnding pr x₁) with DecidableSet.eq (decidableFreeCompletion decA) (ofInv x) (ofInv x)
|
||||
badPrepend' {x} (wordEnding pr ()) | inl x₂
|
||||
badPrepend' {x} (wordEnding pr x₁) | inr pr2 = pr2 refl
|
||||
|
||||
freeGroupGenerators : {a : _} (A : Set a) (decA : DecidableSet A) (w : FreeGroupGenerators A) → SymmetryGroupElements (reflSetoid (ReducedWord decA))
|
||||
freeGroupGenerators A decA (χ x) = sym {f = f} bij
|
||||
data FreeGroupGenerators : Set a where
|
||||
χ : A → FreeGroupGenerators
|
||||
|
||||
freeGroupGenerators : (w : FreeGroupGenerators) → SymmetryGroupElements (reflSetoid (ReducedWord))
|
||||
freeGroupGenerators (χ x) = sym {f = f} bij
|
||||
where
|
||||
open DecidableSet decA
|
||||
f : ReducedWord decA → ReducedWord decA
|
||||
f : ReducedWord → ReducedWord
|
||||
f empty = prependLetter (ofLetter x) empty (wordEmpty refl)
|
||||
f (prependLetter (ofLetter startLetter) w pr) = prependLetter (ofLetter x) (prependLetter (ofLetter startLetter) w pr) (wordEnding (succIsPositive _) ans)
|
||||
where
|
||||
@@ -124,13 +78,13 @@ freeGroupGenerators A decA (χ x) = sym {f = f} bij
|
||||
ans | bl with DecidableSet.eq decA x startLetter
|
||||
ans | bl | inl x=sl = exFalso (startLetter!=x (equalityCommutative x=sl))
|
||||
ans | bl | inr x!=sl = refl
|
||||
bij : SetoidBijection (reflSetoid (ReducedWord decA)) (reflSetoid (ReducedWord decA)) f
|
||||
bij : SetoidBijection (reflSetoid ReducedWord) (reflSetoid ReducedWord) f
|
||||
SetoidInjection.wellDefined (SetoidBijection.inj bij) x=y rewrite x=y = refl
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {empty} fx=fy = refl
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofLetter x₁) y pr1} ()
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) y pr1} fx=fy with DecidableSet.eq decA l x
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) empty (wordEmpty x₁)} () | inl l=x
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) empty (wordEnding (le x ()) x₁)} fx=fy | inl l=x
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) empty (wordEnding () x₁)} fx=fy | inl l=x
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) (prependLetter (ofLetter x₂) y x₁) (wordEmpty ())} fx=fy | inl l=x
|
||||
SetoidInjection.injective (SetoidBijection.inj bij) {empty} {prependLetter (ofInv l) (prependLetter (ofLetter x₂) y x₁) (wordEnding pr bad)} fx=fy | inl refl with ofLetterInjective (prependLetterInjective' fx=fy)
|
||||
... | l=x2 = exFalso ((freeCompletionEqualFalse' decA bad) (applyEquality ofInv l=x2))
|
||||
@@ -186,7 +140,7 @@ freeGroupGenerators A decA (χ x) = sym {f = f} bij
|
||||
needed | inr x!=x = exFalso (x!=x refl)
|
||||
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter l) empty pr} with DecidableSet.eq decA x l
|
||||
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter x) empty (wordEmpty refl)} | inl refl = empty , refl
|
||||
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter x) empty (wordEnding (le x₂ ()) x₁)} | inl refl
|
||||
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter x) empty (wordEnding () x₁)} | inl refl
|
||||
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter l) empty pr} | inr x!=l = prependLetter (ofInv x) (prependLetter (ofLetter l) empty pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA (λ p → x!=l (ofInvInjective p)))) , needed
|
||||
where
|
||||
needed : f (prependLetter (ofInv x) (prependLetter (ofLetter l) empty pr) (wordEnding (succIsPositive 0) (freeCompletionEqualFalse decA {ofInv x} {ofInv l} λ p → x!=l (ofInvInjective p)))) ≡ prependLetter (ofLetter l) empty pr
|
@@ -1,3 +1,4 @@
|
||||
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
@@ -69,4 +70,5 @@ increasingBoundedLimit x increasing (K , kIsBound) with lub (sequenceSubset x) (
|
||||
... | a , _ = a
|
||||
|
||||
increasingBoundedConverges : (x : Sequence A) → (increasing : Increasing x) → (bounded : Bounded x) → x ~> (increasingBoundedLimit x increasing bounded)
|
||||
increasingBoundedConverges x increasing bounded = {!!}
|
||||
increasingBoundedConverges x increasing (K , prBound) with lub (sequenceSubset x) (Sequence.head x , (0 , reflexive)) (K , boundedSequenceBounds K x prBound)
|
||||
... | lub , isLub = {!!}
|
||||
|
Reference in New Issue
Block a user