]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/logic/equality.ma
715423143e7185debaf41b69046ba5e1ccde202b
[helm.git] / helm / software / matita / nlibrary / logic / equality.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 "logic/connectives.ma".
16 include "properties/relations.ma".
17
18 ninductive eq (A: Type[0]) (a: A) : A → CProp[0] ≝
19  refl: eq A a a.
20
21 nlemma eq_rect_CProp0_r':
22  ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (refl A a) → P x p.
23  #A; #a; #x; #p; ncases p; #P; #H; nassumption.
24 nqed.
25
26 nlemma eq_rect_CProp0_r:
27  ∀A.∀a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
28  #A; #a; #P; #p; #x0; #p0; napply (eq_rect_CProp0_r' ??? p0); nassumption.
29 nqed.
30
31 interpretation "leibnitz's equality" 'eq t x y = (eq t x y).
32
33 interpretation "leibnitz's non-equality" 'neq t x y = (Not (eq t x y)).
34
35 ndefinition EQ: ∀A:Type[0]. equivalence_relation A.
36  #A; napply mk_equivalence_relation
37   [ napply eq
38   | napply refl
39   | #x; #y; #H; nrewrite < H; napply refl
40   | #x; #y; #z; #Hyx; #Hxz; nrewrite < Hxz; nassumption]
41 nqed.