X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Ftests%2Fng_commands.ma;h=2241043a9fe05d09f723d02ba1f0ef54c3714bdc;hb=b6afef7e73324824025a6d7f313129d55b72cfc6;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..2241043a9 100644 --- a/helm/software/matita/tests/ng_commands.ma +++ b/helm/software/matita/tests/ng_commands.ma @@ -12,44 +12,51 @@ (* *) (**************************************************************************) -include "nat/plus.ma". +include "ng_pts.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. +include "nat/nat.ma". 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. -(* -ninductive nnat: Type ≝ - nO: nnat - | nS: nnat → nnat. *) \ No newline at end of file +naxiom DT: nat → Type. +naxiom dt: ∀n. DT n. + +ninductive nnat (n: nat) (A:DT n): Type ≝ + nO: nnat n A + | nS: mat n A → mat n A → nnat n A +with mat: Type ≝ + |mS : nnat n A → mat n A. + +nlet rec nnzero (n:nnat 0 (dt ?)) : nnat 0 (dt ?) ≝ + match n return ? with + [ nO ⇒ nO 0 (dt ?) + | nS m _ ⇒ nmzero m ] +and nmzero (m:mat 0 (dt ?)) : nnat 0 (dt ?) ≝ + match m return ? with + [ mS n ⇒ nnzero n ]. + +nrecord pair (n: nat) (x: DT n) (label: Type): Type ≝ + { lbl:label; l: pair n x label; r: pair n x label}.