X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Flogic.ma;h=70e743e993ac96c12cdbf7a4a56cdc23bc8280b6;hb=85a42e4a2a4c62818b6a98eff545e58ceb8770a4;hp=c80fa31f55c5469a97212e3c49431360e3e805fa;hpb=94188b0cbaff6340464d90cc13ee246ea7ec3284;p=helm.git diff --git a/matita/matita/lib/basics/logic.ma b/matita/matita/lib/basics/logic.ma index c80fa31f5..70e743e99 100644 --- a/matita/matita/lib/basics/logic.ma +++ b/matita/matita/lib/basics/logic.ma @@ -153,7 +153,28 @@ inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝ definition iff := λ A,B. (A → B) ∧ (B → A). -interpretation "iff" 'iff a b = (iff a b). +interpretation "iff" 'iff a b = (iff a b). + +lemma iff_sym: ∀A,B. A ↔ B → B ↔ A. +#A #B * /3/ qed. + +lemma iff_trans:∀A,B,C. A ↔ B → B ↔ C → A ↔ C. +#A #B #C * #H1 #H2 * #H3 #H4 % /3/ qed. + +lemma iff_not: ∀A,B. A ↔ B → ¬A ↔ ¬B. +#A #B * #H1 #H2 % /3/ qed. + +lemma iff_and_l: ∀A,B,C. A ↔ B → C ∧ A ↔ C ∧ B. +#A #B #C * #H1 #H2 % * /3/ qed. + +lemma iff_and_r: ∀A,B,C. A ↔ B → A ∧ C ↔ B ∧ C. +#A #B #C * #H1 #H2 % * /3/ qed. + +lemma iff_or_l: ∀A,B,C. A ↔ B → C ∨ A ↔ C ∨ B. +#A #B #C * #H1 #H2 % * /3/ qed. + +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 *) @@ -242,10 +263,10 @@ 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). +lemma lemmaK : ∀A.∀x:A.∀h:x=x. eqProp ? h (refl A x). #A #x #h @(refl ? h: eqProp ? ? ?). -qed. +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. +qed-.