X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fng_commands.ma;h=b597156499e152dc8533c6e17bf7e2a870f9800a;hb=7233348f05485c2ee317df9c3407cf1ce7e56927;hp=6a65ea039a68774d0ec0e759b04e1cdf36e7c3a6;hpb=070b11daefc90ecc20ebee73acc550aeac1c627b;p=helm.git diff --git a/helm/software/matita/tests/ng_commands.ma b/helm/software/matita/tests/ng_commands.ma index 6a65ea039..b59715649 100644 --- a/helm/software/matita/tests/ng_commands.ma +++ b/helm/software/matita/tests/ng_commands.ma @@ -12,39 +12,33 @@ (* *) (**************************************************************************) +(* ATTENZIONE: si puo' eseguire solo con MatitaC e dopo aver scelto + Debug → always show all disambiguation errors (that, moreover, slows + down the library a lot) *) + include "nat/plus.ma". ndefinition thesis0: ∀A:Type.Type ≝ λA. A → A. -alias id "thesis0" = "cic:/matita/tests/ng_commands/thesis0.con". - ndefinition thesis: ∀A:Type.Type ≝ λA. ?. napply (A → A); nqed. -alias id "thesis" = "cic:/matita/tests/ng_commands/thesis.con". - ntheorem foo: ∀A:Type.thesis A.#A;#x;napply x; nqed. -alias id "foo" = "cic:/matita/tests/ng_commands/foo.con". - ntheorem goo: ∀A:Type.A → A. #A; napply (foo ?); nqed. -naxiom P: Prop. - -alias id "P" = "cic:/matita/tests/ng_commands/P.con". +naxiom NP: Prop. -ndefinition Q: Prop ≝ P. +ndefinition Q: Prop ≝ NP. nlet rec nzero (n:nat) : nat ≝ match n with [ O ⇒ O | S m ⇒ nzero m]. -alias id "nzero" = "cic:/matita/tests/ng_commands/nzero.con". - ntheorem nzero_ok: nzero (S (S O)) = O. napply (refl_eq ? O); nqed.