mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 06:38:39 +00:00
46 lines
1.7 KiB
Agda
46 lines
1.7 KiB
Agda
{-# 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 Setoids.Subset
|
||
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.Lemmas
|
||
open import Groups.Abelian.Definition
|
||
open import Groups.QuotientGroup.Definition
|
||
|
||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||
|
||
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
|
||
|
||
imageGroupPred : B → Set (a ⊔ d)
|
||
imageGroupPred b = Sg A (λ a → Setoid._∼_ T (f a) b)
|
||
|
||
imageGroupSubset : subset T imageGroupPred
|
||
imageGroupSubset {x} {y} x=y (a , fa=x) = a , transitive fa=x x=y
|
||
where
|
||
open Setoid T
|
||
open Equivalence eq
|
||
|
||
imageGroupSubgroup : subgroup H imageGroupPred
|
||
_&&_.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)
|
||
where
|
||
open Setoid T
|
||
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
|