From 44c1079dabf1d3c0b69d0155ddbaea8627ec901c Mon Sep 17 00:00:00 2001 From: matitaweb Date: Thu, 23 Feb 2012 10:46:38 +0000 Subject: [PATCH] commit by user andrea --- weblib/basics/logic.ma | 178 +++++++++++++++++++++--------------- weblib/basics/pts.ma | 3 +- weblib/basics/types.ma | 2 +- weblib/tutorial/chapter1.ma | 16 ++-- weblib/tutorial/chapter4.ma | 41 ++++++--- 5 files changed, 143 insertions(+), 97 deletions(-) diff --git a/weblib/basics/logic.ma b/weblib/basics/logic.ma index aeee17a73..8f981d469 100644 --- a/weblib/basics/logic.ma +++ b/weblib/basics/logic.ma @@ -1,4 +1,4 @@ -(* + (* ||M|| This file is part of HELM, an Hypertextual, Electronic ||A|| Library of Mathematics, developed at the Computer Science ||T|| Department of the University of Bologna, Italy. @@ -10,54 +10,71 @@ V_______________________________________________________________ *) include "basics/pts.ma". -(*include "hints_declaration.ma".*) +include "hints_declaration.ma". (* propositional equality *) -inductive eq (A:Type[1]) (x:A) : A → Prop ≝ +inductive eq (A:Type[2]) (x:A) : A → Prop ≝ refl: eq A x x. - + interpretation "leibnitz's equality" 'eq t x y = (eq t x y). +interpretation "leibniz reflexivity" 'refl = refl. lemma eq_rect_r: - ∀A.∀a,x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.∀P: - ∀x:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a → Type[2]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → P x p. + ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[3]. P a (refl A a) → P x p. #A #a #x #p (cases p) // qed. lemma eq_ind_r : - ∀A.∀a.∀P: ∀x:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a → Prop. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → - ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. - #A #a #P #p #x0 #p0; @(a href="cic:/matita/basics/logic/eq_rect_r.def(1)"eq_rect_r/a ? ? ? p0) //; qed. + ∀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; @(eq_rect_r ? ? ? p0) //; qed. + +lemma eq_rect_Type0_r: + ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + #A #a #P #H #x #p lapply H lapply P + cases p; //; qed. + +lemma eq_rect_Type1_r: + ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[1]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + #A #a #P #H #x #p lapply H lapply P + cases p; //; qed. lemma eq_rect_Type2_r: - ∀A.∀a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[2]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → - ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. - #A #a #P #H #x #p (generalize in match H) (generalize in match P) + ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[2]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + #A #a #P #H #x #p lapply H lapply P + cases p; //; qed. + +lemma eq_rect_Type3_r: + ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[3]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. + #A #a #P #H #x #p lapply H lapply P cases p; //; qed. -theorem rewrite_l: ∀A:Type[1].∀x.∀P:A → Type[1]. P x → ∀y. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → P y. +theorem rewrite_l: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. x = y → P y. #A #x #P #Hx #y #Heq (cases Heq); //; qed. -theorem sym_eq: ∀A.∀x,y:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. -#A #x #y #Heq @(a href="cic:/matita/basics/logic/rewrite_l.def(1)"rewrite_l/a A x (λz.za title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax)); //; qed. +theorem sym_eq: ∀A.∀x,y:A. x = y → y = x. +#A #x #y #Heq @(rewrite_l A x (λz.z=x)) // qed. -theorem rewrite_r: ∀A:Type[1].∀x.∀P:A → Type[1]. P x → ∀y. y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x → P y. -#A #x #P #Hx #y #Heq (cases (a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a ? ? ? Heq)); //; qed. +theorem rewrite_r: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. y = x → P y. +#A #x #P #Hx #y #Heq (cases (sym_eq ? ? ? Heq)); //; qed. -theorem eq_coerc: ∀A,B:Type[0].A→(Aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aB)→B. +theorem eq_coerc: ∀A,B:Type[0].A→(A=B)→B. #A #B #Ha #Heq (elim Heq); //; qed. -theorem trans_eq : ∀A.∀x,y,z:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z → x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z. +theorem trans_eq : ∀A.∀x,y,z:A. x = y → y = z → x = z. #A #x #y #z #H1 #H2 >H1; //; qed. -theorem eq_f: ∀A,B.∀f:A→B.∀x,y:A. xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay → f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y. +theorem eq_f: ∀A,B.∀f:A→B.∀x,y:A. x=y → f x = f y. #A #B #f #x #y #H >H; //; qed. (* deleterio per auto? *) theorem eq_f2: ∀A,B,C.∀f:A→B→C. -∀x1,x2:A.∀y1,y2:B. x1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax2 → y1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay2 → f x1 y1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x2 y2. +∀x1,x2:A.∀y1,y2:B. x1=x2 → y1=y2 → f x1 y1 = f x2 y2. #A #B #C #f #x1 #x2 #y1 #y2 #E1 #E2 >E1; >E2; //; qed. +lemma eq_f3: ∀A,B,C,D.∀f:A→B→C->D. +∀x1,x2:A.∀y1,y2:B. ∀z1,z2:C. x1=x2 → y1=y2 → z1=z2 → f x1 y1 z1 = f x2 y2 z2. +#A #B #C #D #f #x1 #x2 #y1 #y2 #z1 #z2 #E1 #E2 #E3 >E1; >E2; >E3 //; qed. + (* hint to genereric equality definition eq_equality: equality ≝ mk_equality eq refl rewrite_l rewrite_r. @@ -80,12 +97,11 @@ inductive False: Prop ≝ . λA. A → False. *) inductive Not (A:Prop): Prop ≝ -nmk: (A → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a) → Not A. - +nmk: (A → False) → Not A. interpretation "logical not" 'not x = (Not x). -theorem absurd : ∀A:Prop. A → a title="logical not" href="cic:/fakeuri.def(1)"¬/aA → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. +theorem absurd : ∀A:Prop. A → ¬A → False. #A #H #Hn (elim Hn); /2/; qed. (* @@ -93,13 +109,13 @@ ntheorem absurd : ∀ A,C:Prop. A → ¬A → C. #A; #C; #H; #Hn; nelim (Hn H). nqed. *) -theorem not_to_not : ∀A,B:Prop. (A → B) → a title="logical not" href="cic:/fakeuri.def(1)"¬/aB →a title="logical not" href="cic:/fakeuri.def(1)"¬/aA. +theorem not_to_not : ∀A,B:Prop. (A → B) → ¬B →¬A. /4/; qed. (* inequality *) interpretation "leibnitz's non-equality" 'neq t x y = (Not (eq t x y)). -theorem sym_not_eq: ∀A.∀x,y:A. x a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a y → y a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a x. +theorem sym_not_eq: ∀A.∀x,y:A. x ≠ y → y ≠ x. /3/; qed. (* and *) @@ -108,10 +124,10 @@ inductive And (A,B:Prop) : Prop ≝ interpretation "logical and" 'and x y = (And x y). -theorem proj1: ∀A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → A. +theorem proj1: ∀A,B:Prop. A ∧ B → A. #A #B #AB (elim AB) //; qed. -theorem proj2: ∀ A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → B. +theorem proj2: ∀ A,B:Prop. A ∧ B → B. #A #B #AB (elim AB) //; qed. (* or *) @@ -122,7 +138,7 @@ inductive Or (A,B:Prop) : Prop ≝ interpretation "logical or" 'or x y = (Or x y). definition decidable : Prop → Prop ≝ -λ A:Prop. A a title="logical or" href="cic:/fakeuri.def(1)"∨/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a A. +λ A:Prop. A ∨ ¬ A. (* exists *) inductive ex (A:Type[0]) (P:A → Prop) : Prop ≝ @@ -135,10 +151,10 @@ inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝ (* iff *) definition iff := - λ A,B. (A → B) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (B → A). + λ A,B. (A → B) ∧ (B → A). interpretation "iff" 'iff a b = (iff a b). - + lemma iff_sym: ∀A,B. A ↔ B → B ↔ A. #A #B * /3/ qed. @@ -159,88 +175,98 @@ lemma iff_or_l: ∀A,B,C. A ↔ B → C ∨ A ↔ C ∨ B. lemma iff_or_r: ∀A,B,C. A ↔ B → A ∨ C ↔ B ∨ C. #A #B #C * #H1 #H2 % * /3/ qed. + (* cose per destruct: da rivedere *) definition R0 ≝ λT:Type[0].λt:T.t. -definition R1 ≝ a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a. +definition R1 ≝ eq_rect_Type0. -(* useless stuff *) +(* used for lambda-delta *) definition R2 : ∀T0:Type[0]. ∀a0:T0. - ∀T1:∀x0:T0. a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0 → Type[0]. - ∀a1:T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a0). - ∀T2:∀x0:T0. ∀p0:a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0. ∀x1:T1 x0 p0. a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? T1 a1 ? p0 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x1 → Type[0]. - ∀a2:T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a1). + ∀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 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b0. + ∀e0:a0 = b0. ∀b1: T1 b0 e0. - ∀e1:a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? T1 a1 ? e0 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b1. + ∀e1:R1 ?? T1 a1 ? e0 = b1. T2 b0 e0 b1 e1. #T0 #a0 #T1 #a1 #T2 #a2 #b0 #e0 #b1 #e1 -@(a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a ????? e1) -@(a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? ? ?? e0) +@(eq_rect_Type0 ????? e1) +@(R1 ?? ? ?? e0) @a2 qed. definition R3 : ∀T0:Type[0]. ∀a0:T0. - ∀T1:∀x0:T0. a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0 → Type[0]. - ∀a1:T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a0). - ∀T2:∀x0:T0. ∀p0:a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0. ∀x1:T1 x0 p0. a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? T1 a1 ? p0 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x1 → Type[0]. - ∀a2:T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a1). - ∀T3:∀x0:T0. ∀p0:a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0. ∀x1:T1 x0 p0.∀p1:a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? T1 a1 ? p0 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x1. - ∀x2:T2 x0 p0 x1 p1.a href="cic:/matita/basics/logic/R2.def(3)"R2/a ???? T2 a2 x0 p0 ? p1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x2 → Type[0]. - ∀a3:T3 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a1) a2 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? a2). + ∀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 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b0. + ∀e0:a0 = b0. ∀b1: T1 b0 e0. - ∀e1:a href="cic:/matita/basics/logic/R1.def(2)"R1/a ?? T1 a1 ? e0 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b1. + ∀e1:R1 ?? T1 a1 ? e0 = b1. ∀b2: T2 b0 e0 b1 e1. - ∀e2:a href="cic:/matita/basics/logic/R2.def(3)"R2/a ???? T2 a2 b0 e0 ? e1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b2. + ∀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 -@(a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a ????? e2) -@(a href="cic:/matita/basics/logic/R2.def(3)"R2/a ?? ? ???? e0 ? e1) +@(eq_rect_Type0 ????? e2) +@(R2 ?? ? ???? e0 ? e1) @a3 qed. definition R4 : ∀T0:Type[0]. ∀a0:T0. - ∀T1:∀x0:T0. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a T0 a0 x0 → Type[0]. - ∀a1:T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0). - ∀T2:∀x0:T0. ∀p0:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T0 …) a0 x0. ∀x1:T1 x0 p0.a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T1 …) (a href="cic:/matita/basics/logic/R1.def(2)"R1/a T0 a0 T1 a1 x0 p0) x1 → Type[0]. - ∀a2:T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1). - ∀T3:∀x0:T0. ∀p0:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T0 …) a0 x0. ∀x1:T1 x0 p0.∀p1:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T1 …) (a href="cic:/matita/basics/logic/R1.def(2)"R1/a T0 a0 T1 a1 x0 p0) x1. - ∀x2:T2 x0 p0 x1 p1.a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T2 …) (a href="cic:/matita/basics/logic/R2.def(3)"R2/a T0 a0 T1 a1 T2 a2 x0 p0 x1 p1) x2 → Type[0]. - ∀a3:T3 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1) - a2 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1)) a2). - ∀T4:∀x0:T0. ∀p0:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T0 …) a0 x0. ∀x1:T1 x0 p0.∀p1:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T1 …) (a href="cic:/matita/basics/logic/R1.def(2)"R1/a T0 a0 T1 a1 x0 p0) x1. - ∀x2:T2 x0 p0 x1 p1.∀p2:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T2 …) (a href="cic:/matita/basics/logic/R2.def(3)"R2/a T0 a0 T1 a1 T2 a2 x0 p0 x1 p1) x2. - ∀x3:T3 x0 p0 x1 p1 x2 p2.∀p3:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T3 …) (a href="cic:/matita/basics/logic/R3.def(4)"R3/a T0 a0 T1 a1 T2 a2 T3 a3 x0 p0 x1 p1 x2 p2) x3. + ∀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 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1) - a2 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1)) a2) - a3 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T3 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1) - a2 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T2 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0) a1 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a (T1 a0 (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a T0 a0)) a1)) a2)) + ∀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:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T0 …) a0 b0. + ∀e0:eq (T0 …) a0 b0. ∀b1: T1 b0 e0. - ∀e1:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T1 …) (a href="cic:/matita/basics/logic/R1.def(2)"R1/a T0 a0 T1 a1 b0 e0) b1. + ∀e1:eq (T1 …) (R1 T0 a0 T1 a1 b0 e0) b1. ∀b2: T2 b0 e0 b1 e1. - ∀e2:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T2 …) (a href="cic:/matita/basics/logic/R2.def(3)"R2/a T0 a0 T1 a1 T2 a2 b0 e0 b1 e1) b2. + ∀e2:eq (T2 …) (R2 T0 a0 T1 a1 T2 a2 b0 e0 b1 e1) b2. ∀b3: T3 b0 e0 b1 e1 b2 e2. - ∀e3:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a (T3 …) (a href="cic:/matita/basics/logic/R3.def(4)"R3/a T0 a0 T1 a1 T2 a2 T3 a3 b0 e0 b1 e1 b2 e2) b3. + ∀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 -@(a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a ????? e3) -@(a href="cic:/matita/basics/logic/R3.def(4)"R3/a ????????? e0 ? e1 ? e2) +@(eq_rect_Type0 ????? e3) +@(R3 ????????? e0 ? e1 ? e2) @a4 qed. -(* TODO concrete definition by means of proof irrelevance *) -axiom streicherK : ∀T:Type[1].∀t:T.∀P:t a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a t → Type[2].P (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? t) → ∀p.P p. \ No newline at end of file +definition eqProp ≝ λA:Prop.eq A. + +(* Example to avoid indexing and the consequential creation of ill typed + terms during paramodulation *) +example lemmaK : ∀A.∀x:A.∀h:x=x. eqProp ? h (refl A x). +#A #x #h @(refl ? h: eqProp ? ? ?). +qed. + +theorem streicherK : ∀T:Type[2].∀t:T.∀P:t = t → Type[3].P (refl ? t) → ∀p.P p. + #T #t #P #H #p >(lemmaK T t p) @H +qed. diff --git a/weblib/basics/pts.ma b/weblib/basics/pts.ma index b7d9a5da9..c0ec34b71 100644 --- a/weblib/basics/pts.ma +++ b/weblib/basics/pts.ma @@ -18,4 +18,5 @@ universe constraint Type[0] < Type[1]. universe constraint Type[1] < Type[2]. universe constraint Type[2] < Type[3]. universe constraint Type[3] < Type[4]. - +universe constraint Type[4] < Type[5]. + \ No newline at end of file diff --git a/weblib/basics/types.ma b/weblib/basics/types.ma index 0fd4eb588..586ce31ce 100644 --- a/weblib/basics/types.ma +++ b/weblib/basics/types.ma @@ -58,4 +58,4 @@ inductive option (A:Type[0]) : Type[0] ≝ inductive Sig (A:Type[0]) (f:A→Type[0]) : Type[0] ≝ dp: ∀a:A.(f a)→Sig A f. -interpretation "Sigma" 'sigma x = (Sig ? x). +interpretation "Sigma" 'sigma x = (Sig ? x). \ No newline at end of file diff --git a/weblib/tutorial/chapter1.ma b/weblib/tutorial/chapter1.ma index b21fd3e17..9043bdad7 100644 --- a/weblib/tutorial/chapter1.ma +++ b/weblib/tutorial/chapter1.ma @@ -243,7 +243,7 @@ need a few more applications to handle reachability, and side conditions. The magic number to let automation work is, in this case, 9. *) lemma problem: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. -normalize /9/ qed. +normalize /span class="autotactic"9span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a, a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"move_wolf/a, a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a, a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ qed. (* Let us now try to derive the proof in a more interactive way. Of course, we expect to need several moves to transfer all items from a bank to the other, so @@ -304,7 +304,7 @@ requires /2/ since move_goat opens an additional subgoal. By applying "]" we refocus on the skipped goal, going back to a situation similar to the one we started with. *) - | /2/ ] + | /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ ] (* Let us perform the next step, namely moving back the boat, in a sligtly different way. The more operation expects as second argument the new @@ -319,7 +319,7 @@ the system. *) trivial. Wespan style="font-family: Verdana,sans-serif;" /spancan just apply automation to all of them, and it will close the two trivial goals. *) -/2/ +/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ (* Let us come to the next step, that consists in moving the wolf. Suppose that instead of specifying the next intermediate state, we prefer to specify the next @@ -338,12 +338,12 @@ be arbitrary). The simplest way to proceed is to focus on the bank, that is the fourth subgoal, and explicitly instatiate it. Instead of repeatedly using "|", we can perform focusing by typing "4:" as described by the following command. *) -[4: @a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a] /2/ +[4: @a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a] /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ (* Alternatively, we can directly instantiate the bank into the move. Let us complete the proof in this, very readable way. *) -@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /2/ -@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a ?? a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a … )) /2/ -@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a ??? a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /2/ -@a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a /2/ qed. \ No newline at end of file +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a ?? a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a ??? a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ qed. \ No newline at end of file diff --git a/weblib/tutorial/chapter4.ma b/weblib/tutorial/chapter4.ma index b358f8232..8457aa484 100644 --- a/weblib/tutorial/chapter4.ma +++ b/weblib/tutorial/chapter4.ma @@ -1,44 +1,63 @@ +(* In this Chapter we shall develop a naif theory of sets represented as characteristic +predicates over some universe codeA/code, that is as objects of type A→Prop *) + include "basics/logic.ma". -(**** a subset of A is just an object of type A→Prop ****) +(**** For instance the empty set is defined by the always function predicate *) -definition empty_set ≝ λA:Type[0].λa:A.False. +definition empty_set ≝ λA:Type[0].λa:A.a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. notation "\emptyv" non associative with precedence 90 for @{'empty_set}. interpretation "empty set" 'empty_set = (empty_set ?). -definition singleton ≝ λA.λx,a:A.x=a. +(* Similarly, a singleton set contaning containing an element a, is defined +by by the characteristic function asserting equality with a *) + +definition singleton ≝ λA.λx,a:A.xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa. (* notation "{x}" non associative with precedence 90 for @{'sing_lang $x}. *) interpretation "singleton" 'singl x = (singleton ? x). -definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a ∨ Q a. +(* The operations of union, intersection, complement and substraction +are easily defined in terms of the propositional connectives of dijunction, +conjunction and negation *) + +definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a a title="logical or" href="cic:/fakeuri.def(1)"∨/a Q a. interpretation "union" 'union a b = (union ? a b). -definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a ∧ Q a. +definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a a title="logical and" href="cic:/fakeuri.def(1)"∧/a Q a. interpretation "intersection" 'intersects a b = (intersection ? a b). -definition complement ≝ λU:Type[0].λA:U → Prop.λw.¬ A w. +definition complement ≝ λU:Type[0].λA:U → Prop.λw.a title="logical not" href="cic:/fakeuri.def(1)"¬/a A w. interpretation "complement" 'not a = (complement ? a). -definition substraction := λU:Type[0].λA,B:U → Prop.λw.A w ∧ ¬ B w. +definition substraction := λU:Type[0].λA,B:U → Prop.λw.A w a title="logical and" href="cic:/fakeuri.def(1)"∧/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a B w. interpretation "substraction" 'minus a b = (substraction ? a b). +(* Finally, we use implication to define the inclusion relation between +sets *) + definition subset: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA,P,Q.∀a:A.(P a → Q a). interpretation "subset" 'subseteq a b = (subset ? a b). -(* extensional equality *) -definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a ↔ Q a. +(* Two sets are equals if and only if they have the same elements, that is, +if the two characteristic functions are extensionally equivalent: *) + +definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a a title="iff" href="cic:/fakeuri.def(1)"↔/a Q a. notation "A =1 B" non associative with precedence 45 for @{'eqP $A $B}. interpretation "extensional equality" 'eqP a b = (eqP ? a b). +(* This notion of equality is different from the intensional equality of +fucntions, hence we have to prove the usual properties: *) + lemma eqP_sym: ∀U.∀A,B:U →Prop. - A =1 B → B =1 A. + A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B → B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #U #A #B #eqAB #a @iff_sym @eqAB qed. lemma eqP_trans: ∀U.∀A,B,C:U →Prop. A =1 B → B =1 C → A =1 C. #U #A #B #C #eqAB #eqBC #a @iff_trans // qed. -lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. +(* For the same reason, we must also prove that all the operations we have + lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. A =1 C → A ∪ B =1 C ∪ B. #U #A #B #C #eqAB #a @iff_or_r @eqAB qed. -- 2.39.2