]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/nlibrary/logic/destruct_bb.ma
Release 0.5.9.
[helm.git] / helm / software / matita / nlibrary / logic / destruct_bb.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "logic/equality.ma".
16
17 (* nlemma prova : ∀T:Type[0].∀a,b:T.∀e:a = b.
18                ∀P:∀x,y:T.x=y→Prop.P a a (refl T a) → P a b e.
19 #T;#a;#b;#e;#P;#H;
20 nrewrite < e;*)
21
22 ninductive I1 : Type[0] ≝
23 | k1 : I1.
24
25 ninductive I2 : I1 → Type[0] ≝
26 | k2 : ∀x.I2 x.
27
28 ninductive I3 : ∀x:I1.I2 x → Type[0] ≝
29 | k3 : ∀x,y.I3 x y.
30
31 ninductive I4 : ∀x,y.I3 x y → Type[0] ≝
32 | k4 : ∀x,y.∀z:I3 x y.I4 x y z.
33
34 (*alias id "eq" = "cic:/matita/ng/logic/equality/eq.ind(1,0,2)".
35 alias id "refl" = "cic:/matita/ng/logic/equality/eq.con(0,1,2)".*)
36
37 ninductive II : Type[0] ≝
38 | kII1 : ∀x,y,z.∀w:I4 x y z.II
39 | kII2 : ∀x,y,z.∀w:I4 x y z.II.
40
41 (* nlemma foo : ∀a,b,c,d,e,f,g,h. kII1 a b c d = kII2 e f g h → True.
42 #a;#b;#c;#d;#e;#f;#g;#h;#H;
43 ndiscriminate H;
44 nqed. *)
45
46 nlemma faof : ∀a,b,c,d,e,f,g,h.∀Heq:kII1 a b c d = kII1 e f g h.
47               ∀P:(∀a,b,c,d.kII1 a b c d = kII1 e f g h → Prop).
48               P e f g h (refl ??) → P a b c d Heq.
49 #a;#b;#c;#d;#e;#f;#g;#h;#Heq;#P;#HP;
50 ndestruct;
51 napply HP;
52 nqed.