]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cnx.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/notation/relations/predtynormal_5.ma".
16 include "static_2/syntax/tdeq.ma".
17 include "basic_2/rt_transition/cpx.ma".
18
19 (* NORMAL TERMS FOR UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION ********)
20
21 definition cnx: ∀h. sd h → relation3 genv lenv term ≝
22                 λh,o,G,L. NF … (cpx h G L) (tdeq h o …).
23
24 interpretation
25    "normality for unbound context-sensitive parallel rt-transition (term)"
26    'PRedTyNormal h o G L T = (cnx h o G L T).
27
28 (* Basic inversion lemmas ***************************************************)
29
30 lemma cnx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆s⦄ → deg h o s 0.
31 #h #o #G #L #s #H
32 lapply (H (⋆(next h s)) ?) -H /2 width=2 by cpx_ess/ -G -L #H
33 elim (tdeq_inv_sort1 … H) -H #s0 #d #H1 #H2 #H destruct
34 lapply (deg_next … H1) #H0
35 lapply (deg_mono … H0 … H2) -H0 -H2 #H
36 <(pred_inv_refl … H) -H //
37 qed-.
38
39 lemma cnx_inv_abst: ∀h,o,p,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓛ{p}V.T⦄ →
40                     ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄.
41 #h #o #p #G #L #V1 #T1 #HVT1 @conj
42 [ #V2 #HV2 lapply (HVT1 (ⓛ{p}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2
43 | #T2 #HT2 lapply (HVT1 (ⓛ{p}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2
44 ]
45 #H elim (tdeq_inv_pair … H) -H //
46 qed-.
47
48 (* Basic_2A1: was: cnx_inv_abbr *)
49 lemma cnx_inv_abbr_neg: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃-ⓓV.T⦄ →
50                         ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄.
51 #h #o #G #L #V1 #T1 #HVT1 @conj
52 [ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 
53 | #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2
54 ]
55 #H elim (tdeq_inv_pair … H) -H //
56 qed-.
57
58 (* Basic_2A1: was: cnx_inv_eps *)
59 lemma cnx_inv_cast: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓝV.T⦄ → ⊥.
60 #h #o #G #L #V #T #H lapply (H T ?) -H
61 /2 width=6 by cpx_eps, tdeq_inv_pair_xy_y/
62 qed-.
63
64 (* Basic properties *********************************************************)
65
66 lemma cnx_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆s⦄.
67 #h #o #G #L #s #Hs #X #H elim (cpx_inv_sort1 … H) -H
68 /3 width=3 by tdeq_sort, deg_next/
69 qed.
70
71 lemma cnx_sort_iter: ∀h,o,G,L,s,d. deg h o s d → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆((next h)^d s)⦄.
72 #h #o #G #L #s #d #Hs lapply (deg_iter … d Hs) -Hs
73 <minus_n_n /2 width=6 by cnx_sort/
74 qed.
75
76 lemma cnx_abst: ∀h,o,p,G,L,W,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃W⦄ → ⦃G, L.ⓛW⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄ →
77                 ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓛ{p}W.T⦄.
78 #h #o #p #G #L #W #T #HW #HT #X #H
79 elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct
80 @tdeq_pair [ @HW | @HT ] // (**) (* auto fails because δ-expansion gets in the way *)
81 qed.