Restructure towards ideals

This commit is contained in:
Smaug123
2019-11-20 21:20:03 +00:00
parent f2f4e867fc
commit b03a5279bc
22 changed files with 417 additions and 296 deletions

View File

@@ -16,12 +16,14 @@ open import Groups.Groups
open import Groups.Abelian.Lemmas open import Groups.Abelian.Lemmas
open import Groups.DirectSum.Definition open import Groups.DirectSum.Definition
open import Groups.QuotientGroup.Definition open import Groups.QuotientGroup.Definition
open import Groups.QuotientGroup.Lemmas
open import Groups.FiniteGroups.Definition open import Groups.FiniteGroups.Definition
open import Groups.Homomorphisms.Lemmas open import Groups.Homomorphisms.Lemmas
open import Groups.Homomorphisms.Examples
open import Groups.Isomorphisms.Lemmas open import Groups.Isomorphisms.Lemmas
open import Groups.FinitePermutations open import Groups.FinitePermutations
open import Groups.Lemmas open import Groups.Lemmas
open import Groups.Groups2 open import Groups.FirstIsomorphismTheorem
open import Groups.SymmetricGroups.Definition open import Groups.SymmetricGroups.Definition
open import Groups.Actions.Stabiliser open import Groups.Actions.Stabiliser
open import Groups.Actions.Orbit open import Groups.Actions.Orbit
@@ -30,6 +32,7 @@ open import Groups.ActionIsSymmetry
open import Groups.Cyclic.Definition open import Groups.Cyclic.Definition
open import Groups.Cyclic.DefinitionLemmas open import Groups.Cyclic.DefinitionLemmas
open import Groups.Polynomials.Examples open import Groups.Polynomials.Examples
open import Groups.Cosets
open import Fields.Fields open import Fields.Fields
open import Fields.Orders.Partial.Definition open import Fields.Orders.Partial.Definition
@@ -50,6 +53,7 @@ open import Rings.Polynomial.Ring
open import Rings.Polynomial.Evaluation open import Rings.Polynomial.Evaluation
open import Rings.Ideals.Definition open import Rings.Ideals.Definition
open import Rings.Isomorphisms.Definition open import Rings.Isomorphisms.Definition
open import Rings.Quotients.Definition
open import Setoids.Setoids open import Setoids.Setoids
open import Setoids.DirectSum open import Setoids.DirectSum

View File

@@ -11,7 +11,6 @@ open import Groups.Lemmas
open import Groups.Homomorphisms.Definition open import Groups.Homomorphisms.Definition
open import Groups.Groups open import Groups.Groups
open import Groups.SymmetricGroups.Definition open import Groups.SymmetricGroups.Definition
open import Groups.Groups2
open import Groups.Actions.Definition open import Groups.Actions.Definition
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations

View File

@@ -9,7 +9,6 @@ open import Sets.FinSet
open import Groups.Definition open import Groups.Definition
open import Groups.Lemmas open import Groups.Lemmas
open import Groups.Groups open import Groups.Groups
open import Groups.Groups2
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
module Groups.Actions.Definition where module Groups.Actions.Definition where

View File

@@ -10,7 +10,6 @@ open import Groups.Definition
open import Groups.Lemmas open import Groups.Lemmas
open import Groups.Groups open import Groups.Groups
open import Groups.Actions.Definition open import Groups.Actions.Definition
open import Groups.Groups2
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
module Groups.Actions.Orbit where module Groups.Actions.Orbit where

View File

@@ -12,7 +12,6 @@ open import Groups.Groups
open import Groups.Subgroups.Definition open import Groups.Subgroups.Definition
open import Groups.Homomorphisms.Definition open import Groups.Homomorphisms.Definition
open import Groups.Actions.Definition open import Groups.Actions.Definition
open import Groups.Groups2
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
open import Groups.Actions.Definition open import Groups.Actions.Definition
@@ -29,12 +28,13 @@ stabiliserWellDefined x {g} {h} g=h gx=x = transitive (actionWellDefined1 (Equiv
where where
open Equivalence eq open Equivalence eq
stabiliserSubgroup : (x : B) subgroup G (stabiliserWellDefined x) open Setoid T
_&_&_.one (stabiliserSubgroup x) gx=x hx=x = transitive associativeAction (transitive (actionWellDefined2 hx=x) gx=x) open Equivalence (Setoid.eq T)
stabiliserSubgroup : (x : B) subgroup G (stabiliserPred x)
_&&_.fst (stabiliserSubgroup x) = stabiliserWellDefined x
_&_&_.one (_&&_.snd (stabiliserSubgroup x)) gx=x hx=x = transitive associativeAction (transitive (actionWellDefined2 hx=x) gx=x)
_&_&_.two (_&&_.snd (stabiliserSubgroup x)) = identityAction
_&_&_.three (_&&_.snd (stabiliserSubgroup x)) {g} gx=x = transitive (transitive (transitive (actionWellDefined2 (symmetric gx=x)) (symmetric associativeAction)) (actionWellDefined1 (invLeft {g}))) identityAction
where where
open Equivalence eq
_&_&_.two (stabiliserSubgroup x) = identityAction
_&_&_.three (stabiliserSubgroup x) {g = g} gx=x = transitive (transitive (transitive (actionWellDefined2 (symmetric gx=x)) (symmetric associativeAction)) (actionWellDefined1 (invLeft {g}))) identityAction
where
open Equivalence eq
open Group G open Group G

View File

@@ -1,51 +1,50 @@
{-# OPTIONS --safe --warning=error #-} {-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae open import LogicalFormulae
open import Sets.EquivalenceRelations
open import Setoids.Setoids open import Setoids.Setoids
open import Setoids.Subset
open import Functions open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals open import Groups.Definition
open import Numbers.Integers open import Groups.Subgroups.Definition
open import Sets.FinSet open import Groups.Subgroups.Normal.Definition
open import Groups.GroupDefinition
open import Groups.Groups
module Groups.GroupCosets where module Groups.Cosets {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) {c : _} {pred : A Set c} (subgrp : subgroup G pred) where
data GroupCoset {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) : Set (a b c d) where
cosetOfElt : (x : A) GroupCoset subgrp
groupCosetSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) Setoid (GroupCoset subgrp) open Equivalence (Setoid.eq S)
Setoid.__ (groupCosetSetoid {A = A} {S = S} {_+A_ = _+A_} {G = G} subgrp) (cosetOfElt x) (cosetOfElt y) = Sg (A && A) (λ p x +A (_&&_.fst p) y +A (_&&_.snd p)) open import Groups.Lemmas G
where open _&_&_ (_&&_.snd subgrp)
open Setoid S open Group G
Reflexive.reflexive (Equivalence.reflexiveEq (Setoid.eq (groupCosetSetoid {S = S} {G = G} subgrp))) {cosetOfElt x} = (identity ,, identity) , transitive (multIdentRight) (symmetric multIdentRight)
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq (groupCosetSetoid {S = S} subgrp))) {cosetOfElt x} {cosetOfElt y} ((xMul ,, yMul) , pr) = (yMul ,, xMul) , Symmetric.symmetric (Equivalence.symmetricEq (Setoid.eq S)) pr
Transitive.transitive (Equivalence.transitiveEq (Setoid.eq (groupCosetSetoid {S = S} {_+A_ = _+A_} {G = G} subgrp))) {cosetOfElt x} {cosetOfElt y} {cosetOfElt z} ((xMul ,, yMul) , pr) ((yMul' ,, zMul) , pr2) = (((xMul +A (inverse yMul)) +A yMul') ,, zMul) , transitive (transitive multAssoc (wellDefined (transitive multAssoc (transitive (wellDefined pr reflexive) (transitive (symmetric multAssoc) (transitive (wellDefined reflexive invRight) multIdentRight)))) reflexive)) pr2
where
open Group G
open Setoid S
open Transitive (Equivalence.transitiveEq eq)
open Reflexive (Equivalence.reflexiveEq eq)
open Symmetric (Equivalence.symmetricEq eq)
groupCosetBijection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) SetoidsBiject T (groupCosetSetoid subgrp) private
SetoidsBiject.bij (groupCosetBijection {f = f} subgrp) x = cosetOfElt (f x) subs = _&&_.fst subgrp
SetoidInjection.wellDefined (SetoidBijection.inj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {G = G} {fHom = fHom} subgrp))) x~y = (identity ,, identity) , transitive multIdentRight (transitive (GroupHom.wellDefined fHom x~y) (symmetric multIdentRight))
where cosetSetoid : Setoid A
open Group G Setoid.__ cosetSetoid g h = pred ((Group.inverse G h) + g)
open Setoid S Equivalence.reflexive (Setoid.eq cosetSetoid) = subs (symmetric (Group.invLeft G)) two
open Transitive (Equivalence.transitiveEq eq) Equivalence.symmetric (Setoid.eq cosetSetoid) yx = subs (transitive invContravariant (+WellDefined reflexive invInv)) (three yx)
open Symmetric (Equivalence.symmetricEq eq) Equivalence.transitive (Setoid.eq cosetSetoid) yx zy = subs (transitive +Associative (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) reflexive)) (one zy yx)
SetoidInjection.injective (SetoidBijection.inj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {_+A_ = _+A_} {G = G} {f = f} subgrp))) {x} {y} ((xH ,, yH) , b) = {!!}
SetoidSurjection.wellDefined (SetoidBijection.surj (SetoidsBiject.bijIsBijective (groupCosetBijection {S = S} {G = G} {fHom = fHom} subgrp))) x~y = (identity ,, identity) , transitive multIdentRight (transitive (GroupHom.wellDefined fHom x~y) (symmetric multIdentRight)) cosetGroup : normalSubgroup G subgrp Group cosetSetoid _+_
where Group.+WellDefined (cosetGroup norm) {m} {n} {x} {y} m=x n=y = ans
open Group G where
open Setoid S t : pred (inverse y + n)
open Transitive (Equivalence.transitiveEq eq) t = n=y
open Symmetric (Equivalence.symmetricEq eq) u : pred (inverse x + m)
SetoidSurjection.surjective (SetoidBijection.surj (SetoidsBiject.bijIsBijective (groupCosetBijection subgrp))) {cosetOfElt x} = {!!} u = m=x
v : pred (m + inverse x)
v = subs (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) (norm u)
ans' : pred ((inverse y) + ((inverse x + m) + inverse (inverse y)))
ans' = norm u
ans'' : pred ((inverse y) + ((inverse x + m) + y))
ans'' = subs (+WellDefined reflexive (+WellDefined reflexive (invTwice y))) ans'
ans : pred (inverse (x + y) + (m + n))
ans = subs (transitive (transitive +Associative (transitive (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) +Associative)) reflexive) (symmetric +Associative))) (symmetric (+WellDefined invContravariant reflexive))) (one ans'' t)
Group.0G (cosetGroup norm) = 0G
Group.inverse (cosetGroup norm) = inverse
Group.+Associative (cosetGroup norm) {a} {b} {c} = subs (symmetric (transitive (+WellDefined (inverseWellDefined (symmetric +Associative)) reflexive) (invLeft {a + (b + c)}))) two
Group.identRight (cosetGroup norm) = subs (symmetric (transitive +Associative (transitive (+WellDefined invLeft reflexive) identRight))) two
Group.identLeft (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive identLeft) invLeft)) two
Group.invLeft (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive invLeft) invLeft)) two
Group.invRight (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive invRight) invLeft)) two

View File

@@ -0,0 +1,49 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Groups.Groups
open import Groups.Definition
open import Orders
open import Numbers.Integers.Integers
open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
open import Sets.EquivalenceRelations
open import Numbers.Naturals.Naturals
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Definition
open import Groups.Subgroups.Definition
open import Groups.Subgroups.Normal.Definition
open import Groups.Subgroups.Normal.Lemmas
open import Groups.Lemmas
open import Groups.Abelian.Definition
open import Groups.QuotientGroup.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.FirstIsomorphismTheorem {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) where
open import Groups.Homomorphisms.Image fHom
open import Groups.Homomorphisms.Kernel fHom
open import Groups.Cosets G groupKernelIsSubgroup
open Setoid T
open Equivalence eq
open import Setoids.Subset T
groupFirstIsomorphismTheoremWellDefined : {x y : A} groupKernelPred (Group.inverse G y +G x) Setoid.__ (subsetSetoid imageGroupSubset) (f x , (x , reflexive)) (f y , (y , reflexive))
groupFirstIsomorphismTheoremWellDefined {x} {y} pr = transitive (symmetric (invTwice H _)) (transitive (symmetric u) (invTwice H _))
where
t : Setoid.__ T (Group.inverse H (f y) +H (f x)) (Group.0G H)
t = transitive (transitive (Group.+WellDefined H (symmetric (homRespectsInverse fHom)) reflexive) (symmetric (GroupHom.groupHom fHom))) pr
u : Setoid.__ T (Group.inverse H (Group.inverse H (f y))) (Group.inverse H (Group.inverse H (f x)))
u = inverseWellDefined H (transferToRight' H t)
groupFirstIsomorphismTheorem : GroupsIsomorphic (cosetGroup groupKernelIsNormalSubgroup) (subgroupIsGroup H imageGroupSubset imageGroupSubgroup)
GroupsIsomorphic.isomorphism groupFirstIsomorphismTheorem x = f x , (x , reflexive)
GroupHom.groupHom (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem)) {x} {y} = GroupHom.groupHom fHom
GroupHom.wellDefined (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem)) {x} {y} = groupFirstIsomorphismTheoremWellDefined {x} {y}
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) = groupFirstIsomorphismTheoremWellDefined
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H reflexive fx=fy) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom))))
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) = groupFirstIsomorphismTheoremWellDefined
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) {b , (a , fa=b)} = a , fa=b

View File

@@ -73,6 +73,20 @@ quotientGroupSetoid {A = A} {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H
k : f (m ·A ((Group.0G G) ·A Group.inverse G o)) 0G k : f (m ·A ((Group.0G G) ·A Group.inverse G o)) 0G
k = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.reflexive (Setoid.eq S))))) j k = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.reflexive (Setoid.eq S))))) j
quotientGroupLemma : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) {x y : A} Setoid.__ T (underf x) (underf y) Setoid.__ (quotientGroupSetoid G f) x y
quotientGroupLemma {S = S} {T = T} G {H = H} fHom {x} {y} fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (homRespectsInverse fHom)) (transferToRight'' H fx=fy))
where
open Group G
open Setoid T
open Equivalence eq
quotientGroupLemma' : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) {x y : A} Setoid.__ (quotientGroupSetoid G f) x y Setoid.__ T (underf x) (underf y)
quotientGroupLemma' {S = S} {T = T} G {H = H} f fx=fy = transferToRight H (transitive (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (symmetric (homRespectsInverse f))) (symmetric (GroupHom.groupHom f))) fx=fy)
where
open Group G
open Setoid T
open Equivalence eq
{- {-
quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → A → A quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_}{f : A → B} → (fHom : GroupHom G H f) → A → A
quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!} quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!}

View File

@@ -1,54 +0,0 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Groups.Groups
open import Groups.Definition
open import Orders
open import Numbers.Integers.Integers
open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
open import Sets.EquivalenceRelations
open import Numbers.Naturals.Naturals
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Homomorphisms.Image
open import Groups.Isomorphisms.Definition
open import Groups.Subgroups.Definition
open import Groups.Subgroups.Normal.Definition
open import Groups.Subgroups.Normal.Lemmas
open import Groups.Lemmas
open import Groups.Abelian.Definition
open import Groups.QuotientGroup.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.Groups2 where
groupFirstIsomorphismIso : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom A
groupFirstIsomorphismIso fHom (ofElt a) = a
groupFirstIsomorphismIsoHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHom (imageGroup fHom) (quotientGroup G fHom) (groupFirstIsomorphismIso fHom)
GroupHom.groupHom (groupFirstIsomorphismIsoHom {G = G} fHom) {ofElt a} {ofElt b} = Equivalence.reflexive (Setoid.eq (quotientGroupSetoid G fHom))
GroupHom.wellDefined (groupFirstIsomorphismIsoHom {T = T} {_+G_ = _+G_} {G = G} {H = H} {f = f} fHom) {ofElt a} {ofElt b} pr = ans
where
open Setoid T
open Equivalence (Setoid.eq T)
ans : f (a +G Group.inverse G b) Group.0G H
ans = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H pr reflexive) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invRight G)) (imageOfIdentityIsIdentity fHom))))
groupFirstIsomorphismTheorem' : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupIso (imageGroup fHom) (quotientGroup G fHom) (groupFirstIsomorphismIso fHom)
GroupIso.groupHom (groupFirstIsomorphismTheorem' fHom) = groupFirstIsomorphismIsoHom fHom
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (groupFirstIsomorphismTheorem' fHom))) {x} {y} x~y = GroupHom.wellDefined (groupFirstIsomorphismIsoHom fHom) {x} {y} x~y
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (groupFirstIsomorphismTheorem' {T = T} {H = H} {f = f} fHom))) {ofElt a} {ofElt b} pr = need
where
open Setoid T
open Equivalence eq
need : f a f b
need = transferToRight H (transitive (transitive (Group.+WellDefined H reflexive (symmetric (homRespectsInverse fHom))) (symmetric (GroupHom.groupHom fHom))) pr)
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (groupFirstIsomorphismTheorem' fHom))) {x} {y} x~y = GroupHom.wellDefined (groupFirstIsomorphismIsoHom fHom) {x} {y} x~y
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (groupFirstIsomorphismTheorem' {G = G} fHom))) {a} = ofElt a , Equivalence.reflexive (Setoid.eq (quotientGroupSetoid G fHom))
groupFirstIsomorphismTheorem : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupsIsomorphic (imageGroup fHom) (quotientGroup G fHom)
groupFirstIsomorphismTheorem fHom = record { isomorphism = groupFirstIsomorphismIso fHom ; proof = groupFirstIsomorphismTheorem' fHom }

View File

@@ -0,0 +1,18 @@
{-# 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.Naturals
open import Sets.FinSet
open import Groups.Definition
open import Sets.EquivalenceRelations
open import Groups.Homomorphisms.Definition
open import Groups.Lemmas
module Groups.Homomorphisms.Examples where
identityHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) GroupHom G G id
GroupHom.groupHom (identityHom {S = S} G) = Equivalence.reflexive (Setoid.eq S)
GroupHom.wellDefined (identityHom G) = id

View File

@@ -5,6 +5,7 @@ open import Groups.Definition
open import Orders open import Orders
open import Numbers.Integers.Integers open import Numbers.Integers.Integers
open import Setoids.Setoids open import Setoids.Setoids
open import Setoids.Subset
open import LogicalFormulae open import LogicalFormulae
open import Sets.FinSet open import Sets.FinSet
open import Functions open import Functions
@@ -20,45 +21,25 @@ open import Groups.QuotientGroup.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.Homomorphisms.Image where module Groups.Homomorphisms.Image {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) where
data GroupHomImageElement {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) : Set (a b c d) where imageGroupPred : B Set (a d)
ofElt : (x : A) GroupHomImageElement fHom imageGroupPred b = Sg A (λ a Setoid.__ T (f a) b)
imageGroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) Setoid (GroupHomImageElement fHom) imageGroupSubset : subset T imageGroupPred
(imageGroupSetoid {T = T} {f = f} fHom Setoid. ofElt b1) (ofElt b2) = Setoid.__ T (f b1) (f b2) imageGroupSubset {x} {y} x=y (a , fa=x) = a , transitive fa=x x=y
Equivalence.reflexive (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} = Equivalence.reflexive (Setoid.eq T)
Equivalence.symmetric (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} {ofElt b2} = Equivalence.symmetric (Setoid.eq T)
Equivalence.transitive (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} {ofElt b2} {ofElt b3} = Equivalence.transitive (Setoid.eq T)
imageGroupOp : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom GroupHomImageElement fHom GroupHomImageElement fHom
imageGroupOp {_+A_ = _+A_} fHom (ofElt a) (ofElt b) = ofElt (a +A b)
imageGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (fHom : GroupHom G H f) Group (imageGroupSetoid fHom) (imageGroupOp fHom)
Group.+WellDefined (imageGroup {T = T} {_+A_ = _+A_} {H = H} {f = f} fHom) {ofElt x} {ofElt y} {ofElt a} {ofElt b} x~a y~b = ans
where where
open Setoid T open Setoid T
open Equivalence eq open Equivalence eq
ans : f (x +A y) f (a +A b)
ans = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H x~a y~b) (symmetric (GroupHom.groupHom fHom)))
Group.0G (imageGroup {G = G} fHom) = ofElt (Group.0G G)
Group.inverse (imageGroup {G = G} fHom) (ofElt a) = ofElt (Group.inverse G a)
Group.+Associative (imageGroup {G = G} fHom) {ofElt a} {ofElt b} {ofElt c} = GroupHom.wellDefined fHom (Group.+Associative G)
Group.identRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.identRight G)
Group.identLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.identLeft G)
Group.invLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invLeft G)
Group.invRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invRight G)
groupHomImageInclusion : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHomImageElement fHom B imageGroupSubgroup : subgroup H imageGroupPred
groupHomImageInclusion {f = f} fHom (ofElt x) = f x _&&_.fst imageGroupSubgroup = imageGroupSubset
_&_&_.one (_&&_.snd imageGroupSubgroup) {x} {y} (a , fa=x) (b , fb=y) = (a +A b) , transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fa=x fb=y)
groupHomImageIncludes : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) GroupHom (imageGroup fHom) H (groupHomImageInclusion fHom)
GroupHom.groupHom (groupHomImageIncludes fHom) {ofElt x} {ofElt y} = GroupHom.groupHom fHom
GroupHom.wellDefined (groupHomImageIncludes fHom) {ofElt x} {ofElt y} x~y = x~y
groupHomImageInjects : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) SetoidInjection (imageGroupSetoid fHom) T (groupHomImageInclusion fHom)
groupHomImageInjects {S = S} {T} {_+G_} {_+H_} {G} {H} {f} fHom = record { wellDefined = λ {x} {y} GroupHom.wellDefined (groupHomImageIncludes fHom) {x} {y} ; injective = λ {x} {y} inj {x} {y} }
where where
inj : {x y : GroupHomImageElement fHom} (Setoid.__ T (groupHomImageInclusion fHom x) (groupHomImageInclusion fHom y)) Setoid.__ (imageGroupSetoid fHom) x y open Setoid T
inj {ofElt x} {ofElt y} x~y = x~y open Equivalence eq
_&_&_.two (_&&_.snd imageGroupSubgroup) = Group.0G G , imageOfIdentityIsIdentity fHom
_&_&_.three (_&&_.snd imageGroupSubgroup) {x} (a , fa=x) = Group.inverse G a , transitive (homRespectsInverse fHom) (inverseWellDefined H fa=x)
where
open Setoid T
open Equivalence eq

View File

@@ -0,0 +1,44 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Groups.Groups
open import Groups.Definition
open import Orders
open import Numbers.Integers.Integers
open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
open import Sets.EquivalenceRelations
open import Numbers.Naturals.Naturals
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Groups.Isomorphisms.Definition
open import Groups.Subgroups.Definition
open import Groups.Subgroups.Normal.Definition
open import Groups.Subgroups.Normal.Lemmas
open import Groups.Lemmas
open import Groups.Abelian.Definition
open import Groups.QuotientGroup.Definition
open import Groups.Cosets
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.Homomorphisms.Kernel {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A A A} {_+H_ : B B B} {G : Group S _+G_} {H : Group T _+H_} {f : A B} (fHom : GroupHom G H f) where
open Setoid T
open Equivalence eq
groupKernelPred : A Set d
groupKernelPred a = Setoid.__ T (f a) (Group.0G H)
groupKernelPredWd : {x y : A} (Setoid.__ S x y) groupKernelPred x groupKernelPred y
groupKernelPredWd x=y fx=0 = transitive (GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) x=y)) fx=0
groupKernelIsSubgroup : subgroup G groupKernelPred
_&_&_.one (_&&_.snd groupKernelIsSubgroup) fg=0 fh=0 = transitive (transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fg=0 fh=0)) (Group.identLeft H)
_&_&_.two (_&&_.snd groupKernelIsSubgroup) = imageOfIdentityIsIdentity fHom
_&_&_.three (_&&_.snd groupKernelIsSubgroup) fg=0 = transitive (homRespectsInverse fHom) (transitive (inverseWellDefined H fg=0) (invIdent H))
_&&_.fst groupKernelIsSubgroup = groupKernelPredWd
groupKernelIsNormalSubgroup : normalSubgroup G groupKernelIsSubgroup
groupKernelIsNormalSubgroup {g} fk=0 = transitive (transitive (transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H reflexive (transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H fk=0 reflexive) (Group.identLeft H)))) (symmetric (GroupHom.groupHom fHom)))) (GroupHom.wellDefined fHom (Group.invRight G {g}))) (imageOfIdentityIsIdentity fHom)

View File

@@ -11,36 +11,38 @@ open import Sets.EquivalenceRelations
open import Groups.Homomorphisms.Definition open import Groups.Homomorphisms.Definition
open import Groups.Lemmas open import Groups.Lemmas
module Groups.Homomorphisms.Lemmas where module Groups.Homomorphisms.Lemmas {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {B : Set b} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : A B} (hom : GroupHom G H f) where
imageOfIdentityIsIdentity : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {B : Set n} {T : Setoid {n} {p} B} {_·A_ : A A A} {_·B_ : B B B} {G : Group S _·A_} {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) Setoid.__ T (f (Group.0G G)) (Group.0G H) imageOfIdentityIsIdentity : Setoid.__ T (f (Group.0G G)) (Group.0G H)
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t imageOfIdentityIsIdentity = Equivalence.symmetric (Setoid.eq T) t
where where
open Group H open Group H
open Setoid T open Setoid T
id2 : Setoid.__ S (Group.0G G) ((Group.0G G) ·A (Group.0G G)) id2 : Setoid.__ S (Group.0G G) ((Group.0G G) +A (Group.0G G))
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G) id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
r : f (Group.0G G) f (Group.0G G) ·B f (Group.0G G) r : f (Group.0G G) f (Group.0G G) +B f (Group.0G G)
s : 0G ·B f (Group.0G G) f (Group.0G G) ·B f (Group.0G G) s : 0G +B f (Group.0G G) f (Group.0G G) +B f (Group.0G G)
t : 0G f (Group.0G G) t : 0G f (Group.0G G)
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
s = Equivalence.transitive (Setoid.eq T) identLeft r s = Equivalence.transitive (Setoid.eq T) identLeft r
r = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom) r = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom)
groupHomsCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A A A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B B B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C C C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A B} {g : B C} (fHom : GroupHom G H f) (gHom : GroupHom H I g) GroupHom G I (g f) groupHomsCompose : {o t : _} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C C C} {I : Group U _+C_} {g : B C} (gHom : GroupHom H I g) GroupHom G I (g f)
GroupHom.wellDefined (groupHomsCompose {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined fHom pr) GroupHom.wellDefined (groupHomsCompose {I} {f} gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined hom pr)
GroupHom.groupHom (groupHomsCompose {S = S} {_+A_ = _·A_} {T = T} {U = U} {_+C_ = _·C_} {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} = answer GroupHom.groupHom (groupHomsCompose {U = U} {_+C_ = _·C_} {I} {g} gHom) {x} {y} = answer
where where
open Group I open Group I
answer : (Setoid.__ U) ((g f) (x ·A y)) ((g f) x ·C (g f) y) answer : (Setoid.__ U) ((g f) (x +A y)) ((g f) x ·C (g f) y)
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y}) answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom hom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
homRespectsInverse : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} {G : Group S _·A_} {H : Group T _·B_} {underF : A B} (f : GroupHom G H underF) {x : A} Setoid.__ T (underF (Group.inverse G x)) (Group.inverse H (underF x)) homRespectsInverse : {x : A} Setoid.__ T (f (Group.inverse G x)) (Group.inverse H (f x))
homRespectsInverse {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {underF = f} fHom {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom))) 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 where
open Setoid T open Setoid T
open Equivalence eq open Equivalence eq
identityHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A A A} (G : Group S _+A_) GroupHom G G id zeroImpliesInverseZero : {x : A} Setoid.__ T (f x) (Group.0G H) Setoid.__ T (f (Group.inverse G x)) (Group.0G H)
GroupHom.groupHom (identityHom {S = S} G) = Equivalence.reflexive (Setoid.eq S) zeroImpliesInverseZero {x} fx=0 = transitive homRespectsInverse (transitive (inverseWellDefined H fx=0) (invIdent H))
GroupHom.wellDefined (identityHom G) = id where
open Setoid T
open Equivalence eq

View File

@@ -14,11 +14,13 @@ open import Fields.FieldOfFractions.Setoid
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
open import Groups.Lemmas open import Groups.Lemmas
open import Groups.Homomorphisms.Lemmas open import Groups.Homomorphisms.Lemmas
open import Groups.Subgroups.Definition
open import Groups.Subgroups.Normal.Definition
module Groups.QuotientGroup.Definition where module Groups.QuotientGroup.Definition where
quotientGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) Group (quotientGroupSetoid G f) _·A_ quotientGroupByHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {underf : A B} (f : GroupHom G H underf) Group (quotientGroupSetoid G f) _·A_
Group.+WellDefined (quotientGroup {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans Group.+WellDefined (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans
where where
open Setoid T open Setoid T
open Equivalence (Setoid.eq T) open Equivalence (Setoid.eq T)
@@ -40,15 +42,15 @@ Group.+WellDefined (quotientGroup {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_
p8 = x~m p8 = x~m
ans : f ((x ·A y) ·A (Group.inverse G (m ·A n))) Group.0G H ans : f ((x ·A y) ·A (Group.inverse G (m ·A n))) Group.0G H
ans = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 (transitive p6 (transitive p7 p8)))))) ans = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 (transitive p6 (transitive p7 p8))))))
Group.0G (quotientGroup G fHom) = Group.0G G Group.0G (quotientGroupByHom G fHom) = Group.0G G
Group.inverse (quotientGroup G fHom) = Group.inverse G Group.inverse (quotientGroupByHom G fHom) = Group.inverse G
Group.+Associative (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans Group.+Associative (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans
where where
open Setoid T open Setoid T
open Equivalence (Setoid.eq T) open Equivalence (Setoid.eq T)
ans : f ((a ·A (b ·A c)) ·A (Group.inverse G ((a ·A b) ·A c))) Group.0G H ans : f ((a ·A (b ·A c)) ·A (Group.inverse G ((a ·A b) ·A c))) Group.0G H
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.+Associative G))) (imageOfIdentityIsIdentity fHom) ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.+Associative G))) (imageOfIdentityIsIdentity fHom)
Group.identRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans Group.identRight (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
where where
open Group G open Group G
open Setoid T open Setoid T
@@ -56,7 +58,7 @@ Group.identRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {under
transitiveG = Equivalence.transitive (Setoid.eq S) transitiveG = Equivalence.transitive (Setoid.eq S)
ans : f ((a ·A 0G) ·A inverse a) Group.0G H ans : f ((a ·A 0G) ·A inverse a) Group.0G H
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identRight G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom) ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identRight G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
Group.identLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans Group.identLeft (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
where where
open Group G open Group G
open Setoid T open Setoid T
@@ -64,14 +66,14 @@ Group.identLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf
transitiveG = Equivalence.transitive (Setoid.eq S) transitiveG = Equivalence.transitive (Setoid.eq S)
ans : f ((0G ·A a) ·A (inverse a)) Group.0G H ans : f ((0G ·A a) ·A (inverse a)) Group.0G H
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identLeft G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom) ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identLeft G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
Group.invLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans Group.invLeft (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
where where
open Group G open Group G
open Setoid T open Setoid T
open Equivalence eq open Equivalence eq
ans : f ((inverse x ·A x) ·A (inverse 0G)) (Group.0G H) ans : f ((inverse x ·A x) ·A (inverse 0G)) (Group.0G H)
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdent G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom) ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdent G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
Group.invRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans Group.invRight (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
where where
open Group G open Group G
open Setoid T open Setoid T

View File

@@ -0,0 +1,21 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import Functions
open import Sets.FinSet
open import LogicalFormulae
open import Groups.Definition
open import Groups.Groups
open import Groups.FiniteGroups.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Abelian.Definition
open import Setoids.Setoids
open import Fields.FieldOfFractions.Setoid
open import Sets.EquivalenceRelations
open import Groups.Lemmas
open import Groups.QuotientGroup.Definition
module Groups.QuotientGroup.Lemmas {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} (G : Group S _+A_) (H : Group T _+B_) {f : A B} (fHom : GroupHom G H f) where
projectionMapIsGroupHom : GroupHom G (quotientGroupByHom G fHom) id
GroupHom.groupHom projectionMapIsGroupHom {x} {y} = quotientGroupLemma G fHom (Equivalence.reflexive (Setoid.eq T))
GroupHom.wellDefined projectionMapIsGroupHom x=y = quotientGroupLemma G fHom (GroupHom.wellDefined fHom x=y)

View File

@@ -11,7 +11,21 @@ open import Groups.Homomorphisms.Definition
module Groups.Subgroups.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where module Groups.Subgroups.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where
open import Setoids.Subset S
open Group G open Group G
subgroup : {c : _} {pred : A Set c} (wd : {x y : A} (Setoid.__ S x y) (pred x pred y)) Set (a c) subgroup : {c : _} (pred : A Set c) Set (a b c)
subgroup {pred = pred} wd = ({g h : A} (pred g) (pred h) pred (g + h)) & pred 0G & ({g : A} (pred g) (pred (inverse g))) subgroup pred = subset pred && (({g h : A} (pred g) (pred h) pred (g + h)) & pred 0G & ({g : A} (pred g) (pred (inverse g))))
subgroupOp : {c : _} {pred : A Set c} (s : subgroup pred) Sg A pred Sg A pred Sg A pred
subgroupOp {pred = pred} (_ ,, record { one = one ; two = two ; three = three }) (a , prA) (b , prB) = (a + b) , one prA prB
subgroupIsGroup : {c : _} {pred : A Set c} (subs : subset pred) (s : subgroup pred) Group (subsetSetoid subs) (subgroupOp s)
Group.+WellDefined (subgroupIsGroup _ s) {m , prM} {n , prN} {x , prX} {y , prY} m=x n=y = +WellDefined m=x n=y
Group.0G (subgroupIsGroup _ (_ ,, record { two = two })) = 0G , two
Group.inverse (subgroupIsGroup _ (_ ,, record { three = three })) (a , prA) = (inverse a) , three prA
Group.+Associative (subgroupIsGroup _ s) {a , prA} {b , prB} {c , prC} = +Associative
Group.identRight (subgroupIsGroup _ s) {a , prA} = identRight
Group.identLeft (subgroupIsGroup _ s) {a , prA} = identLeft
Group.invLeft (subgroupIsGroup _ s) {a , prA} = invLeft
Group.invRight (subgroupIsGroup _ s) {a , prA} = invRight

View File

@@ -16,11 +16,10 @@ open import Groups.Isomorphisms.Definition
open import Groups.Subgroups.Definition open import Groups.Subgroups.Definition
open import Groups.Lemmas open import Groups.Lemmas
open import Groups.Abelian.Definition open import Groups.Abelian.Definition
open import Groups.QuotientGroup.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Groups.Subgroups.Normal.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where module Groups.Subgroups.Normal.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) where
normalSubgroup : {c : _} {pred : A Set c} (wd : {x y : A} (Setoid.__ S x y) (pred x pred y)) (sub : subgroup G wd) Set (a c) normalSubgroup : {c : _} {pred : A Set c} (sub : subgroup G pred) Set (a c)
normalSubgroup {pred = pred} wd sub = {g k : A} pred k pred (g + (k + Group.inverse G g)) normalSubgroup {pred = pred} sub = {g k : A} pred k pred (g + (k + Group.inverse G g))

View File

@@ -64,12 +64,13 @@ groupKernelGroupPred {T = T} G {H = H} {f = f} hom a = Setoid.__ T (f a) (Gro
groupKernelGroupPredWd : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) {x y : A} (Setoid.__ S x y) (groupKernelGroupPred G hom x groupKernelGroupPred G hom y) groupKernelGroupPredWd : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) {x y : A} (Setoid.__ S x y) (groupKernelGroupPred G hom x groupKernelGroupPred G hom y)
groupKernelGroupPredWd {S = S} {T = T} G hom {x} {y} x=y fx=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom (Equivalence.symmetric (Setoid.eq S) x=y)) fx=0 groupKernelGroupPredWd {S = S} {T = T} G hom {x} {y} x=y fx=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom (Equivalence.symmetric (Setoid.eq S) x=y)) fx=0
groupKernelGroupIsSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) subgroup G (groupKernelGroupPredWd G hom) groupKernelGroupIsSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) subgroup G (groupKernelGroupPred G hom)
_&_&_.one (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom) g=0 h=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.groupHom hom) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H g=0 h=0) (Group.identLeft H)) _&_&_.one (_&&_.snd (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom)) g=0 h=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.groupHom hom) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H g=0 h=0) (Group.identLeft H))
_&_&_.two (groupKernelGroupIsSubgroup G hom) = imageOfIdentityIsIdentity hom _&_&_.two (_&&_.snd (groupKernelGroupIsSubgroup G hom)) = imageOfIdentityIsIdentity hom
_&_&_.three (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom) g=0 = Equivalence.transitive (Setoid.eq T) (homRespectsInverse hom) (Equivalence.transitive (Setoid.eq T) (inverseWellDefined H g=0) (invIdent H)) _&_&_.three (_&&_.snd (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom)) g=0 = Equivalence.transitive (Setoid.eq T) (homRespectsInverse hom) (Equivalence.transitive (Setoid.eq T) (inverseWellDefined H g=0) (invIdent H))
_&&_.fst (groupKernelGroupIsSubgroup G hom) = groupKernelGroupPredWd G hom
groupKernelGroupIsNormalSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) normalSubgroup G (groupKernelGroupPredWd G hom) (groupKernelGroupIsSubgroup G hom) groupKernelGroupIsNormalSubgroup : {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A A A} {_·B_ : B B B} (G : Group S _·A_) {H : Group T _·B_} {f : A B} (hom : GroupHom G H f) normalSubgroup G (groupKernelGroupIsSubgroup G hom)
groupKernelGroupIsNormalSubgroup {T = T} G {H = H} hom k=0 = transitive groupHom (transitive (+WellDefined reflexive groupHom) (transitive (+WellDefined reflexive (transitive (+WellDefined k=0 reflexive) identLeft)) (transitive (symmetric groupHom) (transitive (wellDefined (Group.invRight G)) (imageOfIdentityIsIdentity hom))))) groupKernelGroupIsNormalSubgroup {T = T} G {H = H} hom k=0 = transitive groupHom (transitive (+WellDefined reflexive groupHom) (transitive (+WellDefined reflexive (transitive (+WellDefined k=0 reflexive) identLeft)) (transitive (symmetric groupHom) (transitive (wellDefined (Group.invRight G)) (imageOfIdentityIsIdentity hom)))))
where where
open Setoid T open Setoid T
@@ -77,8 +78,8 @@ groupKernelGroupIsNormalSubgroup {T = T} G {H = H} hom k=0 = transitive groupHom
open Equivalence eq open Equivalence eq
open GroupHom hom open GroupHom hom
abelianGroupSubgroupIsNormal : {a b c : _} {A : Set a} {B : Set b} {S : Setoid {a} {b} A} {_+_ : A A A} {G : Group S _+_} {pred : A Set c} (predWd : {x y : A} (Setoid.__ S x y) (pred x pred y)) (s : subgroup G predWd) AbelianGroup G normalSubgroup G predWd s abelianGroupSubgroupIsNormal : {a b c : _} {A : Set a} {B : Set b} {S : Setoid {a} {b} A} {_+_ : A A A} {G : Group S _+_} {pred : A Set c} (s : subgroup G pred) AbelianGroup G normalSubgroup G s
abelianGroupSubgroupIsNormal {S = S} {_+_ = _+_} {G = G} predWd record { one = respectsPlus ; two = respectsId ; three = respectsInv } abelian {k} {l} prK = predWd (transitive (transitive (transitive (symmetric identLeft) (+WellDefined (symmetric invRight) reflexive)) (symmetric +Associative)) (+WellDefined reflexive commutative)) prK abelianGroupSubgroupIsNormal {S = S} {_+_ = _+_} {G = G} (predWd ,, record { one = respectsPlus ; two = respectsId ; three = respectsInv }) abelian {k} {l} prK = predWd (transitive (transitive (transitive (symmetric identLeft) (+WellDefined (symmetric invRight) reflexive)) (symmetric +Associative)) (+WellDefined reflexive commutative)) prK
where where
open Group G open Group G
open AbelianGroup abelian open AbelianGroup abelian

View File

@@ -13,7 +13,6 @@ open import Groups.Subgroups.Definition
open import Groups.Homomorphisms.Definition open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas open import Groups.Homomorphisms.Lemmas
open import Groups.Actions.Definition open import Groups.Actions.Definition
open import Groups.Groups2
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
module Groups.SymmetricGroups.Lemmas where module Groups.SymmetricGroups.Lemmas where

View File

@@ -17,11 +17,13 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Ideals.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where module Rings.Ideals.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open import Groups.Subgroups.Definition (Ring.additiveGroup R)
ringKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) Set (a d) ringKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) Set (a d)
ringKernel {T = T} R2 {f} fHom = Sg A (λ a Setoid.__ T (f a) (Ring.0R R2)) ringKernel {T = T} R2 {f} fHom = Sg A (λ a Setoid.__ T (f a) (Ring.0R R2))
ideal : {c : _} {pred : A Set c} (wd : {x y : A} (Setoid.__ S x y) (pred x pred y)) Set (a c) ideal : {c : _} (pred : A Set c) Set (a b c)
ideal {pred = pred} wd = (pred (Ring.0R R)) & ({x y : A} pred x pred y pred (x + y)) & ({x : A} {y : A} pred x pred (x * y)) ideal pred = subgroup pred && ({x : A} {y : A} pred x pred (x * y))
idealPredForKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) A Set d idealPredForKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) A Set d
idealPredForKernel {T = T} R2 {f} fHom a = Setoid.__ T (f a) (Ring.0R R2) idealPredForKernel {T = T} R2 {f} fHom a = Setoid.__ T (f a) (Ring.0R R2)
@@ -29,17 +31,24 @@ idealPredForKernel {T = T} R2 {f} fHom a = Setoid.__ T (f a) (Ring.0R R2)
idealPredForKernelWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) {x y : A} (Setoid.__ S x y) (idealPredForKernel R2 fHom x idealPredForKernel R2 fHom y) idealPredForKernelWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) {x y : A} (Setoid.__ S x y) (idealPredForKernel R2 fHom x idealPredForKernel R2 fHom y)
idealPredForKernelWellDefined {T = T} R2 {f} fHom a x=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (RingHom.groupHom fHom) (Equivalence.symmetric (Setoid.eq S) a)) x=0 idealPredForKernelWellDefined {T = T} R2 {f} fHom a x=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (RingHom.groupHom fHom) (Equivalence.symmetric (Setoid.eq S) a)) x=0
kernelIdealIsIdeal : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} {R2 : Ring T _+2_ _*2_} {f : A C} (fHom : RingHom R R2 f) ideal {pred = idealPredForKernel R2 fHom} (idealPredForKernelWellDefined R2 fHom) kernelIdealIsIdeal : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} {R2 : Ring T _+2_ _*2_} {f : A C} (fHom : RingHom R R2 f) ideal (idealPredForKernel R2 fHom)
_&_&_.one (kernelIdealIsIdeal fHom) = imageOfIdentityIsIdentity (RingHom.groupHom fHom) _&&_.fst (_&&_.fst (kernelIdealIsIdeal {R2 = R2} fHom)) = idealPredForKernelWellDefined R2 fHom
_&_&_.two (kernelIdealIsIdeal {T = T} {R2 = R2} fHom) fx=0 fy=0 = transitive (transitive (GroupHom.groupHom (RingHom.groupHom fHom)) (+WellDefined fx=0 fy=0)) identLeft _&_&_.one (_&&_.snd (_&&_.fst (kernelIdealIsIdeal {T = T} {R2 = R2} fHom))) {x} {y} fx=0 fy=0 = transitive (transitive (GroupHom.groupHom (RingHom.groupHom fHom)) (+WellDefined fx=0 fy=0)) identLeft
where where
open Ring R2 open Ring R2
open Group (Ring.additiveGroup R2) open Group (Ring.additiveGroup R2)
open Setoid T open Setoid T
open Equivalence eq open Equivalence eq
_&_&_.three (kernelIdealIsIdeal {T = T} {R2 = R2} fHom) fx=0 = transitive (RingHom.ringHom fHom) (transitive (Ring.*WellDefined R2 fx=0 reflexive) (transitive (Ring.*Commutative R2) (Ring.timesZero R2))) _&_&_.two (_&&_.snd (_&&_.fst (kernelIdealIsIdeal fHom))) = imageOfIdentityIsIdentity (RingHom.groupHom fHom)
_&_&_.three (_&&_.snd (_&&_.fst (kernelIdealIsIdeal {T = T} {R2 = R2} fHom))) {x} fx=0 = zeroImpliesInverseZero (RingHom.groupHom fHom) fx=0
where where
open Ring R2 open Ring R2
open Group (Ring.additiveGroup R2) open Group (Ring.additiveGroup R2)
open Setoid T open Setoid T
open Equivalence eq open Equivalence eq
_&&_.snd (kernelIdealIsIdeal {T = T} {R2 = R2} {f = f} fHom) {x} {y} fx=0 = transitive (RingHom.ringHom fHom) (transitive (Ring.*WellDefined R2 fx=0 reflexive) (transitive (Ring.*Commutative R2) (Ring.timesZero R2 {f y})))
where
open Setoid T
open Equivalence eq
-- TODO : define the quotient by an ideal; note that the result is a ring

View File

@@ -6,115 +6,116 @@ open import LogicalFormulae
open import Sets.EquivalenceRelations open import Sets.EquivalenceRelations
module Setoids.Setoids where module Setoids.Setoids where
record Setoid {a} {b} (A : Set a) : Set (a lsuc b) where
infix 1 __
field
__ : A A Set b
eq : Equivalence __
open Equivalence eq record Setoid {a} {b} (A : Set a) : Set (a lsuc b) where
infix 1 __
field
__ : A A Set b
eq : Equivalence __
~refl : {r : A} (r r) open Equivalence eq
~refl {r} = reflexive {r}
record Quotient {a} {b} {c} {A : Set a} {image : Set b} (S : Setoid {a} {c} A) : Set (a b c) where ~refl : {r : A} (r r)
~refl {r} = reflexive {r}
record Quotient {a} {b} {c} {A : Set a} {image : Set b} (S : Setoid {a} {c} A) : Set (a b c) where
open Setoid S
field
map : A image
mapWellDefined : {x y : A} x y map x map y
mapSurjective : Surjection map
mapInjective : {x y : A} map x map y x y
record SetoidToSet {a b c : _} {A : Set a} (S : Setoid {a} {c} A) (quotient : Set b) : Set (a b c) where
open Setoid S
field
project : A quotient
wellDefined : (x y : A) (x y) project x project y
surj : Surjection project
inj : (x y : A) project x project y x y
open Setoid
reflSetoid : {a : _} (A : Set a) Setoid A
__ (reflSetoid A) a b = a b
Equivalence.reflexive (eq (reflSetoid A)) = refl
Equivalence.symmetric (eq (reflSetoid A)) {b} {.b} refl = refl
Equivalence.transitive (eq (reflSetoid A)) {b} {.b} {.b} refl refl = refl
record SetoidInjection {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 b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
injective : {x y : A} (f x) B (f y) x A y
record SetoidSurjection {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 b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
surjective : {x : B} Sg A (λ a f a B x)
record SetoidBijection {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 b c d) where
field
inj : SetoidInjection S T f
surj : SetoidSurjection S T f
wellDefined : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
wellDefined = SetoidInjection.wellDefined inj
record SetoidsBiject {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) : Set (a b c d) where
field
bij : A B
bijIsBijective : SetoidBijection S T bij
setoidInjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidInjection S T g) (hB : SetoidInjection T U h) SetoidInjection S U (h g)
SetoidInjection.wellDefined (setoidInjComp gI hI) x~y = SetoidInjection.wellDefined hI (SetoidInjection.wellDefined gI x~y)
SetoidInjection.injective (setoidInjComp gI hI) hgx~hgy = SetoidInjection.injective gI (SetoidInjection.injective hI hgx~hgy)
setoidSurjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidSurjection S T g) (hB : SetoidSurjection T U h) SetoidSurjection S U (h g)
SetoidSurjection.wellDefined (setoidSurjComp gI hI) x~y = SetoidSurjection.wellDefined hI (SetoidSurjection.wellDefined gI x~y)
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} with SetoidSurjection.surjective hI {x}
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} | a , prA with SetoidSurjection.surjective gI {a}
SetoidSurjection.surjective (setoidSurjComp {U = U} gI hI) {x} | a , prA | b , prB = b , Equivalence.transitive (Setoid.eq U) (SetoidSurjection.wellDefined hI prB) prA
where
open Setoid U
setoidBijComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidBijection S T g) (hB : SetoidBijection T U h) SetoidBijection S U (h g)
SetoidBijection.inj (setoidBijComp gB hB) = setoidInjComp (SetoidBijection.inj gB) (SetoidBijection.inj hB)
SetoidBijection.surj (setoidBijComp gB hB) = setoidSurjComp (SetoidBijection.surj gB) (SetoidBijection.surj hB)
setoidIdIsBijective : {a b : _} {A : Set a} {S : Setoid {a} {b} A} SetoidBijection S S (λ i i)
SetoidInjection.wellDefined (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidInjection.injective (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.wellDefined (SetoidBijection.surj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.surjective (SetoidBijection.surj (setoidIdIsBijective {S = S})) {x} = x , Equivalence.reflexive (Setoid.eq S)
record SetoidInvertible {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 b c d) where
field
fWellDefined : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
inverse : B A
inverseWellDefined : {x y : B} Setoid.__ T x y Setoid.__ S (inverse x) (inverse y)
isLeft : (b : B) Setoid.__ T (f (inverse b)) b
isRight : (a : A) Setoid.__ S (inverse (f a)) a
setoidBijectiveImpliesInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (bij : SetoidBijection S T f) SetoidInvertible S T f
SetoidInvertible.fWellDefined (setoidBijectiveImpliesInvertible bij) x~y = SetoidInjection.wellDefined (SetoidBijection.inj bij) x~y
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x | a , b = a
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible bij) {x} {y} x~y with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible {T = T} bij) {x} {y} x~y | a , prA with SetoidSurjection.surjective (SetoidBijection.surj bij) {y}
... | b , prB = SetoidInjection.injective (SetoidBijection.inj bij) (Equivalence.transitive (Setoid.eq T) prA (Equivalence.transitive (Setoid.eq T) x~y (Equivalence.symmetric (Setoid.eq T) prB)))
where
open Setoid T
SetoidInvertible.isLeft (setoidBijectiveImpliesInvertible bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {b}
... | a , prA = prA
SetoidInvertible.isRight (setoidBijectiveImpliesInvertible {f = f} bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {f b}
... | fb , prFb = SetoidInjection.injective (SetoidBijection.inj bij) prFb
setoidInvertibleImpliesBijective : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (inv : SetoidInvertible S T f) SetoidBijection S T f
SetoidInjection.wellDefined (SetoidBijection.inj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
SetoidInjection.injective (SetoidBijection.inj (setoidInvertibleImpliesBijective {S = S} {f = f} inv)) {x} {y} pr = Equivalence.transitive (Setoid.eq S) (Equivalence.symmetric (Setoid.eq S) (SetoidInvertible.isRight inv x)) (Equivalence.transitive (Setoid.eq S) (SetoidInvertible.inverseWellDefined inv pr) (SetoidInvertible.isRight inv y))
where
open Setoid S open Setoid S
field SetoidSurjection.wellDefined (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
map : A image SetoidSurjection.surjective (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) {x} = SetoidInvertible.inverse inv x , SetoidInvertible.isLeft inv x
mapWellDefined : {x y : A} x y map x map y
mapSurjective : Surjection map
mapInjective : {x y : A} map x map y x y
record SetoidToSet {a b c : _} {A : Set a} (S : Setoid {a} {c} A) (quotient : Set b) : Set (a b c) where
open Setoid S
field
project : A quotient
wellDefined : (x y : A) (x y) project x project y
surj : Surjection project
inj : (x y : A) project x project y x y
open Setoid
reflSetoid : {a : _} (A : Set a) Setoid A
__ (reflSetoid A) a b = a b
Equivalence.reflexive (eq (reflSetoid A)) = refl
Equivalence.symmetric (eq (reflSetoid A)) {b} {.b} refl = refl
Equivalence.transitive (eq (reflSetoid A)) {b} {.b} {.b} refl refl = refl
record SetoidInjection {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 b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
injective : {x y : A} (f x) B (f y) x A y
record SetoidSurjection {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 b c d) where
open Setoid S renaming (__ to _A_)
open Setoid T renaming (__ to _B_)
field
wellDefined : {x y : A} x A y (f x) B (f y)
surjective : {x : B} Sg A (λ a f a B x)
record SetoidBijection {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 b c d) where
field
inj : SetoidInjection S T f
surj : SetoidSurjection S T f
wellDefined : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
wellDefined = SetoidInjection.wellDefined inj
record SetoidsBiject {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) : Set (a b c d) where
field
bij : A B
bijIsBijective : SetoidBijection S T bij
setoidInjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidInjection S T g) (hB : SetoidInjection T U h) SetoidInjection S U (h g)
SetoidInjection.wellDefined (setoidInjComp gI hI) x~y = SetoidInjection.wellDefined hI (SetoidInjection.wellDefined gI x~y)
SetoidInjection.injective (setoidInjComp gI hI) hgx~hgy = SetoidInjection.injective gI (SetoidInjection.injective hI hgx~hgy)
setoidSurjComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidSurjection S T g) (hB : SetoidSurjection T U h) SetoidSurjection S U (h g)
SetoidSurjection.wellDefined (setoidSurjComp gI hI) x~y = SetoidSurjection.wellDefined hI (SetoidSurjection.wellDefined gI x~y)
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} with SetoidSurjection.surjective hI {x}
SetoidSurjection.surjective (setoidSurjComp gI hI) {x} | a , prA with SetoidSurjection.surjective gI {a}
SetoidSurjection.surjective (setoidSurjComp {U = U} gI hI) {x} | a , prA | b , prB = b , Equivalence.transitive (Setoid.eq U) (SetoidSurjection.wellDefined hI prB) prA
where
open Setoid U
setoidBijComp : {a b c d e f : _} {A : Set a} {B : Set b} {C : Set c} {S : Setoid {a} {d} A} {T : Setoid {b} {e} B} {U : Setoid {c} {f} C} {g : A B} {h : B C} (gB : SetoidBijection S T g) (hB : SetoidBijection T U h) SetoidBijection S U (h g)
SetoidBijection.inj (setoidBijComp gB hB) = setoidInjComp (SetoidBijection.inj gB) (SetoidBijection.inj hB)
SetoidBijection.surj (setoidBijComp gB hB) = setoidSurjComp (SetoidBijection.surj gB) (SetoidBijection.surj hB)
setoidIdIsBijective : {a b : _} {A : Set a} {S : Setoid {a} {b} A} SetoidBijection S S (λ i i)
SetoidInjection.wellDefined (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidInjection.injective (SetoidBijection.inj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.wellDefined (SetoidBijection.surj (setoidIdIsBijective {A = A})) = id
SetoidSurjection.surjective (SetoidBijection.surj (setoidIdIsBijective {S = S})) {x} = x , Equivalence.reflexive (Setoid.eq S)
record SetoidInvertible {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 b c d) where
field
fWellDefined : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)
inverse : B A
inverseWellDefined : {x y : B} Setoid.__ T x y Setoid.__ S (inverse x) (inverse y)
isLeft : (b : B) Setoid.__ T (f (inverse b)) b
isRight : (a : A) Setoid.__ S (inverse (f a)) a
setoidBijectiveImpliesInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (bij : SetoidBijection S T f) SetoidInvertible S T f
SetoidInvertible.fWellDefined (setoidBijectiveImpliesInvertible bij) x~y = SetoidInjection.wellDefined (SetoidBijection.inj bij) x~y
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverse (setoidBijectiveImpliesInvertible bij) x | a , b = a
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible bij) {x} {y} x~y with SetoidSurjection.surjective (SetoidBijection.surj bij) {x}
SetoidInvertible.inverseWellDefined (setoidBijectiveImpliesInvertible {T = T} bij) {x} {y} x~y | a , prA with SetoidSurjection.surjective (SetoidBijection.surj bij) {y}
... | b , prB = SetoidInjection.injective (SetoidBijection.inj bij) (Equivalence.transitive (Setoid.eq T) prA (Equivalence.transitive (Setoid.eq T) x~y (Equivalence.symmetric (Setoid.eq T) prB)))
where
open Setoid T
SetoidInvertible.isLeft (setoidBijectiveImpliesInvertible bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {b}
... | a , prA = prA
SetoidInvertible.isRight (setoidBijectiveImpliesInvertible {f = f} bij) b with SetoidSurjection.surjective (SetoidBijection.surj bij) {f b}
... | fb , prFb = SetoidInjection.injective (SetoidBijection.inj bij) prFb
setoidInvertibleImpliesBijective : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {f : A B} (inv : SetoidInvertible S T f) SetoidBijection S T f
SetoidInjection.wellDefined (SetoidBijection.inj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
SetoidInjection.injective (SetoidBijection.inj (setoidInvertibleImpliesBijective {S = S} {f = f} inv)) {x} {y} pr = Equivalence.transitive (Setoid.eq S) (Equivalence.symmetric (Setoid.eq S) (SetoidInvertible.isRight inv x)) (Equivalence.transitive (Setoid.eq S) (SetoidInvertible.inverseWellDefined inv pr) (SetoidInvertible.isRight inv y))
where
open Setoid S
SetoidSurjection.wellDefined (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) x~y = SetoidInvertible.fWellDefined inv x~y
SetoidSurjection.surjective (SetoidBijection.surj (setoidInvertibleImpliesBijective inv)) {x} = SetoidInvertible.inverse inv x , SetoidInvertible.isLeft inv x

21
Setoids/Subset.agda Normal file
View File

@@ -0,0 +1,21 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Sets.EquivalenceRelations
open import Setoids.Setoids
module Setoids.Subset {a b : _} {A : Set a} (S : Setoid {a} {b} A) where
open Setoid S
open Equivalence eq
subset : {c : _} (pred : A Set c) Set (a b c)
subset pred = ({x y : A} x y pred x pred y)
subsetSetoid : {c : _} {pred : A Set c} (subs : subset pred) Setoid (Sg A pred)
Setoid.__ (subsetSetoid subs) (x , predX) (y , predY) = Setoid.__ S x y
Equivalence.reflexive (Setoid.eq (subsetSetoid subs)) {a , b} = reflexive
Equivalence.symmetric (Setoid.eq (subsetSetoid subs)) {a , prA} {b , prB} x = symmetric x
Equivalence.transitive (Setoid.eq (subsetSetoid subs)) {a , prA} {b , prB} {c , prC} x y = transitive x y