X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Flib%2Fbasics%2Flogic.ma;h=4b90f75181300bdaced20673e08e5d6420f4fbc1;hb=f5fa6554c93bec72a5bb098c0a2991fe294883b6;hp=8f981d46924556ee885407a1320d147ef07aaad0;hpb=2225846a2ee76026bef7c3e8eacc42e86c3ae6a2;p=helm.git diff --git a/matita/matita/lib/basics/logic.ma b/matita/matita/lib/basics/logic.ma index 8f981d469..4b90f7518 100644 --- a/matita/matita/lib/basics/logic.ma +++ b/matita/matita/lib/basics/logic.ma @@ -147,7 +147,13 @@ inductive ex (A:Type[0]) (P:A → Prop) : Prop ≝ interpretation "exists" 'exists x = (ex ? x). inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝ - ex_intro2: ∀ x:A. P x → Q x → ex2 A P Q. + ex2_intro: ∀ x:A. P x → Q x → ex2 A P Q. + +interpretation "exists on two predicates" 'exists2 x1 x2 = (ex2 ? x1 x2). + +lemma ex2_commute: ∀A0. ∀P0,P1:A0→Prop. (∃∃x0. P0 x0 & P1 x0) → ∃∃x0. P1 x0 & P0 x0. +#A0 #P0 #P1 * /2 width=3/ +qed-. (* iff *) definition iff := @@ -263,10 +269,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-.