X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Flogic%2Fcprop.ma;h=0302264c690a7575a63cce643fd9940933d23a12;hb=06bf8cbce7fd66562954c002d4058fb18fa366cb;hp=e18811da416feab641aab7f79f6b2324fe5ee9b7;hpb=bc4cc41e813fc850a7f981cd60ba22e14485b7d1;p=helm.git diff --git a/helm/software/matita/nlibrary/logic/cprop.ma b/helm/software/matita/nlibrary/logic/cprop.ma index e18811da4..0302264c6 100644 --- a/helm/software/matita/nlibrary/logic/cprop.ma +++ b/helm/software/matita/nlibrary/logic/cprop.ma @@ -52,15 +52,13 @@ ndefinition and_morphism: binary_morphism1 CPROP CPROP CPROP. | napply (H4 K2)]##] nqed. -unification hint 0 ≔ A,B ⊢ fun21 … and_morphism A B ≡ And A B. +unification hint 0 ≔ A,B ⊢ fun21 … (mk_binary_morphism1 … And (prop21 … and_morphism)) A B ≡ And A B. (*nlemma test: ∀A,A',B: CProp[0]. A=A' → (B ∨ A) = B → (B ∧ A) ∧ B. #A; #A'; #B; #H1; #H2; napply (. ((#‡H1)‡H2^-1)); nnormalize; nqed.*) -(*interpretation "and_morphism" 'and a b = (fun21 ??? and_morphism a b).*) - ndefinition or_morphism: binary_morphism1 CPROP CPROP CPROP. napply mk_binary_morphism1 [ napply Or @@ -72,9 +70,7 @@ ndefinition or_morphism: binary_morphism1 CPROP CPROP CPROP. | napply (H4 H)]##] nqed. -unification hint 0 ≔ A,B ⊢ fun21 … or_morphism A B ≡ Or A B. - -(*interpretation "or_morphism" 'or a b = (fun21 ??? or_morphism a b).*) +unification hint 0 ≔ A,B ⊢ fun21 … (mk_binary_morphism1 … Or (prop21 … or_morphism)) A B ≡ Or A B. ndefinition if_morphism: binary_morphism1 CPROP CPROP CPROP. napply mk_binary_morphism1 @@ -82,4 +78,4 @@ ndefinition if_morphism: binary_morphism1 CPROP CPROP CPROP. | #a; #a'; #b; #b'; #H1; #H2; napply mk_iff; #H; #w [ napply (if … H2); napply H; napply (fi … H1); nassumption | napply (fi … H2); napply H; napply (if … H1); nassumption]##] -nqed. \ No newline at end of file +nqed.