X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpx_fqus.ma;h=18a5023359979f8b9358de78e70226c6c03393f9;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hp=41763dfa2c97aa65d8b5657d5a0235f97d9a8094;hpb=f694e3336cbdabdeefd86f85d827edfd26bf3464;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fqus.ma index 41763dfa2..18a502335 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fqus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_fqus.ma @@ -12,37 +12,43 @@ (* *) (**************************************************************************) -(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************) +(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) + +include "static_2/relocation/lifts_teqx.ma". +include "static_2/s_computation/fqus_fqup.ma". +include "basic_2/rt_transition/cpx_drops.ma". +include "basic_2/rt_transition/cpx_lsubr.ma". (* Properties on supclosure *************************************************) -lemma fqu_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 -/3 width=3 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, cpx_pair_sn, cpx_bind, cpx_flat, ex2_intro/ -[ #I #G #L #V2 #U2 #HVU2 - elim (lift_total U2 0 1) - /4 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop, ex2_intro/ -| #G #L #K #T1 #U1 #k #HLK1 #HTU1 #T2 #HTU2 - elim (lift_total T2 0 (k+1)) - /3 width=11 by cpx_lift, fqu_drop, ex2_intro/ +lemma fqu_cpx_trans: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & ❪G1,L1,U1❫ ⬂[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +/3 width=3 by cpx_pair_sn, cpx_bind, cpx_flat, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, ex2_intro/ +[ #I #G #L2 #V2 #X2 #HVX2 + elim (lifts_total X2 (𝐔❨1❩)) + /3 width=3 by fqu_drop, cpx_delta, ex2_intro/ +| /5 width=4 by lsubr_cpx_trans, cpx_bind, lsubr_unit, fqu_clear, ex2_intro/ +| #I #G #L2 #T2 #X2 #HTX2 #U2 #HTU2 + elim (cpx_lifts_sn … HTU2 (Ⓣ) … (L2.ⓘ[I]) … HTX2) + /3 width=3 by fqu_drop, drops_refl, drops_drop, ex2_intro/ ] qed-. -lemma fquq_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fquq_inv_gen … H) -H -[ #HT12 elim (fqu_cpx_trans … HT12 … HTU2) /3 width=3 by fqu_fquq, ex2_intro/ +lemma fquq_cpx_trans: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂⸮[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & ❪G1,L1,U1❫ ⬂⸮[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H +[ #HT12 #U2 #HTU2 elim (fqu_cpx_trans … HT12 … HTU2) /3 width=3 by fqu_fquq, ex2_intro/ | * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/ ] qed-. -lemma fqup_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2 +lemma fqup_cpx_trans: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & ❪G1,L1,U1❫ ⬂+[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2 [ #G2 #L2 #T2 #H12 #U2 #HTU2 elim (fqu_cpx_trans … H12 … HTU2) -T2 /3 width=3 by fqu_fqup, ex2_intro/ | #G #G2 #L #L2 #T #T2 #_ #HT2 #IHT1 #U2 #HTU2 @@ -51,71 +57,74 @@ lemma fqup_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, ] qed-. -lemma fqus_cpx_trans: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 elim (fqus_inv_gen … H) -H -[ #HT12 elim (fqup_cpx_trans … HT12 … HTU2) /3 width=3 by fqup_fqus, ex2_intro/ +lemma fqus_cpx_trans: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂*[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & ❪G1,L1,U1❫ ⬂*[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqus_inv_fqup … H) -H +[ #HT12 #U2 #HTU2 elim (fqup_cpx_trans … HT12 … HTU2) /3 width=3 by fqup_fqus, ex2_intro/ | * #H1 #H2 #H3 destruct /2 width=3 by ex2_intro/ ] qed-. -lemma fqu_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → (T2 = U2 → ⊥) → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 -[ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1) +lemma fqu_cpx_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → (T2 ≛ U2 → ⊥) → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & T1 ≛ U1 → ⊥ & ❪G1,L1,U1❫ ⬂[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +[ #I #G #L #V1 #V2 #HV12 #_ elim (lifts_total V2 𝐔❨1❩) #U2 #HVU2 @(ex3_intro … U2) - [1,3: /3 width=7 by fqu_drop, cpx_delta, drop_pair, drop_drop/ - | #H destruct - lapply (lift_inv_lref2_be … HVU2 ? ?) -HVU2 // + [1,3: /3 width=7 by cpx_delta, fqu_drop/ + | #H lapply (teqx_inv_lref1 … H) -H + #H destruct /2 width=5 by lifts_inv_lref2_uni_lt/ ] -| #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T)) +| #I #G #L #V1 #T #V2 #HV12 #H0 @(ex3_intro … (②[I]V2.T)) [1,3: /2 width=4 by fqu_pair_sn, cpx_pair_sn/ - | #H0 destruct /2 width=1 by/ + | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ ] -| #a #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓑ{a,I}V.T2)) +| #p #I #G #L #V #T1 #Hb #T2 #HT12 #H0 @(ex3_intro … (ⓑ[p,I]V.T2)) [1,3: /2 width=4 by fqu_bind_dx, cpx_bind/ - | #H0 destruct /2 width=1 by/ + | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ + ] +| #p #I #G #L #V #T1 #Hb #T2 #HT12 #H0 @(ex3_intro … (ⓑ[p,I]V.T2)) + [1,3: /4 width=4 by lsubr_cpx_trans, cpx_bind, lsubr_unit, fqu_clear/ + | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ ] -| #I #G #L #V #T1 #T2 #HT12 #H @(ex3_intro … (ⓕ{I}V.T2)) +| #I #G #L #V #T1 #T2 #HT12 #H0 @(ex3_intro … (ⓕ[I]V.T2)) [1,3: /2 width=4 by fqu_flat_dx, cpx_flat/ - | #H0 destruct /2 width=1 by/ + | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ ] -| #G #L #K #T1 #U1 #k #HLK #HTU1 #T2 #HT12 #H elim (lift_total T2 0 (k+1)) - #U2 #HTU2 @(ex3_intro … U2) - [1,3: /2 width=10 by cpx_lift, fqu_drop/ - | #H0 destruct /3 width=5 by lift_inj/ +| #I #G #L #T1 #U1 #HTU1 #T2 #HT12 #H0 + elim (cpx_lifts_sn … HT12 (Ⓣ) … (L.ⓘ[I]) … HTU1) -HT12 + /4 width=6 by fqu_drop, drops_refl, drops_drop, teqx_inv_lifts_bi, ex3_intro/ ] qed-. -lemma fquq_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → (T2 = U2 → ⊥) → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐⸮ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fquq_inv_gen … H12) -H12 -[ #H12 elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2 +lemma fquq_cpx_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂⸮[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → (T2 ≛ U2 → ⊥) → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & T1 ≛ U1 → ⊥ & ❪G1,L1,U1❫ ⬂⸮[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H12 elim H12 -H12 +[ #H12 #U2 #HTU2 #H elim (fqu_cpx_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqu_fquq, ex3_intro/ | * #HG #HL #HT destruct /3 width=4 by ex3_intro/ ] qed-. -lemma fqup_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → (T2 = U2 → ⊥) → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐+ ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 -[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpx_trans_neq … H12 … HTU2 H) -T2 +lemma fqup_cpx_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → (T2 ≛ U2 → ⊥) → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & T1 ≛ U1 → ⊥ & ❪G1,L1,U1❫ ⬂+[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind_dx … H) -G1 -L1 -T1 +[ #G1 #L1 #T1 #H12 #U2 #HTU2 #H elim (fqu_cpx_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqu_fqup, ex3_intro/ | #G #G1 #L #L1 #T #T1 #H1 #_ #IH12 #U2 #HTU2 #H elim (IH12 … HTU2 H) -T2 - #U1 #HTU1 #H #H12 elim (fqu_cpx_trans_neq … H1 … HTU1 H) -T1 + #U1 #HTU1 #H #H12 elim (fqu_cpx_trans_tneqx … H1 … HTU1 H) -T1 /3 width=8 by fqup_strap2, ex3_intro/ ] qed-. -lemma fqus_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ⬈[h, o] U2 → (T2 = U2 → ⊥) → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ⬈[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ⊐* ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_gen … H12) -H12 -[ #H12 elim (fqup_cpx_trans_neq … H12 … HTU2 H) -T2 +lemma fqus_cpx_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ⬂*[b] ❪G2,L2,T2❫ → + ∀U2. ❪G2,L2❫ ⊢ T2 ⬈[h] U2 → (T2 ≛ U2 → ⊥) → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈[h] U1 & T1 ≛ U1 → ⊥ & ❪G1,L1,U1❫ ⬂*[b] ❪G2,L2,U2❫. +#h #b #G1 #G2 #L1 #L2 #T1 #T2 #H12 #U2 #HTU2 #H elim (fqus_inv_fqup … H12) -H12 +[ #H12 elim (fqup_cpx_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqup_fqus, ex3_intro/ | * #HG #HL #HT destruct /3 width=4 by ex3_intro/ ]