1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "nat/plus.ma".
17 ndefinition thesis0: ∀A:Type.Type ≝ λA. A → A.
19 ndefinition thesis: ∀A:Type.Type ≝ λA. ?.
23 ntheorem foo: ∀A:Type.thesis A.#A;#x;napply x;
26 ntheorem goo: ∀A:Type.A → A. #A; napply (foo ?);
31 ndefinition Q: Prop ≝ NP.
33 nlet rec nzero (n:nat) : nat ≝
38 ntheorem nzero_ok: nzero (S (S O)) = O.
42 naxiom DT: nat → Type.
45 ninductive nnat (n: nat) (A:DT n): Type ≝
47 | nS: mat n A → mat n A → nnat n A
49 |mS : nnat n A → mat n A.
51 nlet rec nnzero (n:nnat 0 (dt ?)) : nnat 0 (dt ?) ≝
55 and nmzero (m:mat 0 (dt ?)) : nnat 0 (dt ?) ≝
59 nrecord pair (n: nat) (x: DT n) (label: Type): Type ≝
60 { lbl:label; l: pair n x label; r: pair n x label}.