]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_2A1/cpr/cnf_cif.etc
update in binaries for λδ
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_2A1 / cpr / cnf_cif.etc
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/cif.ma".
16 include "basic_2/reducibility/cnf_lift.ma".
17
18 (* CONTEXT-SENSITIVE NORMAL TERMS *******************************************)
19
20 (* Main properties **********************************************************)
21
22 lemma tps_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2.
23 #L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e
24 [ //
25 | #L #K #V #W #i #d #e #_ #_ #HLK #_ #H -d -e
26   elim (cif_inv_delta … HLK ?) //
27 | #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H
28   elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct
29   lapply (IHV12 … HV1) -IHV12 -HV1 #H destruct /3 width=1/
30 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H
31   elim (cif_inv_flat … H) -H #HV1 #HT1 #_ #_ /3 width=1/
32 ]
33 qed.
34
35 lemma tpss_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶*[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2.
36 #L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 //
37 #T #T2 #_ #HT2 #IHT1 #HT1
38 lapply (IHT1 HT1) -IHT1 #H destruct /2 width=5/
39 qed.
40
41 lemma tpr_cif_eq: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ 𝐈⦃T1⦄ → T1 = T2.
42 #T1 #T2 #H elim H -T1 -T2
43 [ //
44 | * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #L #H
45   [ elim (cif_inv_appl … H) -H #HV1 #HT1 #_
46     >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 //
47   | elim (cif_inv_ri2 … H) /2 width=1/
48   ]
49 | #a #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #L #H
50   elim (cif_inv_appl … H) -H #_ #_ #H
51   elim (simple_inv_bind … H)
52 | #a * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #L #H
53   [ lapply (tps_lsubr_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2
54     elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct
55     lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct
56     lapply (IHT1 … HT1) -IHT1 #H destruct
57     lapply (tps_cif_eq … HT2 ?) -HT2 //
58   | <(tps_inv_refl_SO2 … HT2 ?) -HT2 //
59     elim (cif_inv_ib2 … H) -H /2 width=1/ /3 width=2/
60   ]
61 | #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #L #H
62   elim (cif_inv_appl … H) -H #_ #_ #H
63   elim (simple_inv_bind … H)
64 | #V1 #T1 #T #T2 #_ #_ #_ #L #H
65   elim (cif_inv_ri2 … H) /2 width=1/
66 | #V1 #T1 #T2 #_ #_ #L #H
67   elim (cif_inv_ri2 … H) /2 width=1/
68 ]
69 qed.
70
71 lemma cpr_cif_eq: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2.
72 #L #T1 #T2 * #T0 #HT10 #HT02 #HT1
73 lapply (tpr_cif_eq … HT10 … HT1) -HT10 #H destruct /2 width=5/
74 qed.
75
76 theorem cif_cnf: ∀L,T. L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐍⦃T⦄.
77 /3 width=3/ qed.
78
79 (* Note: this property is unusual *)
80 lemma cnf_crf_false: ∀L,T. L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ⊥.
81 #L #T #H elim H -L -T
82 [ #L #K #V #i #HLK #H
83   elim (cnf_inv_delta … HLK H)
84 | #L #V #T #_ #IHV #H
85   elim (cnf_inv_appl … H) -H /2 width=1/
86 | #L #V #T #_ #IHT #H
87   elim (cnf_inv_appl … H) -H /2 width=1/
88 | #I #L #V #T * #H1 #H2 destruct
89   [ elim (cnf_inv_zeta … H2)
90   | elim (cnf_inv_tau … H2)
91   ]
92 |5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct
93   [1,3: elim (cnf_inv_abbr … H2) -H2 /2 width=1/
94   |*: elim (cnf_inv_abst … H2) -H2 /2 width=1/
95   ]
96 | #a #L #V #W #T #H
97   elim (cnf_inv_appl … H) -H #_ #_ #H
98   elim (simple_inv_bind … H)
99 | #a #L #V #W #T #H
100   elim (cnf_inv_appl … H) -H #_ #_ #H
101   elim (simple_inv_bind … H)
102 ]
103 qed.
104
105 theorem cnf_cif: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ 𝐈⦃T⦄.
106 /2 width=4/ qed.