]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma
- predefined_virtuals: some additions
[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/reducibility/tpr.ma".
16
17 (* CONTEXT-FREE NORMAL TERMS ************************************************)
18
19 definition tnf: predicate term ≝ NF … tpr (eq …).
20
21 interpretation
22    "context-free normality (term)"
23    'Normal T = (tnf T).
24
25 (* Basic inversion lemmas ***************************************************)
26
27 lemma tnf_inv_abst: ∀V,T. 𝐍⦃ⓛV.T⦄ → 𝐍⦃V⦄ ∧ 𝐍⦃T⦄.
28 #V1 #T1 #HVT1 @conj
29 [ #V2 #HV2 lapply (HVT1 (ⓛV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
30 | #T2 #HT2 lapply (HVT1 (ⓛV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
31 ]
32 qed-.
33
34 lemma tnf_inv_appl: ∀V,T. 𝐍⦃ⓐV.T⦄ → ∧∧ 𝐍⦃V⦄ & 𝐍⦃T⦄ & 𝐒⦃T⦄.
35 #V1 #T1 #HVT1 @and3_intro
36 [ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct //
37 | #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct //
38 | generalize in match HVT1; -HVT1 elim T1 -T1 * // * #W1 #U1 #_ #_ #H
39   [ elim (lift_total V1 0 1) #V2 #HV12
40     lapply (H (ⓓW1.ⓐV2.U1) ?) -H /2 width=3/ -HV12 #H destruct
41   | lapply (H (ⓓV1.U1) ?) -H /2 width=1/ #H destruct
42 ]
43 qed-.
44
45 lemma tnf_inv_abbr: ∀V,T. 𝐍⦃ⓓV.T⦄ → ⊥.
46 #V #T #H elim (is_lift_dec T 0 1)
47 [ * #U #HTU
48   lapply (H U ?) -H /2 width=3/ #H destruct
49   elim (lift_inv_pair_xy_y … HTU)
50 | #HT
51   elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12
52   lapply (H (ⓓV.T2) ?) -H /2 width=3/ -HT2 #H destruct /3 width=2/
53 ]
54 qed.
55
56 lemma tnf_inv_cast: ∀V,T. 𝐍⦃ⓝV.T⦄ → ⊥.
57 #V #T #H lapply (H T ?) -H /2 width=1/ #H
58 @discr_tpair_xy_y //
59 qed-.