]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/limits/Class/eq.ma
limits: reorganized and attached to nightly tests (cow compiles fully)
[helm.git] / helm / software / matita / contribs / limits / Class / eq.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "Class/defs.ma".
16
17 (* ACZEL CATEGORIES:
18    - Here we prove the standard properties of the equality.
19 *)
20
21 theorem ceq_cin_sx: ∀C,c1,c2. ceq C c1 c2 → cin ? c1.
22 intros; elim H; clear H c1 c2; autobatch.
23 qed.
24
25 theorem ceq_cin_dx: ∀C,c1,c2. ceq C c1 c2 → cin ? c2.
26 intros; elim H; clear H c1 c2; autobatch.
27 qed.
28
29 theorem ceq_trans: ∀C,c1,c2. ceq C c1 c2 → ∀c3. ceq ? c2 c3 → ceq ? c1 c3.
30 intros 4; elim H; clear H c1 c2; autobatch.
31 qed.
32
33 theorem ceq_conf: ∀C,c1,c2. ceq C c1 c2 → ∀c3. ceq ? c1 c3 → ceq ? c2 c3.
34 intros 4; elim H; clear H c1 c2; autobatch depth = 4.
35 qed.
36
37 theorem ceq_sym: ∀ C,c1,c2. ceq C c1 c2 → ceq ? c2 c1.
38 intros; autobatch depth = 4.
39 qed.
40
41 theorem ceq_repl: ∀C,c1,c2. ceq C c1 c2 →
42                       ∀d1. ceq ? c1 d1 → ∀d2. ceq ? c2 d2 → ceq ? d1 d2.
43 intros; autobatch.
44 qed.