view src/Fundamental.agda @ 271:c209aebeab2a

Fundamental again
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 24 Jan 2023 16:40:39 +0900
parents 0081e1ed5ead
children ce372f6347d6
line wrap: on
line source

-- fundamental homomorphism theorem
--

open import Level hiding ( suc ; zero )
module Fundamental (c l : Level) where

open import Algebra
open import Algebra.Structures
open import Algebra.Definitions
open import Algebra.Core
open import Algebra.Bundles

open import Data.Product
open import Relation.Binary.PropositionalEquality 
open import Gutil0
import Function.Definitions as FunctionDefinitions

import Algebra.Morphism.Definitions as MorphismDefinitions
open import Algebra.Morphism.Structures 

open import Tactic.MonoidSolver using (solve; solve-macro)

--
-- Given two groups G and H and a group homomorphism f : G → H,
-- let K be a normal subgroup in G and φ the natural surjective homomorphism G → G/K
-- (where G/K is the quotient group of G by K).
-- If K is a subset of ker(f) then there exists a unique homomorphism h: G/K → H such that f = h∘φ.
--     https://en.wikipedia.org/wiki/Fundamental_theorem_on_homomorphisms
--
--       f
--    G --→ H
--    |   /
--  φ |  / h
--    ↓ /
--    G/K
--

import Relation.Binary.Reasoning.Setoid as EqReasoning

_<_∙_> :  (m : Group c l )  →    Group.Carrier m →  Group.Carrier m →  Group.Carrier m
m < x ∙ y > =  Group._∙_ m x y

infixr 9 _<_∙_>

--
--  Coset : N : NormalSubGroup →  { a ∙ n | G ∋ a , N ∋ n }
--


open GroupMorphisms

import Axiom.Extensionality.Propositional
postulate f-extensionality : { n m : Level}  → Axiom.Extensionality.Propositional.Extensionality n m

record NormalSubGroup (A : Group (c ⊔ l) l )  : Set (c ⊔ l)  where
   open Group A
   field  
       ⟦_⟧ : Group.Carrier A → Group.Carrier A 
       normal :  IsGroupHomomorphism (GR A) (GR A)  ⟦_⟧
       comm : {a b :  Carrier } → b ∙ ⟦ a ⟧  ≈ ⟦ a ⟧  ∙ b 

-- Set of a ∙ ∃ n ∈ N
--
record aN {A : Group (c ⊔ l) l }  (N : NormalSubGroup A ) (x : Group.Carrier A  ) : Set (c ⊔ l) where 
    open Group A
    open NormalSubGroup N
    field 
        a n : Group.Carrier A 
        aN=x :  a ∙ ⟦ n ⟧ ≈ x

open import Tactic.MonoidSolver using (solve; solve-macro)

_/_ : (A : Group (c ⊔ l) l ) (N  : NormalSubGroup A ) → Group (c ⊔ l) l
_/_ A N  = record {
    Carrier  = (x : Group.Carrier A  ) → aN N x
    ; _≈_      = ?
    ; _∙_      = _+_
    ; ε        = λ x → record { a = x  ; n = ε  ; x=aN = ?  }
    ; _⁻¹      = λ f x → record { a = x ∙ ⟦ n (f x) ⟧  ⁻¹  ; n = n (f x)  ; x=aN = ?  }
    ; isGroup = record { isMonoid  = record { isSemigroup = record { isMagma = record {
       isEquivalence = record {refl = ? ; trans = ? ; sym = ? }
       ; ∙-cong = ? }
       ; assoc = ? }
       ; identity =  ? , (λ q → ? )  }
       ; inverse   = ( (λ x → ? ) , (λ x → ? ))  
       ; ⁻¹-cong   = λ i=j → ?
      }
   } where
       open Group A
       open aN
       open NormalSubGroup N
       open IsGroupHomomorphism normal 
       _+_ : (f g : (x : Group.Carrier A  )  → aN N x ) → (x : Group.Carrier A  )  → aN N x
       _+_ f g x = record { a = x ⁻¹ ∙ a (f x) ∙ a (g x) ; n = n (f x) ∙ n (g x) ; aN=x = q00 } where
          q00 : ( x ⁻¹ ∙  a (f x) ∙ a (g x)  ) ∙ ⟦ n (f x) ∙ n (g x) ⟧  ≈ x
          q00 = begin 
              ( x ⁻¹ ∙ a (f x) ∙ a (g x) ) ∙ ⟦ n (f x) ∙ n (g x) ⟧  ≈⟨ ∙-cong (assoc _ _ _) (homo _ _ ) ⟩
              x ⁻¹ ∙ (a (f x) ∙ a (g x) ) ∙ ( ⟦ n (f x) ⟧  ∙ ⟦ n (g x) ⟧ ) ≈⟨ solve monoid ⟩
              x ⁻¹ ∙ (a (f x) ∙ ((a (g x)  ∙ ⟦ n (f x) ⟧ ) ∙ ⟦ n (g x) ⟧ )) ≈⟨ ∙-cong (grefl A) (∙-cong (grefl A) (∙-cong comm (grefl A) ))  ⟩
              x ⁻¹ ∙ (a (f x) ∙ ((⟦ n (f x) ⟧  ∙ a (g x)) ∙ ⟦ n (g x) ⟧ ))  ≈⟨ solve monoid ⟩
              x ⁻¹ ∙ (a (f x) ∙ ⟦ n (f x) ⟧ ) ∙ (a (g x) ∙ ⟦ n (g x) ⟧ )  ≈⟨ ∙-cong (grefl A) (aN=x (g x)  ) ⟩
              x ⁻¹ ∙ (a (f x) ∙ ⟦ n (f x) ⟧ ) ∙ x   ≈⟨ ∙-cong (∙-cong (grefl A) (aN=x (f x))) (grefl A)  ⟩
              (x ⁻¹ ∙ x ) ∙ x   ≈⟨ ∙-cong (proj₁ inverse _ ) (grefl A)  ⟩
              ε ∙ x  ≈⟨ solve monoid  ⟩
              x ∎ where 
                open EqReasoning (Algebra.Group.setoid A)
                -- open IsGroup isGroup


-- K ⊂ ker(f)
K⊆ker : (G H : Group (c ⊔ l) l)  (K : NormalSubGroup G) (f : Group.Carrier G → Group.Carrier H ) → Set ( c  Level.⊔  l ) 
K⊆ker G H K f = (x : Group.Carrier G ) → f ( ? K x ) ≈ ε   where
  open Group H

record FundamentalHomomorphism (G H : Group (c ⊔ l) l)  
    (f : Group.Carrier G → Group.Carrier H ) 
    (homo : IsGroupHomomorphism (GR G) (GR H) f ) 
    (K : NormalSubGroup G ) (kf : K⊆ker G H K f) :  Set ( c  Level.⊔ ( Level.suc l) ) where
  open Group H
  field
    h : ? → Group.Carrier H
    h-homo :  IsGroupHomomorphism (GR (G / ?) ) (GR H) h
    unique : (x : Group.Carrier G)  →  f x ≈ h ( ? x )

FundamentalHomomorphismTheorm : (G H : Group (c ⊔ l) l)
    (f : Group.Carrier G → Group.Carrier H )
    (homo : IsGroupHomomorphism (GR G) (GR H) f )
    (K : NormalSubGroup G ) → (kf : K⊆ker G H K f)   → FundamentalHomomorphism G H f homo K kf
FundamentalHomomorphismTheorm G H f homo K kf = record {
     h = h
   ; h-homo = h-homo
   ; unique = unique
  } where
    open Group H
    h : ? G K → Carrier   --  G / K → H
    h r = f ?
    _o_ = Group._∙_ G
    h03 : (x y : Group.Carrier (G / ? ) ) →  h ( ? x y ) ≈ h x ∙ h y
    h03 x y = {!!}
    h-homo :  IsGroupHomomorphism (GR (G / ? ) ) (GR H) h
    h-homo = record {
     isMonoidHomomorphism = record {
          isMagmaHomomorphism = record {
             isRelHomomorphism = record { cong = λ {x} {y} eq → {!!} } 
           ; homo = h03 }
        ; ε-homo = {!!} }
       ; ⁻¹-homo = {!!} }
    unique : (x : Group.Carrier G)  →  f x ≈ h ( ? x )
    unique x = begin
         f x ≈⟨ grefl H ⟩
         h ( ? x ) ∎ where open EqReasoning (Algebra.Group.setoid H )