1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/rt_computation/csx_lsubr.ma".
16 include "basic_2/rt_computation/csx_cpxs.ma".
17 include "basic_2/rt_computation/jsx_rsx.ma".
19 (* STRONGLY NORMALIZING REFERRED LOCAL ENVS FOR EXTENDED RT-TRANSITION ******)
21 (* Forward lemmas with strongly rt-normalizing terms ************************)
23 fact rsx_fwd_lref_pair_csx_aux (G):
25 ∀I,K,V. L = K.ⓑ[I]V → ❪G,K❫ ⊢ ⬈*𝐒 V.
27 @(rsx_ind … H) -L #L #_ #IH #I #K #V1 #H destruct
28 @csx_intro #V2 #HV12 #HnV12
29 @(IH … I) -IH [1,4: // | -HnV12 | -G #H ]
30 [ /2 width=1 by lpx_pair/
31 | elim (reqg_inv_zero_pair_sn … H) -H #Y #X #_ #H1 #H2 destruct -I
36 lemma rsx_fwd_lref_pair_csx (G):
37 ∀I,K,V. G ⊢ ⬈*𝐒[#0] K.ⓑ[I]V → ❪G,K❫ ⊢ ⬈*𝐒 V.
38 /2 width=4 by rsx_fwd_lref_pair_csx_aux/ qed-.
40 lemma rsx_fwd_lref_pair_csx_drops (G):
41 ∀I,K,V,i,L. ⇩[i] L ≘ K.ⓑ[I]V → G ⊢ ⬈*𝐒[#i] L → ❪G,K❫ ⊢ ⬈*𝐒 V.
42 #G #I #K #V #i elim i -i
43 [ #L #H >(drops_fwd_isid … H) -H
44 /2 width=2 by rsx_fwd_lref_pair_csx/
46 elim (drops_inv_bind2_isuni_next … H1) -H1 // #J #Y #HY #H destruct
47 lapply (rsx_inv_lifts … H2 … (𝐔❨1❩) ?????) -H2
48 /3 width=6 by drops_refl, drops_drop/
52 (* Inversion lemmas with strongly rt-normalizing terms **********************)
54 lemma rsx_inv_lref_pair (G):
55 ∀I,K,V. G ⊢ ⬈*𝐒[#0] K.ⓑ[I]V →
56 ∧∧ ❪G,K❫ ⊢ ⬈*𝐒 V & G ⊢ ⬈*𝐒[V] K.
57 /3 width=2 by rsx_fwd_lref_pair_csx, rsx_fwd_pair, conj/ qed-.
59 lemma rsx_inv_lref_pair_drops (G):
60 ∀I,K,V,i,L. ⇩[i] L ≘ K.ⓑ[I]V → G ⊢ ⬈*𝐒[#i] L →
61 ∧∧ ❪G,K❫ ⊢ ⬈*𝐒 V & G ⊢ ⬈*𝐒[V] K.
62 /3 width=5 by rsx_fwd_lref_pair_csx_drops, rsx_fwd_lref_pair_drops, conj/ qed-.
64 lemma rsx_inv_lref_drops (G):
67 | ∃∃I,K. ⇩[i] L ≘ K.ⓤ[I]
68 | ∃∃I,K,V. ⇩[i] L ≘ K.ⓑ[I]V & ❪G,K❫ ⊢ ⬈*𝐒 V & G ⊢ ⬈*𝐒[V] K.
69 #G #L #i #H elim (drops_F_uni L i)
70 [ /2 width=1 by or3_intro0/
71 | * * /4 width=10 by rsx_fwd_lref_pair_csx_drops, rsx_fwd_lref_pair_drops, ex3_3_intro, ex1_2_intro, or3_intro2, or3_intro1/
75 (* Properties with strongly rt-normalizing terms ****************************)
77 (* Note: swapping the eliminations to avoid rsx_cpx_trans: no solution found *)
78 (* Basic_2A1: uses: lsx_lref_be_lpxs *)
79 lemma rsx_lref_pair_lpxs (G):
80 ∀K1,V. ❪G,K1❫ ⊢ ⬈*𝐒 V →
81 ∀K2. G ⊢ ⬈*𝐒[V] K2 → ❪G,K1❫ ⊢ ⬈* K2 →
82 ∀I. G ⊢ ⬈*𝐒[#0] K2.ⓑ[I]V.
84 @(csx_ind_cpxs … H) -V #V0 #_ #IHV0 #K2 #H
85 @(rsx_ind … H) -K2 #K0 #HK0 #IHK0 #HK10 #I
86 @rsx_intro #Y #HY #HnY
87 elim (lpx_inv_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct
88 elim (teqx_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnY ]
89 [ /5 width=5 by rsx_reqx_trans, lpxs_step_dx, reqg_pair, reqg_refl/
90 | @(IHV0 … HnV02) -IHV0 -HnV02
91 [ /2 width=3 by lpxs_cpx_trans/
92 | /3 width=3 by rsx_lpx_trans, rsx_cpx_trans/
93 | /2 width=3 by lpxs_step_dx/
98 lemma rsx_lref_pair (G):
99 ∀K,V. ❪G,K❫ ⊢ ⬈*𝐒 V → G ⊢ ⬈*𝐒[V] K → ∀I. G ⊢ ⬈*𝐒[#0] K.ⓑ[I]V.
100 /2 width=3 by rsx_lref_pair_lpxs/ qed.
102 (* Basic_2A1: uses: lsx_lref_be *)
103 lemma rsx_lref_pair_drops (G):
104 ∀K,V. ❪G,K❫ ⊢ ⬈*𝐒 V → G ⊢ ⬈*𝐒[V] K →
105 ∀I,i,L. ⇩[i] L ≘ K.ⓑ[I]V → G ⊢ ⬈*𝐒[#i] L.
106 #G #K #V #HV #HK #I #i elim i -i
107 [ #L #H >(drops_fwd_isid … H) -H /2 width=1 by rsx_lref_pair/
109 elim (drops_inv_bind2_isuni_next … H) -H // #J #Y #HY #H destruct
110 @(rsx_lifts … (𝐔❨1❩)) /3 width=6 by drops_refl, drops_drop/ (**) (* full auto fails *)
114 (* Main properties with strongly rt-normalizing terms ***********************)
116 (* Basic_2A1: uses: csx_lsx *)
118 ∀L,T. ❪G,L❫ ⊢ ⬈*𝐒 T → G ⊢ ⬈*𝐒[T] L.
119 #G #L #T @(fqup_wf_ind_eq (Ⓣ) … G L T) -G -L -T
120 #Z #Y #X #IH #G #L * *
122 | #i #HG #HL #HT #H destruct
123 elim (csx_inv_lref_drops … H) -H [ |*: * ]
124 [ /2 width=1 by rsx_lref_atom_drops/
125 | /2 width=3 by rsx_lref_unit_drops/
126 | /4 width=6 by rsx_lref_pair_drops, fqup_lref/
129 | #p #I #V #T #HG #HL #HT #H destruct
130 elim (csx_fwd_bind … H) -H /3 width=1 by rsx_bind/
131 | #I #V #T #HG #HL #HT #H destruct
132 elim (csx_fwd_flat … H) -H /3 width=1 by rsx_flat/