Built with Alectryon, running Coq+SerAPI v8.18.0+0.18.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. Section congruence. 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). >> *) Definition OpCompatible {w : SymbolType σ} (f : Operation A w) : Type := (a b : 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). Class OpsCompatible : Type := ops_compatible : (u : Symbol σ), OpCompatible u.#A.
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext
n: trunc_index
H0: forall (s : Sort σ) (x y : A s), IsTrunc n (Φ s x y)

IsTrunc n OpsCompatible
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext
n: trunc_index
H0: forall (s : Sort σ) (x y : 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. *) Class IsCongruence : 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 }. Global Arguments BuildIsCongruence {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.
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext

IsHProp IsCongruence
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext

IsHProp IsCongruence
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext

forall x y : IsCongruence, x = y
σ: Signature
A: Algebra σ
Φ: forall s : Sort σ, Relation (A s)
H: Funext
C1: forall (s : Sort σ) (x y : A s), IsHProp (Φ s x y)
C2: forall s : Sort σ, EquivRel (Φ s)
C3: OpsCompatible
D1: forall (s : Sort σ) (x y : A s), IsHProp (Φ s x y)
D2: forall s : Sort σ, EquivRel (Φ s)
D3: OpsCompatible

{| is_mere_relation_cong := C1; equiv_rel_cong := C2; ops_compatible_cong := C3 |} = {| is_mere_relation_cong := D1; equiv_rel_cong := D2; ops_compatible_cong := D3 |}
by destruct (path_ishprop C1 D1), (path_ishprop C2 D2), (path_ishprop C3 D3). Defined. End congruence. (** 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 σ
Φ: forall s : 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 σ
Φ: forall s : 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 σ
Φ: forall s : 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)
exact (P (x,a) (x,b) (EquivRel_Reflexive x, R)). Defined.