]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/Plogic/equality.ma
Propositional equality
[helm.git] / helm / software / matita / nlibrary / Plogic / equality.ma
1 (**************************************************************************)
2 (*       ___                                                                *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "logic/pts.ma".
16
17 ninductive peq (A:Type[0]) (x:A) : A \to Prop \def
18     refl_peq : peq A x x.
19     
20 interpretation "leibnitz's equality" 'eq t x y = (peq t x y).
21  
22 ntheorem rewrite_l: ∀A.∀x.∀P:A → Prop. P x → ∀y. x = y → P y.
23 #A; #x; #P; #Hx; #y; #Heq;ncases Heq;nassumption.
24 nqed.
25
26 ntheorem sym_peq: ∀A.∀x,y:A. x = y → y = x.
27 #A; #x; #y; #Heq; napply (rewrite_l A x (λz.z=x)); 
28 ##[ @; ##| nassumption; ##]
29 nqed.
30
31 ntheorem rewrite_r: ∀A.∀x.∀P:A → Prop. P x → ∀y. y = x → P y.
32 #A; #x; #P; #Hx; #y; #Heq;ncases (sym_peq ? ? ?Heq);nassumption.
33 nqed.
34
35 (*
36 theorem eq_rect':
37  \forall A. \forall x:A. \forall P: \forall y:A. x=y \to Type.
38   P ? (refl_eq ? x) \to \forall y:A. \forall p:x=y. P y p.
39  intros.
40  exact
41   (match p1 return \lambda y. \lambda p.P y p with
42     [refl_eq \Rightarrow p]).
43 qed.
44  
45 variant reflexive_eq : \forall A:Type. reflexive A (eq A)
46 \def refl_eq.
47     
48 theorem symmetric_eq: \forall A:Type. symmetric A (eq A).
49 unfold symmetric.intros.elim H. apply refl_eq.
50 qed.
51
52 variant sym_eq : \forall A:Type.\forall x,y:A. x=y  \to y=x
53 \def symmetric_eq.
54
55 theorem transitive_eq : \forall A:Type. transitive A (eq A).
56 unfold transitive.intros.elim H1.assumption.
57 qed.
58
59 variant trans_eq : \forall A:Type.\forall x,y,z:A. x=y  \to y=z \to x=z
60 \def transitive_eq.
61
62 theorem symmetric_not_eq: \forall A:Type. symmetric A (λx,y.x ≠ y).
63 unfold symmetric.simplify.intros.unfold.intro.apply H.apply sym_eq.assumption.
64 qed.
65
66 variant sym_neq: ∀A:Type.∀x,y.x ≠ y →y ≠ x
67 ≝ symmetric_not_eq.
68
69 theorem eq_elim_r:
70  \forall A:Type.\forall x:A. \forall P: A \to Prop.
71    P x \to \forall y:A. y=x \to P y.
72 intros. elim (sym_eq ? ? ? H1).assumption.
73 qed.
74
75 theorem eq_elim_r':
76  \forall A:Type.\forall x:A. \forall P: A \to Set.
77    P x \to \forall y:A. y=x \to P y.
78 intros. elim (sym_eq ? ? ? H).assumption.
79 qed.
80
81 theorem eq_elim_r'':
82  \forall A:Type.\forall x:A. \forall P: A \to Type.
83    P x \to \forall y:A. y=x \to P y.
84 intros. elim (sym_eq ? ? ? H).assumption.
85 qed.
86
87 theorem eq_f: \forall  A,B:Type.\forall f:A\to B.
88 \forall x,y:A. x=y \to f x = f y.
89 intros.elim H.apply refl_eq.
90 qed.
91
92 theorem eq_f': \forall  A,B:Type.\forall f:A\to B.
93 \forall x,y:A. x=y \to f y = f x.
94 intros.elim H.apply refl_eq.
95 qed.
96 *)
97
98 (* 
99 coercion sym_eq.
100 coercion eq_f.
101
102
103 default "equality"
104  cic:/matita/logic/equality/eq.ind
105  cic:/matita/logic/equality/sym_eq.con
106  cic:/matita/logic/equality/transitive_eq.con
107  cic:/matita/logic/equality/eq_ind.con
108  cic:/matita/logic/equality/eq_elim_r.con
109  cic:/matita/logic/equality/eq_rec.con
110  cic:/matita/logic/equality/eq_elim_r'.con
111  cic:/matita/logic/equality/eq_rect.con
112  cic:/matita/logic/equality/eq_elim_r''.con
113  cic:/matita/logic/equality/eq_f.con
114 (* *)
115  cic:/matita/logic/equality/eq_OF_eq.con.
116 (* *)
117 (*  
118  cic:/matita/logic/equality/eq_f'.con. (* \x.sym (eq_f x) *)
119  *)
120  
121 theorem eq_f2: \forall  A,B,C:Type.\forall f:A\to B \to C.
122 \forall x1,x2:A. \forall y1,y2:B.
123 x1=x2 \to y1=y2 \to f x1 y1 = f x2 y2.
124 intros.elim H1.elim H.reflexivity.
125 qed.
126
127 definition comp \def
128  \lambda A.
129   \lambda x,y,y':A.
130    \lambda eq1:x=y.
131     \lambda eq2:x=y'.
132      eq_ind ? ? (\lambda a.a=y') eq2 ? eq1.
133      
134 lemma trans_sym_eq:
135  \forall A.
136   \forall x,y:A.
137    \forall u:x=y.
138     comp ? ? ? ? u u = refl_eq ? y.
139  intros.
140  apply (eq_rect' ? ? ? ? ? u).
141  reflexivity.
142 qed.
143
144 definition nu \def
145  \lambda A.
146   \lambda H: \forall x,y:A. decidable (x=y).
147    \lambda x,y. \lambda p:x=y.
148      match H x y with
149       [ (or_introl p') \Rightarrow p'
150       | (or_intror K) \Rightarrow False_ind ? (K p) ].
151
152 theorem nu_constant:
153  \forall A.
154   \forall H: \forall x,y:A. decidable (x=y).
155    \forall x,y:A.
156     \forall u,v:x=y.
157      nu ? H ? ? u = nu ? H ? ? v.
158  intros.
159  unfold nu.
160  unfold decidable in H.
161  apply (Or_ind' ? ? ? ? ? (H x y)); simplify.
162   intro; reflexivity.
163   intro; elim (q u).
164 qed.
165
166 definition nu_inv \def
167  \lambda A.
168   \lambda H: \forall x,y:A. decidable (x=y).
169    \lambda x,y:A.
170     \lambda v:x=y.
171      comp ? ? ? ? (nu ? H ? ? (refl_eq ? x)) v.
172
173 theorem nu_left_inv:
174  \forall A.
175   \forall H: \forall x,y:A. decidable (x=y).
176    \forall x,y:A.
177     \forall u:x=y.
178      nu_inv ? H ? ? (nu ? H ? ? u) = u.
179  intros.
180  apply (eq_rect' ? ? ? ? ? u).
181  unfold nu_inv.
182  apply trans_sym_eq.
183 qed.
184
185 theorem eq_to_eq_to_eq_p_q:
186  \forall A. \forall x,y:A.
187   (\forall x,y:A. decidable (x=y)) \to
188    \forall p,q:x=y. p=q.
189  intros.
190  rewrite < (nu_left_inv ? H ? ? p).
191  rewrite < (nu_left_inv ? H ? ? q).
192  elim (nu_constant ? H ? ? q).
193  reflexivity.
194 qed.
195
196 (*CSC: alternative proof that does not pollute the environment with
197   technical lemmata. Unfortunately, it is a pain to do without proper
198   support for let-ins.
199 theorem eq_to_eq_to_eq_p_q:
200  \forall A. \forall x,y:A.
201   (\forall x,y:A. decidable (x=y)) \to
202    \forall p,q:x=y. p=q.
203 intros.
204 letin nu \def
205  (\lambda x,y. \lambda p:x=y.
206    match H x y with
207     [ (or_introl p') \Rightarrow p'
208     | (or_intror K) \Rightarrow False_ind ? (K p) ]).
209 cut
210  (\forall q:x=y.
211    eq_ind ? ? (\lambda z. z=y) (nu ? ? q) ? (nu ? ? (refl_eq ? x))
212    = q).
213 focus 8.
214  clear q; clear p.
215  intro.
216  apply (eq_rect' ? ? ? ? ? q);
217  fold simplify (nu ? ? (refl_eq ? x)).
218  generalize in match (nu ? ? (refl_eq ? x)); intro.
219  apply
220   (eq_rect' A x
221    (\lambda y. \lambda u.
222     eq_ind A x (\lambda a.a=y) u y u = refl_eq ? y)
223    ? x H1).
224  reflexivity.
225 unfocus.
226 rewrite < (Hcut p); fold simplify (nu ? ? p).
227 rewrite < (Hcut q); fold simplify (nu ? ? q).
228 apply (Or_ind' (x=x) (x \neq x)
229  (\lambda p:decidable (x=x). eq_ind A x (\lambda z.z=y) (nu x y p) x
230    ([\lambda H1.eq A x x]
231     match p with
232     [(or_introl p') \Rightarrow p'
233     |(or_intror K) \Rightarrow False_ind (x=x) (K (refl_eq A x))]) =
234    eq_ind A x (\lambda z.z=y) (nu x y q) x
235     ([\lambda H1.eq A x x]
236      match p with
237     [(or_introl p') \Rightarrow p'
238     |(or_intror K) \Rightarrow False_ind (x=x) (K (refl_eq A x))]))
239  ? ? (H x x)).
240 intro; simplify; reflexivity.
241 intro q; elim (q (refl_eq ? x)).
242 qed.
243 *)
244
245 (*
246 theorem a:\forall x.x=x\land True.
247
248 2:intros;
249   split;
250   [
251     exact (refl_eq Prop x);
252   |
253     exact I;
254   ]
255 1:
256   skip
257 ]
258 qed.
259 *) *)
260