Skip to content

Commit

Permalink
Merge pull request #1443 from andreaslyn/infinitary-universal-algebra
Browse files Browse the repository at this point in the history
Term algebra and congruence
  • Loading branch information
andreaslyn authored Apr 8, 2021
2 parents dd3c399 + 79fd12c commit 6d0c7ac
Show file tree
Hide file tree
Showing 15 changed files with 579 additions and 71 deletions.
2 changes: 2 additions & 0 deletions _CoqProject
Original file line number Diff line number Diff line change
Expand Up @@ -354,6 +354,8 @@ theories/Spectra/Coinductive.v
theories/Algebra/Universal/Algebra.v
theories/Algebra/Universal/Homomorphism.v
theories/Algebra/Universal/Operation.v
theories/Algebra/Universal/Congruence.v
theories/Algebra/Universal/TermAlgebra.v

#
# Algebra
Expand Down
9 changes: 5 additions & 4 deletions theories/Algebra/Universal/Algebra.v
Original file line number Diff line number Diff line change
Expand Up @@ -95,6 +95,8 @@ Global Existing Instance hset_algebra.

Global Coercion carriers : Algebra >-> Funclass.

Bind Scope Algebra_scope with Algebra.

Definition SigAlgebra (σ : Signature) : Type
:= {c : Carriers σ
| { _ : forall (u : Symbol σ), Operation c (σ u)
Expand Down Expand Up @@ -158,8 +160,7 @@ Defined.
Arguments path_path_algebra {_} {σ} {A B}%Algebra_scope (p q r)%path_scope.

Module notations_algebra.
(** Conflicts with inverse notation so cannot be reserved *)
Global Notation "u ^^ A" := (operations A u)
(at level 6, no associativity)
: Algebra_scope.

Global Notation "u .# A" := (operations A u) : Algebra_scope.

End notations_algebra.
91 changes: 91 additions & 0 deletions theories/Algebra/Universal/Congruence.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,91 @@
(** This file implements algebra congruence relation. It serves as a
universal algebra generalization of normal subgroup, ring ideal, etc.
Congruence is used to construct quotients, in similarity with how
normal subgroup and ring ideal is used to construct quotients. *)

Require Export HoTT.Algebra.Universal.Algebra.

Require Import
HoTT.Basics
HoTT.Types
HoTT.HProp
HoTT.Classes.interfaces.canonical_names
HoTT.Algebra.Universal.Homomorphism.

Unset Elimination Schemes.

Import notations_algebra.

Section congruence.
Context {σ : Signature} (A : Algebra σ) (Φ : forall s, Relation (A s)).

(** A finitary operation [f : A s1 * A s2 * ... * A sn -> A t]
satisfies [OpCompatible f] iff
<<
Φ s1 x1 y1 * Φ s2 x2 y2 * ... * Φ sn xn yn
>>
implies
<<
Φ t (f (x1, x2, ..., xn)) (f (y1, y2, ..., yn)).
>>
The below definition generalizes this to infinitary operations.
*)

Definition OpCompatible {w : SymbolType σ} (f : Operation A w) : Type
:= forall (a b : DomOperation A w),
(forall i : Arity w, Φ (sorts_dom w i) (a i) (b i)) ->
Φ (sort_cod w) (f a) (f b).

Class OpsCompatible : Type
:= ops_compatible : forall (u : Symbol σ), OpCompatible u.#A.

Global Instance trunc_ops_compatible `{Funext} {n : trunc_index}
`{!forall s x y, IsTrunc n (Φ s x y)}
: IsTrunc n OpsCompatible.
Proof.
apply trunc_forall.
Defined.

(** A family of relations [Φ] is a congruence iff it is a family of
mere equivalence relations and [OpsCompatible A Φ] holds. *)

Class IsCongruence : Type := Build_IsCongruence
{ is_mere_relation_cong : forall (s : Sort σ), is_mere_relation (A s) (Φ s)
; equiv_rel_cong : forall (s : Sort σ), EquivRel (Φ s)
; ops_compatible_cong : OpsCompatible }.

Global Arguments Build_IsCongruence {is_mere_relation_cong}
{equiv_rel_cong}
{ops_compatible_cong}.

Global Existing Instance is_mere_relation_cong.

Global Existing Instance equiv_rel_cong.

Global Existing Instance ops_compatible_cong.

Global Instance hprop_is_congruence `{Funext} : IsHProp IsCongruence.
Proof.
apply (equiv_hprop_allpath _)^-1.
intros [C1 C2 C3] [D1 D2 D3].
by destruct (path_ishprop C1 D1),
(path_ishprop C2 D2),
(path_ishprop C3 D3).
Defined.

End congruence.

(** A homomorphism [f : forall s, A s -> B s] is compatible
with a congruence [Φ] iff [Φ s x y] implies [f s x = f s y]. *)

Definition HomCompatible {σ : Signature} {A B : Algebra σ}
(Φ : forall s, Relation (A s)) `{!IsCongruence A Φ}
(f : forall s, A s -> B s) `{!IsHomomorphism f}
: Type
:= forall s (x y : A s), Φ s x y -> f s x = f s y.

2 changes: 1 addition & 1 deletion theories/Algebra/Universal/Homomorphism.v
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ Section is_homomorphism.
Qed.

Class IsHomomorphism : Type
:= oppreserving_hom : forall (u : Symbol σ), OpPreserving u^^A u^^B.
:= oppreserving_hom : forall (u : Symbol σ), OpPreserving u.#A u.#B.

Global Instance hprop_is_homomorphism `{Funext}
: IsHProp IsHomomorphism.
Expand Down
Loading

0 comments on commit 6d0c7ac

Please sign in to comment.