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 "logic/equality.ma".
17 ninductive unit: Type[0] ≝ k: unit.
19 ninductive bool: unit → Type[0] ≝ true : bool k | false : bool k.
21 nlemma foo: true = false → False. #H; ndestruct;
24 (* nlemma prova : ∀T:Type[0].∀a,b:T.∀e:a = b.
25 ∀P:∀x,y:T.x=y→Prop.P a a (refl T a) → P a b e.
29 ninductive I1 : Type[0] ≝
32 ninductive I2 : I1 → Type[0] ≝
35 ninductive I3 : ∀x:I1.I2 x → Type[0] ≝
38 ninductive I4 : ∀x,y.I3 x y → Type[0] ≝
39 | k4 : ∀x,y.∀z:I3 x y.I4 x y z.
41 (*alias id "eq" = "cic:/matita/ng/logic/equality/eq.ind(1,0,2)".
42 alias id "refl" = "cic:/matita/ng/logic/equality/eq.con(0,1,2)".*)
44 ninductive II : Type[0] ≝
45 | kII1 : ∀x,y,z.∀w:I4 x y z.II
46 | kII2 : ∀x,y,z.∀w:I4 x y z.II.
48 (* nlemma foo : ∀a,b,c,d,e,f,g,h. kII1 a b c d = kII2 e f g h → True.
49 #a;#b;#c;#d;#e;#f;#g;#h;#H;
53 nlemma faof : ∀a,b,c,d,e,f,g,h.∀Heq:kII1 a b c d = kII1 e f g h.
54 ∀P:(∀a,b,c,d.kII1 a b c d = kII1 e f g h → Prop).
55 P e f g h (refl ??) → P a b c d Heq.
56 #a;#b;#c;#d;#e;#f;#g;#h;#Heq;#P;#HP;