]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf_tif.ma
- context-sensitive computation: more properties
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / tnf_tif.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 "basic_2/substitution/tps_lift.ma".
16 include "basic_2/reducibility/tif.ma".
17 include "basic_2/reducibility/tnf.ma".
18
19 (* CONTEXT-FREE NORMAL TERMS ************************************************)
20
21 (* Main properties properties ***********************************************)
22
23 lemma tpr_tif_eq: ∀T1,T2. T1 ➡ T2 →  𝐈[T1] → T1 = T2.
24 #T1 #T2 #H elim H -T1 -T2
25 [ //
26 | * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
27   [ elim (tif_inv_appl … H) -H #HV1 #HT1 #_
28     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
29   | elim (tif_inv_cast … H)
30   ]
31 | #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H
32   elim (tif_inv_appl … H) -H #_ #_ #H
33   elim (simple_inv_bind … H)
34 | * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #H
35   [ -HT2 -IHV1 -IHT1 elim (tif_inv_abbr … H)
36   | <(tps_inv_refl_SO2 … HT2 ?) -HT2 //
37     elim (tif_inv_abst … H) -H #HV1 #HT1
38     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
39   ]
40 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
41   elim (tif_inv_appl … H) -H #_ #_ #H
42   elim (simple_inv_bind … H)
43 | #V1 #T1 #T2 #T #_ #_ #_ #H
44   elim (tif_inv_abbr … H)
45 | #V1 #T1 #T #_ #_ #H
46   elim (tif_inv_cast … H)
47 ]
48 qed.
49
50 theorem tif_tnf: ∀T1.  𝐈[T1] → 𝐍[T1].
51 /2 width=1/ qed.
52
53 (* Note: this property is unusual *)
54 lemma tnf_trf_false: ∀T1. 𝐑[T1] → 𝐍[T1] → False.
55 #T1 #H elim H -T1
56 [ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2 width=1/
57 | #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2 width=1/
58 | #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/
59 | #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2 width=1/
60 | #V #T #H elim (tnf_inv_abbr … H)
61 | #V #T #H elim (tnf_inv_cast … H)
62 | #V #W #T #H elim (tnf_inv_appl … H) -H #_ #_ #H
63   elim (simple_inv_bind … H)
64 ]
65 qed.
66
67 theorem tnf_tif: ∀T1. 𝐍[T1] → 𝐈[T1].
68 /2 width=3/ qed.
69
70 lemma tnf_abst: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐍[ⓛV.T].
71 /4 width=1/ qed.
72
73 lemma tnf_appl: ∀V,T. 𝐍[V] → 𝐍[T] → 𝐒[T] → 𝐍[ⓐV.T].
74 /4 width=1/ qed.