X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fng_commands.ma;h=f816b36840c2b4684593e011fbf8c430c74433d1;hb=d3548c16f481b14ce94e64c790bc767c59590050;hp=dfe92269e4cf081692637219c498d37d1b2e7b13;hpb=e085135177f7b3b74b410d47a4f3bca1784b60b1;p=helm.git diff --git a/helm/software/matita/tests/ng_commands.ma b/helm/software/matita/tests/ng_commands.ma index dfe92269e..f816b3684 100644 --- a/helm/software/matita/tests/ng_commands.ma +++ b/helm/software/matita/tests/ng_commands.ma @@ -16,34 +16,46 @@ 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. +naxiom NP: Prop. -alias id "P" = "cic:/matita/tests/ng_commands/P.con". +ndefinition Q: Prop ≝ NP. -ndefinition Q: Prop ≝ P. - -(* nlet rec nzero (n:nat) : nat ≝ match n with [ O ⇒ O | S m ⇒ nzero m]. +ntheorem nzero_ok: nzero (S (S O)) = O. + napply (refl_eq ? O); +nqed. + ninductive nnat: Type ≝ nO: nnat - | nS: nnat → nnat. *) \ No newline at end of file + | nS: mat → mat → nnat +with mat: Type ≝ + |mS : nnat → mat +. + +nlet rec nnzero (n:nnat) : nnat ≝ + match n return ? with + [ nO ⇒ nO + | nS m _ ⇒ nmzero m ] +and nmzero (m:mat) : nnat ≝ + match m return ? with + [ mS n ⇒ nnzero n ]. + +(* testare anche i record e le ricorsioni/coricorsioni/(co)induttivi MUTUI *) + +(* +nrecord pair: Type ≝ { l: pair; r: pair }. +*) \ No newline at end of file