]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / cnf.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/cpr.ma".
16
17 (* CONTEXT-SENSITIVE NORMAL TERMS *******************************************)
18
19 definition cnf: lenv → predicate term ≝ λL. NF … (cpr L) (eq …).
20
21 interpretation
22    "context-sensitive normality (term)"
23    'Normal L T = (cnf L T).
24
25 (* Basic inversion lemmas ***************************************************)
26
27 lemma cnf_inv_appl: ∀L,V,T. L ⊢ 𝐍⦃ⓐV.T⦄ → ∧∧ L ⊢ 𝐍⦃V⦄ & L ⊢ 𝐍⦃T⦄ & 𝐒⦃T⦄.
28 #L #V1 #T1 #HVT1 @and3_intro
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 | generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H
32   [ elim (lift_total V1 0 1) #V2 #HV12
33     lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3/ -HV12 #H destruct
34   | lapply (H (ⓓ{a}V1.U1) ?) -H /3 width=1/ #H destruct
35 ]
36 qed-.
37
38 lemma cnf_inv_zeta: ∀L,V,T. L ⊢ 𝐍⦃+ⓓV.T⦄ → ⊥.
39 #L #V #T #H elim (is_lift_dec T 0 1)
40 [ * #U #HTU
41   lapply (H U ?) -H /3 width=3 by cpr_tpr, tpr_zeta/ #H destruct (**) (* auto too slow without trace *)
42   elim (lift_inv_pair_xy_y … HTU)
43 | #HT
44   elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12
45   lapply (H (+ⓓV.T2) ?) -H /3 width=3 by cpr_tpr, tpr_delta/ -HT2 #H destruct /3 width=2/ (**) (* auto too slow without trace *)
46 ]
47 qed.
48
49 lemma cnf_inv_tau: ∀L,V,T. L ⊢ 𝐍⦃ⓝV.T⦄ → ⊥.
50 #L #V #T #H lapply (H T ?) -H /2 width=1/ #H
51 @discr_tpair_xy_y //
52 qed-.
53
54 (* Basic properties *********************************************************)
55
56 (* Basic_1: was: nf2_sort *)
57 lemma cnf_sort: ∀L,k. L ⊢ 𝐍⦃⋆k⦄.
58 #L #k #X #H
59 >(cpr_inv_sort1 … H) //
60 qed.
61
62 (* Basic_1: was: nf2_dec *)
63 axiom cnf_dec: ∀L,T1. L ⊢ 𝐍⦃T1⦄ ∨
64                ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → ⊥).
65
66 (* Basic_1: removed theorems 1: nf2_abst_shift *)