X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Flpxs_lleq.ma;h=6f6045dabd917744886cff100e01f28e8d292795;hb=c60524dec7ace912c416a90d6b926bee8553250b;hp=a3ca7a1acad43e7ec715e1fd324ae9dffb539ba9;hpb=f10cfe417b6b8ec1c7ac85c6ecf5fb1b3fdf37db;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma index a3ca7a1ac..6f6045dab 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma @@ -14,7 +14,7 @@ include "basic_2/multiple/lleq_lleq.ma". include "basic_2/reduction/lpx_lleq.ma". -include "basic_2/computation/cpxs_leq.ma". +include "basic_2/computation/cpxs_lreq.ma". include "basic_2/computation/lpxs_drop.ma". include "basic_2/computation/lpxs_cpxs.ma". @@ -23,19 +23,19 @@ include "basic_2/computation/lpxs_cpxs.ma". (* Properties on lazy equivalence for local environments ********************) lemma lleq_lpxs_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, g] K2 → - ∀L1,T,d. L1 ≡[T, d] L2 → - ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ≡[T, d] K2. + ∀L1,T,l. L1 ≡[T, l] L2 → + ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ≡[T, l] K2. #h #g #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/ -#K #K2 #_ #HK2 #IH #L1 #T #d #HT elim (IH … HT) -L2 +#K #K2 #_ #HK2 #IH #L1 #T #l #HT elim (IH … HT) -L2 #L #HL1 #HT elim (lleq_lpx_trans … HK2 … HT) -K /3 width=3 by lpxs_strap1, ex2_intro/ qed-. -lemma lpxs_nlleq_inv_step_sn: ∀h,g,G,L1,L2,T,d. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, d] L2 → ⊥) → - ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, g] L & L1 ≡[T, d] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, g] L0 & L0 ≡[T, d] L2. -#h #g #G #L1 #L2 #T #d #H @(lpxs_ind_dx … H) -L1 +lemma lpxs_nlleq_inv_step_sn: ∀h,g,G,L1,L2,T,l. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ≡[T, l] L2 → ⊥) → + ∃∃L,L0. ⦃G, L1⦄ ⊢ ➡[h, g] L & L1 ≡[T, l] L → ⊥ & ⦃G, L⦄ ⊢ ➡*[h, g] L0 & L0 ≡[T, l] L2. +#h #g #G #L1 #L2 #T #l #H @(lpxs_ind_dx … H) -L1 [ #H elim H -H // -| #L1 #L #H1 #H2 #IH2 #H12 elim (lleq_dec T L1 L d) #H +| #L1 #L #H1 #H2 #IH2 #H12 elim (lleq_dec T L1 L l) #H [ -H1 -H2 elim IH2 -IH2 /3 width=3 by lleq_trans/ -H12 #L0 #L3 #H1 #H2 #H3 #H4 lapply (lleq_nlleq_trans … H … H2) -H2 #H2 elim (lleq_lpx_trans … H1 … H) -L @@ -64,8 +64,8 @@ lemma lpxs_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, /3 width=4 by lpxs_pair, fqu_bind_dx, ex3_intro/ | #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H elim (lleq_inv_flat … H) -H /2 width=4 by fqu_flat_dx, ex3_intro/ -| #G1 #L1 #L #T1 #U1 #e #HL1 #HTU1 #K1 #H1KL1 #H2KL1 - elim (drop_O1_le (Ⓕ) (e+1) K1) +| #G1 #L1 #L #T1 #U1 #m #HL1 #HTU1 #K1 #H1KL1 #H2KL1 + elim (drop_O1_le (Ⓕ) (m+1) K1) [ #K #HK1 lapply (lleq_inv_lift_le … H2KL1 … HK1 HL1 … HTU1 ?) -H2KL1 // #H2KL elim (lpxs_drop_trans_O1 … H1KL1 … HL1) -L1 #K0 #HK10 #H1KL lapply (drop_mono … HK10 … HK1) -HK10 #H destruct @@ -110,32 +110,32 @@ elim (fqus_inv_gen … H) -H ] qed-. -fact leq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ⩬[d, e] L0 → e = ∞ → - ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → - ∃∃L. L ⩬[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & - (∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L). -#h #g #G #L1 #L0 #d #e #H elim H -L1 -L0 -d -e -[ #d #e #_ #L2 #H >(lpxs_inv_atom1 … H) -H +fact lreq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,l,m. L1 ⩬[l, m] L0 → m = ∞ → + ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → + ∃∃L. L ⩬[l, m] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & + (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L). +#h #g #G #L1 #L0 #l #m #H elim H -L1 -L0 -l -m +[ #l #m #_ #L2 #H >(lpxs_inv_atom1 … H) -H /3 width=5 by ex3_intro, conj/ -| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #He destruct -| #I #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H +| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #Hm destruct +| #I #L1 #L0 #V1 #m #HL10 #IHL10 #Hm #Y #H elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct - lapply (ysucc_inv_Y_dx … He) -He #He + lapply (ysucc_inv_Y_dx … Hm) -Hm #Hm elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH - @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, leq_cpxs_trans, leq_pair/ + @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, lreq_cpxs_trans, lreq_pair/ #T elim (IH T) #HL0dx #HL0sn - @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_pair_O_Y/ -| #I1 #I0 #L1 #L0 #V1 #V0 #d #e #HL10 #IHL10 #He #Y #H + @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_pair_O_Y/ +| #I1 #I0 #L1 #L0 #V1 #V0 #l #m #HL10 #IHL10 #Hm #Y #H elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH - @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, leq_succ/ + @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, lreq_succ/ #T elim (IH T) #HL0dx #HL0sn - @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_succ/ + @conj #H @(lleq_lreq_repl … H) -H /3 width=1 by lreq_sym, lreq_succ/ ] qed-. -lemma leq_lpxs_trans_lleq: ∀h,g,G,L1,L0,d. L1 ⩬[d, ∞] L0 → - ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → - ∃∃L. L ⩬[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & - (∀T. L0 ≡[T, d] L2 ↔ L1 ≡[T, d] L). -/2 width=1 by leq_lpxs_trans_lleq_aux/ qed-. +lemma lreq_lpxs_trans_lleq: ∀h,g,G,L1,L0,l. L1 ⩬[l, ∞] L0 → + ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → + ∃∃L. L ⩬[l, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & + (∀T. L0 ≡[T, l] L2 ↔ L1 ≡[T, l] L). +/2 width=1 by lreq_lpxs_trans_lleq_aux/ qed-.