X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Frelocation%2Flifts_lifts.ma;h=3aef874bc1731e436318e7234529f3d340c50287;hb=25c634037771dff0138e5e8e3d4378183ff49b86;hp=41c641716a280339f1c67ae399d534dead1e23aa;hpb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_lifts.ma b/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_lifts.ma index 41c641716..3aef874bc 100644 --- a/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_lifts.ma +++ b/matita/matita/contribs/lambdadelta/static_2/relocation/lifts_lifts.ma @@ -20,9 +20,9 @@ include "static_2/relocation/lifts.ma". (* Basic_1: includes: lift_gen_lift *) (* Basic_2A1: includes: lift_div_le lift_div_be *) -theorem lifts_div4: ∀f2,Tf,T. ⬆*[f2] Tf ≘ T → ∀g2,Tg. ⬆*[g2] Tg ≘ T → +theorem lifts_div4: ∀f2,Tf,T. ⇧*[f2] Tf ≘ T → ∀g2,Tg. ⇧*[g2] Tg ≘ T → ∀f1,g1. H_at_div f2 g2 f1 g1 → - ∃∃T0. ⬆*[f1] T0 ≘ Tf & ⬆*[g1] T0 ≘ Tg. + ∃∃T0. ⇧*[f1] T0 ≘ Tf & ⇧*[g1] T0 ≘ Tg. #f2 #Tf #T #H elim H -f2 -Tf -T [ #f2 #s #g2 #Tg #H #f1 #g1 #_ lapply (lifts_inv_sort2 … H) -H #H destruct @@ -46,13 +46,13 @@ theorem lifts_div4: ∀f2,Tf,T. ⬆*[f2] Tf ≘ T → ∀g2,Tg. ⬆*[g2] Tg ≘ ] qed-. -lemma lifts_div4_one: ∀f,Tf,T. ⬆*[⫯f] Tf ≘ T → - ∀T1. ⬆*[1] T1 ≘ T → - ∃∃T0. ⬆*[1] T0 ≘ Tf & ⬆*[f] T0 ≘ T1. +lemma lifts_div4_one: ∀f,Tf,T. ⇧*[⫯f] Tf ≘ T → + ∀T1. ⇧[1] T1 ≘ T → + ∃∃T0. ⇧[1] T0 ≘ Tf & ⇧*[f] T0 ≘ T1. /4 width=6 by lifts_div4, at_div_id_dx, at_div_pn/ qed-. -theorem lifts_div3: ∀f2,T,T2. ⬆*[f2] T2 ≘ T → ∀f,T1. ⬆*[f] T1 ≘ T → - ∀f1. f2 ⊚ f1 ≘ f → ⬆*[f1] T1 ≘ T2. +theorem lifts_div3: ∀f2,T,T2. ⇧*[f2] T2 ≘ T → ∀f,T1. ⇧*[f] T1 ≘ T → + ∀f1. f2 ⊚ f1 ≘ f → ⇧*[f1] T1 ≘ T2. #f2 #T #T2 #H elim H -f2 -T -T2 [ #f2 #s #f #T1 #H >(lifts_inv_sort2 … H) -T1 // | #f2 #i2 #i #Hi2 #f #T1 #H #f1 #Ht21 elim (lifts_inv_lref2 … H) -H @@ -70,8 +70,8 @@ qed-. (* Basic_1: was: lift1_lift1 (left to right) *) (* Basic_1: includes: lift_free (left to right) lift_d lift1_xhg (right to left) lift1_free (right to left) *) (* Basic_2A1: includes: lift_trans_be lift_trans_le lift_trans_ge lifts_lift_trans_le lifts_lift_trans *) -theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≘ T → ∀f2,T2. ⬆*[f2] T ≘ T2 → - ∀f. f2 ⊚ f1 ≘ f → ⬆*[f] T1 ≘ T2. +theorem lifts_trans: ∀f1,T1,T. ⇧*[f1] T1 ≘ T → ∀f2,T2. ⇧*[f2] T ≘ T2 → + ∀f. f2 ⊚ f1 ≘ f → ⇧*[f] T1 ≘ T2. #f1 #T1 #T #H elim H -f1 -T1 -T [ #f1 #s #f2 #T2 #H >(lifts_inv_sort1 … H) -T2 // | #f1 #i1 #i #Hi1 #f2 #T2 #H #f #Ht21 elim (lifts_inv_lref1 … H) -H @@ -86,9 +86,14 @@ theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≘ T → ∀f2,T2. ⬆*[f2] T ≘ ] qed-. +lemma lifts_trans4_one (f) (T1) (T2): + ∀T. ⇧[1]T1 ≘ T → ⇧*[⫯f]T ≘ T2 → + ∃∃T0. ⇧*[f]T1 ≘ T0 & ⇧[1]T0 ≘ T2. +/4 width=6 by lifts_trans, lifts_split_trans, after_uni_one_dx/ qed-. + (* Basic_2A1: includes: lift_conf_O1 lift_conf_be *) -theorem lifts_conf: ∀f1,T,T1. ⬆*[f1] T ≘ T1 → ∀f,T2. ⬆*[f] T ≘ T2 → - ∀f2. f2 ⊚ f1 ≘ f → ⬆*[f2] T1 ≘ T2. +theorem lifts_conf: ∀f1,T,T1. ⇧*[f1] T ≘ T1 → ∀f,T2. ⇧*[f] T ≘ T2 → + ∀f2. f2 ⊚ f1 ≘ f → ⇧*[f2] T1 ≘ T2. #f1 #T #T1 #H elim H -f1 -T -T1 [ #f1 #s #f #T2 #H >(lifts_inv_sort1 … H) -T2 // | #f1 #i #i1 #Hi1 #f #T2 #H #f2 #Ht21 elim (lifts_inv_lref1 … H) -H @@ -128,3 +133,10 @@ lemma deliftable2_sn_bi: ∀R. deliftable2_sn R → deliftable2_bi R. elim (HR … HU12 … HTU1) -HR -U1 #X #HUX #HTX <(lifts_inj … HUX … HTU2) -U2 -T2 -f // qed-. + +lemma lifts_trans_uni (T): + ∀l1,T1. ⇧[l1] T1 ≘ T → + ∀l2,T2. ⇧[l2] T ≘ T2 → ⇧[l1+l2] T1 ≘ T2. +#T #l1 #T1 #HT1 #l2 #T2 #HT2 +@(lifts_trans … HT1 … HT2) // +qed-.