]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/tests/ng_commands.ma
- hExtlib: new function "list_assoc_all"
[helm.git] / helm / software / matita / tests / ng_commands.ma
index d6f97499e10d582d9d119685f1edd472a5a91ffd..f816b36840c2b4684593e011fbf8c430c74433d1 100644 (file)
@@ -16,40 +16,46 @@ include "nat/plus.ma".
 
 ndefinition thesis0: ∀A:Type.Type ≝ λA. A → A.
 
-alias id "thesis0" = "cic:/matita/tests/ng_commands/thesis0.def(0)".
-
 ndefinition thesis: ∀A:Type.Type ≝ λA. ?.
  napply (A → A);
 nqed.
 
-alias id "thesis" = "cic:/matita/tests/ng_commands/thesis.def(0)".
-
 ntheorem foo: ∀A:Type.thesis A.#A;#x;napply x;
 nqed.
 
-alias id "foo" = "cic:/matita/tests/ng_commands/foo.def(0)".
-
 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.decl".
-
-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.fix(0,0,0)".
-
 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