X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Flogic%2Fcprop.ma;h=e7ecf01ad5f47a49885a477b1b96138de00d893e;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=9b352d4a4112dcab86e9b922343b67414301fa63;hpb=6f2f5039ef719f60ebcf24d7ee17c83eac6cc635;p=helm.git diff --git a/helm/software/matita/nlibrary/logic/cprop.ma b/helm/software/matita/nlibrary/logic/cprop.ma index 9b352d4a4..e7ecf01ad 100644 --- a/helm/software/matita/nlibrary/logic/cprop.ma +++ b/helm/software/matita/nlibrary/logic/cprop.ma @@ -50,9 +50,9 @@ nqed. unification hint 0 ≔ A,B:CProp[0]; T ≟ CPROP, - MM ≟ mk_unary_morphism1 … - (λX.mk_unary_morphism1 … (And X) (prop11 … (and_morphism X))) - (prop11 … and_morphism) + MM ≟ mk_unary_morphism1 ?? + (λX.mk_unary_morphism1 ?? (And X) (prop11 ?? (fun11 ?? and_morphism X))) + (prop11 ?? and_morphism) (*-------------------------------------------------------------*) ⊢ fun11 T T (fun11 T (unary_morphism1_setoid1 T T) MM A) B ≡ And A B. @@ -76,7 +76,7 @@ nqed. unification hint 0 ≔ A,B:CProp[0]; T ≟ CPROP, MM ≟ mk_unary_morphism1 … - (λX.mk_unary_morphism1 … (Or X) (prop11 … (or_morphism X))) + (λX.mk_unary_morphism1 … (Or X) (prop11 … (fun11 ?? or_morphism X))) (prop11 … or_morphism) (*-------------------------------------------------------------*) ⊢ fun11 T T (fun11 T (unary_morphism1_setoid1 T T) MM A) B ≡ Or A B. @@ -121,33 +121,35 @@ nlemma Ex_morphism : ∀S:setoid.((setoid1_of_setoid S) ⇒_1 CProp[0]) ⇒_1 CP #S; @(λP: (setoid1_of_setoid S) ⇒_1 CProp[0].Ex S P); #P Q E; @; *; #x Px; @x; ncases (E x x #); /2/; nqed. -unification hint 0 ≔ S : setoid, P : (setoid1_of_setoid S) ⇒_1 CProp[0]; +unification hint 0 ≔ S : setoid, P : (setoid1_of_setoid S) ⇒_1 CPROP; A ≟ unary_morphism1_setoid1 (setoid1_of_setoid S) CPROP, B ≟ CPROP, - M ≟ mk_unary_morphism1 ?? (λP: (setoid1_of_setoid S) ⇒_1 CProp[0].Ex S P) - (prop11 ?? (Ex_morphism S)) + M ≟ mk_unary_morphism1 ?? + (λP: (setoid1_of_setoid S) ⇒_1 CPROP.Ex (carr S) (fun11 ?? P)) + (prop11 ?? (Ex_morphism S)) (*------------------------*)⊢ - fun11 A B M P ≡ Ex S (fun11 (setoid1_of_setoid S) CPROP P). + fun11 A B M P ≡ Ex (carr S) (fun11 (setoid1_of_setoid S) CPROP P). nlemma Ex_morphism_eta : ∀S:setoid.((setoid1_of_setoid S) ⇒_1 CProp[0]) ⇒_1 CProp[0]. #S; @(λP: (setoid1_of_setoid S) ⇒_1 CProp[0].Ex S (λx.P x)); #P Q E; @; *; #x Px; @x; ncases (E x x #); /2/; nqed. -unification hint 0 ≔ S : setoid, P : (setoid1_of_setoid S) ⇒_1 CProp[0]; +unification hint 0 ≔ S : setoid, P : (setoid1_of_setoid S) ⇒_1 CPROP; A ≟ unary_morphism1_setoid1 (setoid1_of_setoid S) CPROP, B ≟ CPROP, - M ≟ mk_unary_morphism1 ?? (λP: (setoid1_of_setoid S) ⇒_1 CProp[0].Ex S (λx.P x)) - (prop11 ?? (Ex_morphism_eta S)) + M ≟ mk_unary_morphism1 ?? + (λP: (setoid1_of_setoid S) ⇒_1 CPROP.Ex (carr S) (λx.fun11 ?? P x)) + (prop11 ?? (Ex_morphism_eta S)) (*------------------------*)⊢ - fun11 A B M P ≡ Ex S (λx.fun11 (setoid1_of_setoid S) CPROP P x). + fun11 A B M P ≡ Ex (carr S) (λx.fun11 (setoid1_of_setoid S) CPROP P x). nlemma Ex_setoid : ∀S:setoid.((setoid1_of_setoid S) ⇒_1 CPROP) → setoid. #T P; @ (Ex T (λx:T.P x)); @; ##[ #H1 H2; napply True |##*: //; ##] nqed. -unification hint 0 ≔ T,P ; +unification hint 0 ≔ T : setoid,P ; S ≟ (Ex_setoid T P) (*---------------------------*) ⊢ - Ex T (λx:T.P x) ≡ carr S. + Ex (carr T) (λx:carr T.fun11 ?? P x) ≡ carr S. (* couts how many Ex we are traversing *) ninductive counter : Type[0] ≝