Finite permutations (#23)

This commit is contained in:
Patrick Stevens
2019-02-18 08:30:26 +00:00
committed by GitHub
parent c35fa90951
commit 0650021db1
12 changed files with 454 additions and 109 deletions

View File

@@ -1,35 +1,39 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import FinSet
open import Groups
open import GroupActions
open import Numbers.Naturals
open import Sets.FinSet
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupActions
module SymmetryGroups where
module Groups.SymmetryGroups where
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a b) where
sym : {f : A A} SetoidBijection S S f SymmetryGroupElements S
data ExtensionallyEqual {a b : _} {A : Set a} {B : Set b} (f g : A B) : Set (a b) where
eq : ({x : A} (f x) (g x)) ExtensionallyEqual f g
WellDefined : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) Set (a c d)
WellDefined {A = A} S T f = {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
extensionallyEqualReflexive : {a b : _} {A : Set a} {B : Set b} (f : A B) ExtensionallyEqual f f
extensionallyEqualReflexive f = eq (λ {x} refl)
data ExtensionallyEqual {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g : A B) (fWd : WellDefined S T f) (gWd : WellDefined S T g) : Set (a b c d) where
eq : ({x : A} Setoid.__ T (f x) (g x)) ExtensionallyEqual S T f g fWd gWd
extensionallyEqualSymmetric : {a b : _} {A : Set a} {B : Set b} {f g : A B} ExtensionallyEqual f g ExtensionallyEqual g f
extensionallyEqualSymmetric {f} {g} (eq pr) = eq λ {x} equalityCommutative (pr {x})
extensionallyEqualReflexive : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) (fWD1 fWD2 : WellDefined S T f) ExtensionallyEqual S T f f fWD1 fWD2
extensionallyEqualReflexive S T f fWD1 _ = eq (Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq T)))
extensionallyEqualTransitive : {a b : _} {A : Set a} {B : Set b} {f g h : A B} ExtensionallyEqual f g ExtensionallyEqual g h ExtensionallyEqual f h
extensionallyEqualTransitive (eq pr1) (eq pr2) = eq λ {x} transitivity pr1 pr2
extensionallyEqualSymmetric : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g : A B) (fWD : WellDefined S T f) (gWD : WellDefined S T g) ExtensionallyEqual S T f g fWD gWD ExtensionallyEqual S T g f gWD fWD
extensionallyEqualSymmetric S T f g fWD gWD (eq pr) = eq (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq T)) pr)
extensionallyEqualTransitive : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g h : A B) (fWD : WellDefined S T f) (gWD : WellDefined S T g) (hWD : WellDefined S T h) ExtensionallyEqual S T f g fWD gWD ExtensionallyEqual S T g h gWD hWD ExtensionallyEqual S T f h fWD hWD
extensionallyEqualTransitive S T f g h fWD gWD hWD (eq pr1) (eq pr2) = eq (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq T)) pr1 pr2)
symmetricSetoid : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Setoid (SymmetryGroupElements S)
Setoid.__ (symmetricSetoid A) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual f g
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} = extensionallyEqualReflexive f
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric f~g
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (symmetricSetoid A))) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive f~g g~h
Setoid.__ (symmetricSetoid S) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG)
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (symmetricSetoid S))) {sym {f} bijF} = extensionallyEqualReflexive S S f (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijF)
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (symmetricSetoid S))) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) f~g
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (symmetricSetoid S))) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive S S f g h (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) (SetoidBijection.wellDefined bijH) f~g g~h
symmetricGroupOp : {a b : _} {A : Set a} {S : Setoid {a} {b} A} (f g : SymmetryGroupElements S) SymmetryGroupElements S
symmetricGroupOp (sym {f} bijF) (sym {g} bijG) = sym (setoidBijComp bijF bijG)
@@ -39,23 +43,38 @@ module SymmetryGroups where
... | record { inverse = inverse ; inverseWellDefined = iwd ; isLeft = isLeft ; isRight = isRight } = sym (setoidInvertibleImpliesBijective (record { fWellDefined = iwd ; inverse = f ; inverseWellDefined = SetoidInjection.wellDefined (SetoidBijection.inj bijF) ; isLeft = λ b isRight b ; isRight = λ b isLeft b }))
symmetricGroupInvIsLeft : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp (symmetricGroupInv S x) x) (sym setoidIdIsBijective)
symmetricGroupInvIsLeft {A = A} S {sym fBij} = ExtensionallyEqual.eq λ {x} {!ans (sym fBij)!}
symmetricGroupInvIsLeft {A = A} S {sym {f = f} fBij} = ExtensionallyEqual.eq ans
where
ans : (f : A A) (bij : SetoidBijection S S f) f (symmetricGroupInv S (sym fBij) x) sym fBij
ans elt = {!!}
ans : {x : A} Setoid.__ S (f (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {x}
ans {x} | a , b = b
symmetricGroupInvIsRight : {a b : _} {A : Set a} (S : Setoid {a} {b} A) {x : SymmetryGroupElements S} Setoid.__ (symmetricSetoid S) (symmetricGroupOp x (symmetricGroupInv S x)) (sym setoidIdIsBijective)
symmetricGroupInvIsRight S {sym x} = {!!}
symmetricGroupInvIsRight {A = A} S {sym {f = f} fBij} = ExtensionallyEqual.eq ans
where
ans : {x : A} Setoid.__ S (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) (f x)) x
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {f x}
ans {x} | a , b = SetoidInjection.injective (SetoidBijection.inj fBij) b
symmetricGroup : {a b : _} {A : Set a} (S : Setoid {a} {b} A) Group (symmetricSetoid S) (symmetricGroupOp {A = A})
Group.wellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} (ExtensionallyEqual.eq m~x) (ExtensionallyEqual.eq n~y) = ExtensionallyEqual.eq λ {z} transitivity (applyEquality n m~x) n~y
Group.wellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} (ExtensionallyEqual.eq m~x) (ExtensionallyEqual.eq n~y) = ExtensionallyEqual.eq (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq A)) (SetoidBijection.wellDefined bijN m~x) n~y)
Group.identity (symmetricGroup A) = sym setoidIdIsBijective
Group.inverse (symmetricGroup S) = symmetricGroupInv S
Group.multAssoc (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = ExtensionallyEqual.eq λ {x} refl
Group.multIdentRight (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} refl
Group.multIdentLeft (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} refl
Group.multAssoc (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = ExtensionallyEqual.eq λ {x} Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq A))
Group.multIdentRight (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq A))
Group.multIdentLeft (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq A))
Group.invLeft (symmetricGroup S) {x} = symmetricGroupInvIsLeft S {x}
Group.invRight (symmetricGroup S) {x} = symmetricGroupInvIsRight S {x}
actionInducesHom : {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {_+_ : A A A} {G : Group S _+_} {B : Set b} {X : Setoid {b} {d} B} (GroupAction G X) SymmetryGroupElements X
actionInducesHom = {!!}
actionInducesHom : {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {_+_ : A A A} {G : Group S _+_} {B : Set b} {X : Setoid {b} {d} B} (GroupAction G X) A SymmetryGroupElements X
actionInducesHom {S = S} {_+_ = _+_} {G = G} {X = X} action f = sym {f = GroupAction.action action f} bij
where
bij : SetoidBijection X X (GroupAction.action action f)
SetoidInjection.wellDefined (SetoidBijection.inj bij) = GroupAction.actionWellDefined2 action
SetoidInjection.injective (SetoidBijection.inj bij) {x} {y} fx=fy = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) {x} {GroupAction.action action ((Group.inverse G f) + f) x} (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq X)) (GroupAction.identityAction action)) (GroupAction.actionWellDefined1 action (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) (Group.invLeft G)))) (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) {_} {GroupAction.action action (Group.inverse G f + f) y} (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (GroupAction.associativeAction action) (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (GroupAction.actionWellDefined2 action fx=fy) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq X)) (GroupAction.associativeAction action)))) (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (GroupAction.actionWellDefined1 action (Group.invLeft G)) (GroupAction.identityAction action)))
SetoidSurjection.wellDefined (SetoidBijection.surj bij) = GroupAction.actionWellDefined2 action
SetoidSurjection.surjective (SetoidBijection.surj bij) {x} = GroupAction.action action (Group.inverse G f) x , Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq X)) (GroupAction.associativeAction action)) (Transitive.transitive (Equivalence.transitiveEq (Setoid.eq X)) (GroupAction.actionWellDefined1 action (Group.invRight G)) (GroupAction.identityAction action))
actionInducesHomIsHom : {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {_+_ : A A A} {G : Group S _+_} {B : Set b} {X : Setoid {b} {d} B} (action : GroupAction G X) GroupHom G (symmetricGroup X) (actionInducesHom action)
GroupHom.groupHom (actionInducesHomIsHom action) = {!!}
GroupHom.wellDefined (actionInducesHomIsHom action) x=y = {!!}