]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/tests/ng_commands.ma
updating the structures for sorts
[helm.git] / helm / software / matita / tests / ng_commands.ma
index dfe92269e4cf081692637219c498d37d1b2e7b13..2241043a9fe05d09f723d02ba1f0ef54c3714bdc 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-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].
 
-ninductive nnat: Type ≝
-   nO: nnat
- | nS: nnat → nnat. *)
\ No newline at end of file
+ntheorem nzero_ok: nzero (S (S O)) = O.
+ napply (refl_eq ? O);
+nqed.
+
+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}.