Built with Alectryon, running Coq+SerAPI v8.19.0+0.19.3. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done]
Import algebra_notations ne_list.notations.Sectioncongruence.Context {σ : Signature} (A : Algebra σ) (Φ : ∀s, Relation (A s)).(** An 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).>>*)DefinitionOpCompatible {w : SymbolType σ} (f : Operation A w)
: Type
:= ∀ (ab : FamilyProd A (dom_symboltype w)),
for_all_2_family_prod A A Φ a b ->
Φ (cod_symboltype w) (ap_operation f a) (ap_operation f b).ClassOpsCompatible : Type
:= ops_compatible : ∀ (u : Symbol σ), OpCompatible u.#A.
σ: Signature A: Algebra σ Φ: foralls : Sort σ, Relation (A s) H: Funext n: trunc_index H0: forall (s : Sort σ) (xy : A s),
IsTrunc n (Φ s x y)
IsTrunc n OpsCompatible
σ: Signature A: Algebra σ Φ: foralls : Sort σ, Relation (A s) H: Funext n: trunc_index H0: forall (s : Sort σ) (xy : A s),
IsTrunc n (Φ s x y)
IsTrunc n OpsCompatible
apply istrunc_forall.Qed.(** A family of relations [Φ] is a congruence iff it is a family of mere equivalence relations and [OpsCompatible A Φ] holds. *)ClassIsCongruence : Type := BuildIsCongruence
{ is_mere_relation_cong :: ∀ (s : Sort σ), is_mere_relation (A s) (Φ s)
; equiv_rel_cong :: ∀ (s : Sort σ), EquivRel (Φ s)
; ops_compatible_cong :: OpsCompatible }.GlobalArguments BuildIsCongruence {is_mere_relation_cong}
{equiv_rel_cong}
{ops_compatible_cong}.
bydestruct (path_ishprop C1 D1),
(path_ishprop C2 D2),
(path_ishprop C3 D3).Defined.Endcongruence.(** If [Φ] is a congruence and [f : A s1 → A s2 → ... → A sn] an operation such that [OpCompatible A Φ f] holds. Then [OpCompatible (f x)] holds for all [x : A s1]. *)
σ: Signature A: Algebra σ Φ: foralls : Sort σ, Relation (A s) IsCongruence0: IsCongruence A Φ s: Sort σ w: SymbolType σ f: Operation A (s ::: w) x: A s P: OpCompatible A Φ f
OpCompatible A Φ (f x)
σ: Signature A: Algebra σ Φ: foralls : Sort σ, Relation (A s) IsCongruence0: IsCongruence A Φ s: Sort σ w: SymbolType σ f: Operation A (s ::: w) x: A s P: OpCompatible A Φ f
OpCompatible A Φ (f x)
σ: Signature A: Algebra σ Φ: foralls : Sort σ, Relation (A s) IsCongruence0: IsCongruence A Φ s: Sort σ w: SymbolType σ f: Operation A (s ::: w) x: A s P: OpCompatible A Φ f a, b: FamilyProd A (dom_symboltype w) R: for_all_2_family_prod A A Φ a b
Φ (cod_symboltype w)
(ap_operation (f x) a)
(ap_operation (f x) b)