]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/reducibility/tnf.ma
5868a5c935777d578f0eae1fda58ab43fbb12439
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reducibility / tnf.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/trf.ma".
17 include "Basic_2/reducibility/tpr.ma".
18
19 (* CONTEXT-FREE NORMAL TERMS ************************************************)
20
21 definition tnf: term → Prop ≝
22    NF … tpr (eq …).
23
24 interpretation
25    "context-free normality (term)"
26    'Normal T = (tnf T).
27
28 (* Basic inversion lemmas ***************************************************)
29
30 lemma tnf_inv_abst: ∀V,T. ℕ[𝕔{Abst}V.T] → ℕ[V] ∧ ℕ[T].
31 #V1 #T1 #HVT1 @conj
32 [ #V2 #HV2 lapply (HVT1 (𝕔{Abst}V2.T1) ?) -HVT1 /2/ -HV2 #H destruct -V1 T1 //
33 | #T2 #HT2 lapply (HVT1 (𝕔{Abst}V1.T2) ?) -HVT1 /2/ -HT2 #H destruct -V1 T1 //
34 ]
35 qed.
36
37 lemma tnf_inv_appl: ∀V,T. ℕ[𝕔{Appl}V.T] → ∧∧ ℕ[V] & ℕ[T] & 𝕊[T].
38 #V1 #T1 #HVT1 @and3_intro
39 [ #V2 #HV2 lapply (HVT1 (𝕔{Appl}V2.T1) ?) -HVT1 /2/ -HV2 #H destruct -V1 T1 //
40 | #T2 #HT2 lapply (HVT1 (𝕔{Appl}V1.T2) ?) -HVT1 /2/ -HT2 #H destruct -V1 T1 //
41 | generalize in match HVT1 -HVT1; elim T1 -T1 * // * #W1 #U1 #_ #_ #H
42   [ elim (lift_total V1 0 1) #V2 #HV12
43     lapply (H (𝕔{Abbr}W1.𝕔{Appl}V2.U1) ?) -H /2/ -HV12 #H destruct
44   | lapply (H (𝕔{Abbr}V1.U1) ?) -H /2/ #H destruct
45 ]
46 qed.
47
48 axiom tnf_inv_abbr: ∀V,T. ℕ[𝕔{Abbr}V.T] → False.
49
50 lemma tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
51 #V #T #H lapply (H T ?) -H /2/
52 qed.
53
54 (* Basic properties *********************************************************)
55
56 lemma tpr_tif_eq: ∀T1,T2. T1 ⇒ T2 →  𝕀[T1] → T1 = T2.
57 #T1 #T2 #H elim H -T1 T2
58 [ //
59 | * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H
60   [ elim (tif_inv_appl … H) -H #HV1 #HT1 #_
61     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
62   | elim (tif_inv_cast … H)
63   ]
64 | #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H
65   elim (tif_inv_appl … H) -H #_ #_ #H
66   elim (simple_inv_bind … H)
67 | * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #H
68   [ -HT2 IHV1 IHT1; elim (tif_inv_abbr … H)
69   | <(tps_inv_refl_SO2 … HT2 ?) -HT2 //
70     elim (tif_inv_abst … H) -H #HV1 #HT1
71     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
72   ]
73 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H
74   elim (tif_inv_appl … H) -H #_ #_ #H
75   elim (simple_inv_bind … H)
76 | #V1 #T1 #T2 #T #_ #_ #_ #H
77   elim (tif_inv_abbr … H)
78 | #V1 #T1 #T #_ #_ #H
79   elim (tif_inv_cast … H)
80 ]
81 qed.
82
83 theorem tif_tnf: ∀T1.  𝕀[T1] → ℕ[T1].
84 /2/ qed.
85
86 (* Note: this property is unusual *)
87 theorem tnf_trf_false: ∀T1. ℝ[T1] → ℕ[T1] → False.
88 #T1 #H elim H -T1
89 [ #V #T #_ #IHV #H elim (tnf_inv_abst … H) -H /2/
90 | #V #T #_ #IHT #H elim (tnf_inv_abst … H) -H /2/
91 | #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2/
92 | #V #T #_ #IHV #H elim (tnf_inv_appl … H) -H /2/
93 | #V #T #H elim (tnf_inv_abbr … H)
94 | #V #T #H elim (tnf_inv_cast … H)
95 | #V #W #T #H elim (tnf_inv_appl … H) -H #_ #_ #H
96   elim (simple_inv_bind … H)
97 ]
98 qed.
99
100 theorem tnf_tif: ∀T1. ℕ[T1] → 𝕀[T1].
101 /2/ qed.
102
103 lemma tnf_abst: ∀V,T. ℕ[V] → ℕ[T] → ℕ[𝕔{Abst}V.T].
104 /4 width=1/ qed.
105
106 lemma tnf_appl: ∀V,T. ℕ[V] → ℕ[T] → 𝕊[T] → ℕ[𝕔{Appl}V.T].
107 /4 width=1/ qed.