X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flifts_lifts.ma;h=9d7fc61bfee120789fca2059f3d13885132e1c00;hp=8477714f3d2c9ff7e30c3b8e2b319523a6bbdb88;hb=222044da28742b24584549ba86b1805a87def070;hpb=384b04944ac31922ee41418b106b8e19a19ba9f0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts.ma index 8477714f3..9d7fc61bf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts.ma @@ -20,9 +20,9 @@ include "basic_2/relocation/lifts.ma". (* Basic_1: includes: lift_gen_lift *) (* Basic_2A1: includes: lift_div_le lift_div_be *) -theorem lift_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. +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. #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 lift_div4: ∀f2,Tf,T. ⬆*[f2] Tf ≡ T → ∀g2,Tg. ⬆*[g2] Tg ≡ T ] qed-. -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 lift_div4, at_div_id_dx, at_div_pn/ qed-. +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 @@ -87,8 +87,8 @@ theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≡ T → ∀f2,T2. ⬆*[f2] T ≡ 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 @@ -106,13 +106,25 @@ qed-. (* Advanced proprerties *****************************************************) (* Basic_2A1: includes: lift_inj *) -lemma lifts_inj: ∀f,T1,U. ⬆*[f] T1 ≡ U → ∀T2. ⬆*[f] T2 ≡ U → T1 = T2. -#f #T1 #U #H1 #T2 #H2 lapply (isid_after_dx 𝐈𝐝 … f) +lemma lifts_inj: ∀f. is_inj2 … (lifts f). +#f #T1 #U #H1 #T2 #H2 lapply (after_isid_dx 𝐈𝐝 … f) /3 width=6 by lifts_div3, lifts_fwd_isid/ qed-. (* Basic_2A1: includes: lift_mono *) -lemma lifts_mono: ∀f,T,U1. ⬆*[f] T ≡ U1 → ∀U2. ⬆*[f] T ≡ U2 → U1 = U2. -#f #T #U1 #H1 #U2 #H2 lapply (isid_after_sn 𝐈𝐝 … f) +lemma lifts_mono: ∀f,T. is_mono … (lifts f T). +#f #T #U1 #H1 #U2 #H2 lapply (after_isid_sn 𝐈𝐝 … f) /3 width=6 by lifts_conf, lifts_fwd_isid/ qed-. + +lemma liftable2_sn_bi: ∀R. liftable2_sn R → liftable2_bi R. +#R #HR #T1 #T2 #HT12 #f #U1 #HTU1 #U2 #HTU2 +elim (HR … HT12 … HTU1) -HR -T1 #X #HTX #HUX +<(lifts_mono … HTX … HTU2) -T2 -U2 -f // +qed-. + +lemma deliftable2_sn_bi: ∀R. deliftable2_sn R → deliftable2_bi R. +#R #HR #U1 #U2 #HU12 #f #T1 #HTU1 #T2 #HTU2 +elim (HR … HU12 … HTU1) -HR -U1 #X #HUX #HTX +<(lifts_inj … HUX … HTU2) -U2 -T2 -f // +qed-.