X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcpxs_fqus.ma;h=1dd3b59db63d3bfb937afebd3bea918f83bf9135;hp=eb71bafca9ad6cdde5cc7eafb25064e126405c52;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma index eb71bafca..1dd3b59db 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_fqus.ma @@ -21,33 +21,33 @@ include "basic_2/rt_computation/cpxs_cpxs.ma". (* Properties on supclosure *************************************************) -lemma fqu_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ⦃G2,L2⦄ ⊢ T2 ⬈*[h] U2 → - ∀T1. ⦃G1,L1,T1⦄ ⬂[b] ⦃G2,L2,T2⦄ → - ∃∃U1. ⦃G1,L1⦄ ⊢ T1 ⬈*[h] U1 & ⦃G1,L1,U1⦄ ⬂[b] ⦃G2,L2,U2⦄. +lemma fqu_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ❪G2,L2❫ ⊢ T2 ⬈*[h] U2 → + ∀T1. ❪G1,L1,T1❫ ⬂[b] ❪G2,L2,T2❫ → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈*[h] U1 & ❪G1,L1,U1❫ ⬂[b] ❪G2,L2,U2❫. #h #b #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/ #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqu_cpx_trans … HT1 … HT2) -T #T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/ qed-. -lemma fquq_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ⦃G2,L2⦄ ⊢ T2 ⬈*[h] U2 → - ∀T1. ⦃G1,L1,T1⦄ ⬂⸮[b] ⦃G2,L2,T2⦄ → - ∃∃U1. ⦃G1,L1⦄ ⊢ T1 ⬈*[h] U1 & ⦃G1,L1,U1⦄ ⬂⸮[b] ⦃G2,L2,U2⦄. +lemma fquq_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ❪G2,L2❫ ⊢ T2 ⬈*[h] U2 → + ∀T1. ❪G1,L1,T1❫ ⬂⸮[b] ❪G2,L2,T2❫ → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈*[h] U1 & ❪G1,L1,U1❫ ⬂⸮[b] ❪G2,L2,U2❫. #h #b #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/ #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fquq_cpx_trans … HT1 … HT2) -T #T #HT1 #HT2 elim (IHTU2 … HT2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/ qed-. -lemma fqup_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ⦃G2,L2⦄ ⊢ T2 ⬈*[h] U2 → - ∀T1. ⦃G1,L1,T1⦄ ⬂+[b] ⦃G2,L2,T2⦄ → - ∃∃U1. ⦃G1,L1⦄ ⊢ T1 ⬈*[h] U1 & ⦃G1,L1,U1⦄ ⬂+[b] ⦃G2,L2,U2⦄. +lemma fqup_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ❪G2,L2❫ ⊢ T2 ⬈*[h] U2 → + ∀T1. ❪G1,L1,T1❫ ⬂+[b] ❪G2,L2,T2❫ → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈*[h] U1 & ❪G1,L1,U1❫ ⬂+[b] ❪G2,L2,U2❫. #h #b #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/ #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqup_cpx_trans … HT1 … HT2) -T #U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/ qed-. -lemma fqus_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ⦃G2,L2⦄ ⊢ T2 ⬈*[h] U2 → - ∀T1. ⦃G1,L1,T1⦄ ⬂*[b] ⦃G2,L2,T2⦄ → - ∃∃U1. ⦃G1,L1⦄ ⊢ T1 ⬈*[h] U1 & ⦃G1,L1,U1⦄ ⬂*[b] ⦃G2,L2,U2⦄. +lemma fqus_cpxs_trans: ∀h,b,G1,G2,L1,L2,T2,U2. ❪G2,L2❫ ⊢ T2 ⬈*[h] U2 → + ∀T1. ❪G1,L1,T1❫ ⬂*[b] ❪G2,L2,T2❫ → + ∃∃U1. ❪G1,L1❫ ⊢ T1 ⬈*[h] U1 & ❪G1,L1,U1❫ ⬂*[b] ❪G2,L2,U2❫. #h #b #G1 #G2 #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 /2 width=3 by ex2_intro/ #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fqus_cpx_trans … HT1 … HT2) -T #U1 #HTU1 #H2 elim (IHTU2 … H2) -T2 /3 width=3 by cpxs_strap2, ex2_intro/ @@ -55,42 +55,42 @@ qed-. (* Note: a proof based on fqu_cpx_trans_tneqx might exist *) (* Basic_2A1: uses: fqu_cpxs_trans_neq *) -lemma fqu_cpxs_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⦄. +lemma fqu_cpxs_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❵) +[ #I #G #L #V1 #V2 #HV12 #_ elim (lifts_total V2 𝐔❨1❩) #U2 #HVU2 @(ex3_intro … U2) [1,3: /3 width=7 by cpxs_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 #H0 @(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, cpxs_pair_sn/ | #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)) +| #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, cpxs_bind/ | #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)) +| #p #I #G #L #V #T1 #Hb #T2 #HT12 #H0 @(ex3_intro … (ⓑ[p,I]V.T2)) [1,3: /4 width=4 by lsubr_cpxs_trans, cpxs_bind, lsubr_unit, fqu_clear/ | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ ] -| #I #G #L #V #T1 #T2 #HT12 #H0 @(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, cpxs_flat/ | #H elim (teqx_inv_pair … H) -H /2 width=1 by/ ] | #I #G #L #T1 #U1 #HTU1 #T2 #HT12 #H0 - elim (cpxs_lifts_sn … HT12 (Ⓣ) … (L.ⓘ{I}) … HTU1) -HT12 + elim (cpxs_lifts_sn … HT12 (Ⓣ) … (L.ⓘ[I]) … HTU1) -HT12 /4 width=6 by fqu_drop, drops_refl, drops_drop, teqx_inv_lifts_bi, ex3_intro/ ] qed-. (* Basic_2A1: uses: fquq_cpxs_trans_neq *) -lemma fquq_cpxs_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⦄. +lemma fquq_cpxs_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_cpxs_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqu_fquq, ex3_intro/ @@ -99,9 +99,9 @@ lemma fquq_cpxs_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂⸮[b] qed-. (* Basic_2A1: uses: fqup_cpxs_trans_neq *) -lemma fqup_cpxs_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⦄. +lemma fqup_cpxs_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_cpxs_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqu_fqup, ex3_intro/ @@ -112,9 +112,9 @@ lemma fqup_cpxs_trans_tneqx: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂+[b] qed-. (* Basic_2A1: uses: fqus_cpxs_trans_neq *) -lemma fqus_cpxs_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⦄. +lemma fqus_cpxs_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_cpxs_trans_tneqx … H12 … HTU2 H) -T2 /3 width=4 by fqup_fqus, ex3_intro/