X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2FPlogic%2Fequality.ma;h=68e7aa50917fae4e68e80d702b8bb89a3a035350;hb=b1736386d7728463cfcc7b3539633db9154809b5;hp=275ad8db8159fde3f0fbe24f48df82734be7f95c;hpb=c46cc290b8b7c00411cf5b8761eb65f9dda7a62e;p=helm.git diff --git a/helm/software/matita/nlibrary/Plogic/equality.ma b/helm/software/matita/nlibrary/Plogic/equality.ma index 275ad8db8..68e7aa509 100644 --- a/helm/software/matita/nlibrary/Plogic/equality.ma +++ b/helm/software/matita/nlibrary/Plogic/equality.ma @@ -14,247 +14,124 @@ include "logic/pts.ma". -ninductive peq (A:Type[0]) (x:A) : A \to Prop \def - refl_peq : peq A x x. +ninductive eq (A:Type[2]) (x:A) : A → Prop ≝ + refl: eq A x x. -interpretation "leibnitz's equality" 'eq t x y = (peq t x y). - -ntheorem rewrite_l: ∀A.∀x.∀P:A → Prop. P x → ∀y. x = y → P y. -#A; #x; #P; #Hx; #y; #Heq;ncases Heq;nassumption. -nqed. +interpretation "leibnitz's equality" 'eq t x y = (eq t x y). -ntheorem sym_peq: ∀A.∀x,y:A. x = y → y = x. -#A; #x; #y; #Heq; napply (rewrite_l A x (λz.z=x)); -##[ @; ##| nassumption; ##] +nlemma eq_rect_r: + ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type. P a (refl A a) → P x p. + #A; #a; #x; #p; ncases p; #P; #H; nassumption. nqed. -ntheorem rewrite_r: ∀A.∀x.∀P:A → Prop. P x → ∀y. y = x → P y. -#A; #x; #P; #Hx; #y; #Heq;ncases (sym_peq ? ? ?Heq);nassumption. +nlemma eq_ind_r : + ∀A.∀a.∀P: ∀x:A. x = a → Prop. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + #A; #a; #P; #p; #x0; #p0; napply (eq_rect_r ? ? ? p0); nassumption. nqed. (* -theorem eq_rect': - \forall A. \forall x:A. \forall P: \forall y:A. x=y \to Type. - P ? (refl_eq ? x) \to \forall y:A. \forall p:x=y. P y p. - intros. - exact - (match p1 return \lambda y. \lambda p.P y p with - [refl_eq \Rightarrow p]). -qed. - -variant reflexive_eq : \forall A:Type. reflexive A (eq A) -\def refl_eq. - -theorem symmetric_eq: \forall A:Type. symmetric A (eq A). -unfold symmetric.intros.elim H. apply refl_eq. -qed. - -variant sym_eq : \forall A:Type.\forall x,y:A. x=y \to y=x -\def symmetric_eq. - -theorem transitive_eq : \forall A:Type. transitive A (eq A). -unfold transitive.intros.elim H1.assumption. -qed. - -variant trans_eq : \forall A:Type.\forall x,y,z:A. x=y \to y=z \to x=z -\def transitive_eq. - -theorem symmetric_not_eq: \forall A:Type. symmetric A (λx,y.x ≠ y). -unfold symmetric.simplify.intros.unfold.intro.apply H.apply sym_eq.assumption. -qed. - -variant sym_neq: ∀A:Type.∀x,y.x ≠ y →y ≠ x -≝ symmetric_not_eq. - -theorem eq_elim_r: - \forall A:Type.\forall x:A. \forall P: A \to Prop. - P x \to \forall y:A. y=x \to P y. -intros. elim (sym_eq ? ? ? H1).assumption. -qed. - -theorem eq_elim_r': - \forall A:Type.\forall x:A. \forall P: A \to Set. - P x \to \forall y:A. y=x \to P y. -intros. elim (sym_eq ? ? ? H).assumption. -qed. - -theorem eq_elim_r'': - \forall A:Type.\forall x:A. \forall P: A \to Type. - P x \to \forall y:A. y=x \to P y. -intros. elim (sym_eq ? ? ? H).assumption. -qed. - -theorem eq_f: \forall A,B:Type.\forall f:A\to B. -\forall x,y:A. x=y \to f x = f y. -intros.elim H.apply refl_eq. -qed. - -theorem eq_f': \forall A,B:Type.\forall f:A\to B. -\forall x,y:A. x=y \to f y = f x. -intros.elim H.apply refl_eq. -qed. +nlemma eq_ind_r : + ∀A.∀a.∀P: ∀x:A. x = a → Prop. P a (refl_eq A a) → ∀x.∀p:eq ? x a.P x p. + #A; #a; #P; #p; #x0; #p0; ngeneralize in match p; +ncases p0; #Heq; nassumption. +nqed. *) -(* -coercion sym_eq. -coercion eq_f. - - -default "equality" - cic:/matita/logic/equality/eq.ind - cic:/matita/logic/equality/sym_eq.con - cic:/matita/logic/equality/transitive_eq.con - cic:/matita/logic/equality/eq_ind.con - cic:/matita/logic/equality/eq_elim_r.con - cic:/matita/logic/equality/eq_rec.con - cic:/matita/logic/equality/eq_elim_r'.con - cic:/matita/logic/equality/eq_rect.con - cic:/matita/logic/equality/eq_elim_r''.con - cic:/matita/logic/equality/eq_f.con -(* *) - cic:/matita/logic/equality/eq_OF_eq.con. -(* *) -(* - cic:/matita/logic/equality/eq_f'.con. (* \x.sym (eq_f x) *) - *) - -theorem eq_f2: \forall A,B,C:Type.\forall f:A\to B \to C. -\forall x1,x2:A. \forall y1,y2:B. -x1=x2 \to y1=y2 \to f x1 y1 = f x2 y2. -intros.elim H1.elim H.reflexivity. -qed. - -definition comp \def - \lambda A. - \lambda x,y,y':A. - \lambda eq1:x=y. - \lambda eq2:x=y'. - eq_ind ? ? (\lambda a.a=y') eq2 ? eq1. - -lemma trans_sym_eq: - \forall A. - \forall x,y:A. - \forall u:x=y. - comp ? ? ? ? u u = refl_eq ? y. - intros. - apply (eq_rect' ? ? ? ? ? u). - reflexivity. -qed. - -definition nu \def - \lambda A. - \lambda H: \forall x,y:A. decidable (x=y). - \lambda x,y. \lambda p:x=y. - match H x y with - [ (or_introl p') \Rightarrow p' - | (or_intror K) \Rightarrow False_ind ? (K p) ]. +ntheorem rewrite_l: ∀A:Type[2].∀x.∀P:A → Prop. P x → ∀y. x = y → P y. +#A; #x; #P; #Hx; #y; #Heq;ncases Heq;nassumption. +nqed. -theorem nu_constant: - \forall A. - \forall H: \forall x,y:A. decidable (x=y). - \forall x,y:A. - \forall u,v:x=y. - nu ? H ? ? u = nu ? H ? ? v. - intros. - unfold nu. - unfold decidable in H. - apply (Or_ind' ? ? ? ? ? (H x y)); simplify. - intro; reflexivity. - intro; elim (q u). -qed. +ntheorem sym_eq: ∀A:Type[2].∀x,y:A. x = y → y = x. +#A; #x; #y; #Heq; napply (rewrite_l A x (λz.z=x)); +##[ @; ##| nassumption; ##] +nqed. -definition nu_inv \def - \lambda A. - \lambda H: \forall x,y:A. decidable (x=y). - \lambda x,y:A. - \lambda v:x=y. - comp ? ? ? ? (nu ? H ? ? (refl_eq ? x)) v. +ntheorem rewrite_r: ∀A:Type[2].∀x.∀P:A → Prop. P x → ∀y. y = x → P y. +#A; #x; #P; #Hx; #y; #Heq;ncases (sym_eq ? ? ?Heq);nassumption. +nqed. -theorem nu_left_inv: - \forall A. - \forall H: \forall x,y:A. decidable (x=y). - \forall x,y:A. - \forall u:x=y. - nu_inv ? H ? ? (nu ? H ? ? u) = u. - intros. - apply (eq_rect' ? ? ? ? ? u). - unfold nu_inv. - apply trans_sym_eq. -qed. +ntheorem eq_coerc: ∀A,B:Type[1].A→(A=B)→B. +#A; #B; #Ha; #Heq;nelim Heq; nassumption. +nqed. -theorem eq_to_eq_to_eq_p_q: - \forall A. \forall x,y:A. - (\forall x,y:A. decidable (x=y)) \to - \forall p,q:x=y. p=q. - intros. - rewrite < (nu_left_inv ? H ? ? p). - rewrite < (nu_left_inv ? H ? ? q). - elim (nu_constant ? H ? ? q). - reflexivity. -qed. +ndefinition R0 ≝ λT:Type[0].λt:T.t. + +ndefinition R1 ≝ eq_rect_Type0. + +ndefinition R2 : + ∀T0:Type[0]. + ∀a0:T0. + ∀T1:∀x0:T0. a0=x0 → Type[0]. + ∀a1:T1 a0 (refl ? a0). + ∀T2:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0. R1 ?? T1 a1 ? p0 = x1 → Type[0]. + ∀a2:T2 a0 (refl ? a0) a1 (refl ? a1). + ∀b0:T0. + ∀e0:a0 = b0. + ∀b1: T1 b0 e0. + ∀e1:R1 ?? T1 a1 ? e0 = b1. + T2 b0 e0 b1 e1. +#T0;#a0;#T1;#a1;#T2;#a2;#b0;#e0;#b1;#e1; +napply (eq_rect_Type0 ????? e1); +napply (R1 ?? ? ?? e0); +napply a2; +nqed. -(*CSC: alternative proof that does not pollute the environment with - technical lemmata. Unfortunately, it is a pain to do without proper - support for let-ins. -theorem eq_to_eq_to_eq_p_q: - \forall A. \forall x,y:A. - (\forall x,y:A. decidable (x=y)) \to - \forall p,q:x=y. p=q. -intros. -letin nu \def - (\lambda x,y. \lambda p:x=y. - match H x y with - [ (or_introl p') \Rightarrow p' - | (or_intror K) \Rightarrow False_ind ? (K p) ]). -cut - (\forall q:x=y. - eq_ind ? ? (\lambda z. z=y) (nu ? ? q) ? (nu ? ? (refl_eq ? x)) - = q). -focus 8. - clear q; clear p. - intro. - apply (eq_rect' ? ? ? ? ? q); - fold simplify (nu ? ? (refl_eq ? x)). - generalize in match (nu ? ? (refl_eq ? x)); intro. - apply - (eq_rect' A x - (\lambda y. \lambda u. - eq_ind A x (\lambda a.a=y) u y u = refl_eq ? y) - ? x H1). - reflexivity. -unfocus. -rewrite < (Hcut p); fold simplify (nu ? ? p). -rewrite < (Hcut q); fold simplify (nu ? ? q). -apply (Or_ind' (x=x) (x \neq x) - (\lambda p:decidable (x=x). eq_ind A x (\lambda z.z=y) (nu x y p) x - ([\lambda H1.eq A x x] - match p with - [(or_introl p') \Rightarrow p' - |(or_intror K) \Rightarrow False_ind (x=x) (K (refl_eq A x))]) = - eq_ind A x (\lambda z.z=y) (nu x y q) x - ([\lambda H1.eq A x x] - match p with - [(or_introl p') \Rightarrow p' - |(or_intror K) \Rightarrow False_ind (x=x) (K (refl_eq A x))])) - ? ? (H x x)). -intro; simplify; reflexivity. -intro q; elim (q (refl_eq ? x)). -qed. -*) +ndefinition R3 : + ∀T0:Type[0]. + ∀a0:T0. + ∀T1:∀x0:T0. a0=x0 → Type[0]. + ∀a1:T1 a0 (refl ? a0). + ∀T2:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0. R1 ?? T1 a1 ? p0 = x1 → Type[0]. + ∀a2:T2 a0 (refl ? a0) a1 (refl ? a1). + ∀T3:∀x0:T0. ∀p0:a0=x0. ∀x1:T1 x0 p0.∀p1:R1 ?? T1 a1 ? p0 = x1. + ∀x2:T2 x0 p0 x1 p1.R2 ???? T2 a2 x0 p0 ? p1 = x2 → Type[0]. + ∀a3:T3 a0 (refl ? a0) a1 (refl ? a1) a2 (refl ? a2). + ∀b0:T0. + ∀e0:a0 = b0. + ∀b1: T1 b0 e0. + ∀e1:R1 ?? T1 a1 ? e0 = b1. + ∀b2: T2 b0 e0 b1 e1. + ∀e2:R2 ???? T2 a2 b0 e0 ? e1 = b2. + T3 b0 e0 b1 e1 b2 e2. +#T0;#a0;#T1;#a1;#T2;#a2;#T3;#a3;#b0;#e0;#b1;#e1;#b2;#e2; +napply (eq_rect_Type0 ????? e2); +napply (R2 ?? ? ???? e0 ? e1); +napply a3; +nqed. -(* -theorem a:\forall x.x=x\land True. -[ -2:intros; - split; - [ - exact (refl_eq Prop x); - | - exact I; - ] -1: - skip -] -qed. -*) *) +ndefinition R4 : + ∀T0:Type[0]. + ∀a0:T0. + ∀T1:∀x0:T0. eq T0 a0 x0 → Type[0]. + ∀a1:T1 a0 (refl T0 a0). + ∀T2:∀x0:T0. ∀p0:eq (T0 …) a0 x0. ∀x1:T1 x0 p0.eq (T1 …) (R1 T0 a0 T1 a1 x0 p0) x1 → Type[0]. + ∀a2:T2 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1). + ∀T3:∀x0:T0. ∀p0:eq (T0 …) a0 x0. ∀x1:T1 x0 p0.∀p1:eq (T1 …) (R1 T0 a0 T1 a1 x0 p0) x1. + ∀x2:T2 x0 p0 x1 p1.eq (T2 …) (R2 T0 a0 T1 a1 T2 a2 x0 p0 x1 p1) x2 → Type[0]. + ∀a3:T3 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1) + a2 (refl (T2 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1)) a2). + ∀T4:∀x0:T0. ∀p0:eq (T0 …) a0 x0. ∀x1:T1 x0 p0.∀p1:eq (T1 …) (R1 T0 a0 T1 a1 x0 p0) x1. + ∀x2:T2 x0 p0 x1 p1.∀p2:eq (T2 …) (R2 T0 a0 T1 a1 T2 a2 x0 p0 x1 p1) x2. + ∀x3:T3 x0 p0 x1 p1 x2 p2.∀p3:eq (T3 …) (R3 T0 a0 T1 a1 T2 a2 T3 a3 x0 p0 x1 p1 x2 p2) x3. + Type[0]. + ∀a4:T4 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1) + a2 (refl (T2 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1)) a2) + a3 (refl (T3 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1) + a2 (refl (T2 a0 (refl T0 a0) a1 (refl (T1 a0 (refl T0 a0)) a1)) a2)) + a3). + ∀b0:T0. + ∀e0:eq (T0 …) a0 b0. + ∀b1: T1 b0 e0. + ∀e1:eq (T1 …) (R1 T0 a0 T1 a1 b0 e0) b1. + ∀b2: T2 b0 e0 b1 e1. + ∀e2:eq (T2 …) (R2 T0 a0 T1 a1 T2 a2 b0 e0 b1 e1) b2. + ∀b3: T3 b0 e0 b1 e1 b2 e2. + ∀e3:eq (T3 …) (R3 T0 a0 T1 a1 T2 a2 T3 a3 b0 e0 b1 e1 b2 e2) b3. + T4 b0 e0 b1 e1 b2 e2 b3 e3. +#T0;#a0;#T1;#a1;#T2;#a2;#T3;#a3;#T4;#a4;#b0;#e0;#b1;#e1;#b2;#e2;#b3;#e3; +napply (eq_rect_Type0 ????? e3); +napply (R3 ????????? e0 ? e1 ? e2); +napply a4; +nqed. +naxiom streicherK : ∀T:Type[0].∀t:T.∀P:t = t → Type[2].P (refl ? t) → ∀p.P p. \ No newline at end of file