From f4787814123d74c9504e988137c2c13279838257 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Tue, 17 Jan 2017 18:27:32 +0000 Subject: [PATCH] previous lemma proved ... --- .../basic_2/relocation/drops_lexs.ma | 8 +++ .../basic_2/relocation/drops_lstar.ma | 1 - .../basic_2/rt_transition/lfpx_frees.ma | 18 +---- .../lambdadelta/basic_2/static/frees_drops.ma | 70 +++++-------------- .../lambdadelta/basic_2/static/lfxs_drops.ma | 1 - .../ground_2/relocation/rtmap_coafter.ma | 35 ++++++---- 6 files changed, 49 insertions(+), 84 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma index d55d50dbb..a192a206e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lexs.ma @@ -152,3 +152,11 @@ elim (lexs_liftable_co_dedropable … H1RN H1RP H2RN H2RP … HLK1 … HK12 … #X #HX #HLK2 #H1L12 elim (lexs_inv_push1 … HX) -HX #L2 #V2 #H2L12 #HV12 #H destruct /2 width=6 by ex4_2_intro/ qed-. + +lemma drops_atom2_lexs_conf: ∀RN,RP,b,f1,L1. ⬇*[b, f1] L1 ≡ ⋆ → 𝐔⦃f1⦄ → + ∀f,L2. L1 ⦻*[RN, RP, f] L2 → + ∀f2. f1 ~⊚ f2 ≡f → ⬇*[b, f1] L2 ≡ ⋆. +#RN #RP #b #f1 #L1 #H1 #Hf1 #f #L2 #H2 #f2 #H3 +elim (lexs_co_dropable_sn … H1 … H2 … H3) // -H1 -H2 -H3 -Hf1 +#L #H #HL2 lapply (lexs_inv_atom1 … H) -H // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma index f4734a9b0..8badfaf9d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma @@ -14,7 +14,6 @@ include "ground_2/lib/lstar.ma". include "basic_2/relocation/lreq_lreq.ma". -include "basic_2/relocation/drops.ma". (* GENERIC SLICING FOR LOCAL ENVIRONMENTS ***********************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_frees.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_frees.ma index c786518ee..250760ca7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_frees.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_frees.ma @@ -13,7 +13,6 @@ (**************************************************************************) include "basic_2/relocation/drops_lexs.ma". -include "basic_2/s_computation/fqup_weight.ma". include "basic_2/static/frees_drops.ma". include "basic_2/static/lsubf_frees.ma". include "basic_2/static/lfxs.ma". @@ -23,19 +22,6 @@ include "basic_2/rt_transition/cpx_drops.ma". (* Properties with context-sensitive free variables *************************) -axiom pippo: ∀RN,RP,L1,i. ⬇*[Ⓕ, 𝐔❴i❵] L1 ≡ ⋆ → - ∀f,L2. L1 ⦻*[RN, RP, f] L2 → - ⬇*[Ⓕ, 𝐔❴i❵] L2 ≡ ⋆. -(* -#RN #RP #L1 #i #H1 #f #L2 #H2 -lapply (lexs_co_dropable_sn … H1 … H2) // -HL1 -H2 -*) - - -axiom frees_inv_lifts_SO: ∀b,f,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f → - ∀K. ⬇*[b, 𝐔❴1❵] L ≡ K → ∀T. ⬆*[1] T ≡ U → - K ⊢ 𝐅*⦃T⦄ ≡ ⫱f. - axiom frees_pair_flat: ∀L,T,f1,I1,V1. L.ⓑ{I1}V1 ⊢ 𝐅*⦃T⦄ ≡ f1 → ∀f2,I2,V2. L.ⓑ{I2}V2 ⊢ 𝐅*⦃T⦄ ≡ f2 → ∀f0. f1 ⋓ f2 ≡ f0 → @@ -56,8 +42,8 @@ lemma cpx_frees_conf_lfpx: ∀h,G,L1,T1,f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 → elim (frees_inv_lref_drops … H1) -H1 * [ -IH #HL1 #Hg1 elim (cpx_inv_lref1_drops … H0) -H0 - [ #H destruct lapply (pippo … HL1 … H2) -HL1 -H2 - /3 width=3 by frees_lref_atom, sle_refl, ex2_intro/ + [ #H destruct + /5 width=9 by frees_lref_atom, drops_atom2_lexs_conf, coafter_isuni_isid, sle_refl, ex2_intro/ | * -H2 -Hg1 #I #K1 #V1 #V2 #HLK1 lapply (drops_TF … HLK1) -HLK1 #HLK1 lapply (drops_mono … HLK1 … HL1) -L1 #H destruct diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma index c5762eacb..f153ad53c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -14,6 +14,7 @@ include "ground_2/relocation/nstream_coafter.ma". include "basic_2/relocation/drops_drops.ma". +include "basic_2/static/frees_fqup.ma". include "basic_2/static/frees_frees.ma". (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************) @@ -153,61 +154,26 @@ lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → (* Inversion lemmas with generic slicing for local environments *************) +lemma frees_inv_lifts_ex: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → + ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U → + ∃∃f1. f ~⊚ f1 ≡ f2 & K ⊢ 𝐅*⦃T⦄ ≡ f1. +#b #f2 #L #U #Hf2 #f #K #HLK #T elim (frees_total K T) +/3 width=9 by frees_fwd_coafter, ex2_intro/ +qed-. + +lemma frees_inv_lifts_SO: ∀b,f,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f → + ∀K. ⬇*[b, 𝐔❴1❵] L ≡ K → ∀T. ⬆*[1] T ≡ U → + K ⊢ 𝐅*⦃T⦄ ≡ ⫱f. +#b #f #L #U #H #K #HLK #T #HTU elim(frees_inv_lifts_ex … H … HLK … HTU) -b -L -U +#f1 #Hf #Hf1 elim (coafter_inv_nxx … Hf) -Hf +/3 width=5 by frees_eq_repl_back, coafter_isid_inv_sn/ +qed-. + lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U → ∀f1. f ~⊚ f1 ≡ f2 → K ⊢ 𝐅*⦃T⦄ ≡ f1. -#b #f2 #L #U #H lapply (frees_fwd_isfin … H) elim H -f2 -L -U -[ #f2 #I #Hf2 #_ #f #K #H1 #T #H2 #f1 #H3 - lapply (coafter_fwd_isid2 … H3 … Hf2) -H3 // -Hf2 #Hf1 - elim (drops_inv_atom1 … H1) -H1 #H #_ destruct - elim (lifts_inv_atom2 … H2) -H2 * /2 width=3 by frees_atom/ -| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - lapply (lifts_inv_sort2 … H2) -H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - ] /3 width=4 by frees_sort/ -| #f2 #I #L #W #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_inv_next … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (lifts_inv_lref2 … H2) -H2 #i #H2 #H destruct - lapply (at_inv_xxp … H2 ?) -H2 // * #g #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct - elim (coafter_inv_pxn … H3) -H3 [ |*: // ] #g1 #Hf2 #H destruct - /3 width=4 by frees_zero/ -| #f2 #I #L #W #j #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - elim (lifts_inv_lref2 … H2) -H2 #x #H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct - elim (at_inv_xpn … H2) -H2 [ |*: // ] #j #Hg #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - lapply (at_inv_xnn … H2 ????) -H2 [5: |*: // ] - ] /4 width=4 by lifts_lref, frees_lref/ -| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3 - lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2 - lapply (lifts_inv_gref2 … H2) -H2 #H destruct - elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ] - [ #g #g1 #Hf2 #H #H0 destruct - elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct - | #g #Hf2 #H destruct - lapply (drops_inv_drop1 … H1) -H1 - ] /3 width=4 by frees_gref/ -| #f2W #f2U #f2 #p #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3 - elim (sor_inv_isfin3 … H1f2) // #H1f2W #H - lapply (isfin_inv_tl … H) -H - elim (lifts_inv_bind2 … H2) -H2 #V #T #HVW #HTU #H destruct - elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 // #f1W #f1U #H2f2W #H - elim (coafter_inv_tl0 … H) -H /4 width=5 by frees_bind, drops_skip/ -| #f2W #f2U #f2 #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3 - elim (sor_inv_isfin3 … H1f2) // - elim (lifts_inv_flat2 … H2) -H2 #V #T #HVW #HTU #H destruct - elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 /3 width=5 by frees_flat/ -] +#b #f2 #L #U #H #f #K #HLK #T #HTU #f1 #Hf2 elim (frees_inv_lifts_ex … H … HLK … HTU) -b -L -U +/3 width=7 by frees_eq_repl_back, coafter_inj/ qed-. lemma frees_inv_drops: ∀f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 → diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma index f43eb4e95..0a3f287bb 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_drops.ma @@ -14,7 +14,6 @@ include "basic_2/relocation/drops_ceq.ma". include "basic_2/relocation/drops_lexs.ma". -include "basic_2/static/frees_fqup.ma". include "basic_2/static/frees_drops.ma". include "basic_2/static/lfxs.ma". diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma index 56449d838..c6e4b3a96 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_coafter.ma @@ -326,14 +326,15 @@ lemma coafter_isid_inv_sn: ∀f1,f2,f. f1 ~⊚ f2 ≡ f → 𝐈⦃f1⦄ → f2 lemma coafter_isid_inv_dx: ∀f1,f2,f. f1 ~⊚ f2 ≡ f → 𝐈⦃f2⦄ → 𝐈⦃f⦄. /4 width=4 by eq_id_isid, coafter_isid_dx, coafter_mono/ qed-. -(* Properties with uniform relocations **************************************) +(* Properties with test for uniform relocations *****************************) -lemma coafter_uni_sn: ∀i,f. 𝐔❴i❵ ~⊚ f ≡ ↑*[i] f. -#i elim i -i /2 width=5 by coafter_isid_sn, coafter_next/ +lemma coafter_isuni_isid: ∀f2. 𝐈⦃f2⦄ → ∀f1. 𝐔⦃f1⦄ → f1 ~⊚ f2 ≡ f2. +#f #Hf #g #H elim H -g +/3 width=5 by coafter_isid_sn, coafter_eq_repl_back0, coafter_next, eq_push_inv_isid/ qed. -(* -(* Properties on isuni ******************************************************) + +(* lemma coafter_isid_isuni: ∀f1,f2. 𝐈⦃f2⦄ → 𝐔⦃f1⦄ → f1 ~⊚ ⫯f2 ≡ ⫯f1. #f1 #f2 #Hf2 #H elim H -H /5 width=7 by coafter_isid_dx, coafter_eq_repl_back2, coafter_next, coafter_push, eq_push_inv_isid/ @@ -349,9 +350,15 @@ lemma coafter_uni_next2: ∀f2. 𝐔⦃f2⦄ → ∀f1,f. ⫯f2 ~⊚ f1 ≡ f /3 width=5 by coafter_next/ ] qed. +*) -(* Properties on uni ********************************************************) +(* Properties with uniform relocations **************************************) + +lemma coafter_uni_sn: ∀i,f. 𝐔❴i❵ ~⊚ f ≡ ↑*[i] f. +#i elim i -i /2 width=5 by coafter_isid_sn, coafter_next/ +qed. +(* lemma coafter_uni: ∀n1,n2. 𝐔❴n1❵ ~⊚ 𝐔❴n2❵ ≡ 𝐔❴n1+n2❵. @nat_elim2 /4 width=5 by coafter_uni_next2, coafter_isid_sn, coafter_isid_dx, coafter_next/ @@ -360,7 +367,7 @@ qed. (* Forward lemmas on at *****************************************************) lemma coafter_at_fwd: ∀i,i1,f. @⦃i1, f⦄ ≡ i → ∀f2,f1. f2 ~⊚ f1 ≡ f → - ∃∃i2. @⦃i1, f1⦄ ≡ i2 & @⦃i2, f2⦄ ≡ i. + ∃∃i2. @⦃i1, f1⦄ ≡ i2 & @⦃i2, f2⦄ ≡ i. #i elim i -i [2: #i #IH ] #i1 #f #Hf #f2 #f1 #Hf21 [ elim (at_inv_xxn … Hf) -Hf [1,3:* |*: // ] [1: #g #j1 #Hg #H0 #H |2,4: #g #Hg #H ] @@ -377,7 +384,7 @@ lemma coafter_at_fwd: ∀i,i1,f. @⦃i1, f⦄ ≡ i → ∀f2,f1. f2 ~⊚ f1 ≡ qed-. lemma coafter_fwd_at: ∀i,i2,i1,f1,f2. @⦃i1, f1⦄ ≡ i2 → @⦃i2, f2⦄ ≡ i → - ∀f. f2 ~⊚ f1 ≡ f → @⦃i1, f⦄ ≡ i. + ∀f. f2 ~⊚ f1 ≡ f → @⦃i1, f⦄ ≡ i. #i elim i -i [2: #i #IH ] #i2 #i1 #f1 #f2 #Hf1 #Hf2 #f #Hf [ elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ] #g2 [ #j2 ] #Hg2 [ #H22 ] #H20 @@ -395,13 +402,13 @@ lemma coafter_fwd_at: ∀i,i2,i1,f1,f2. @⦃i1, f1⦄ ≡ i2 → @⦃i2, f2⦄ qed-. lemma coafter_fwd_at2: ∀f,i1,i. @⦃i1, f⦄ ≡ i → ∀f1,i2. @⦃i1, f1⦄ ≡ i2 → - ∀f2. f2 ~⊚ f1 ≡ f → @⦃i2, f2⦄ ≡ i. + ∀f2. f2 ~⊚ f1 ≡ f → @⦃i2, f2⦄ ≡ i. #f #i1 #i #Hf #f1 #i2 #Hf1 #f2 #H elim (coafter_at_fwd … Hf … H) -f #j1 #H #Hf2 <(at_mono … Hf1 … H) -i1 -i2 // qed-. lemma coafter_fwd_at1: ∀i,i2,i1,f,f2. @⦃i1, f⦄ ≡ i → @⦃i2, f2⦄ ≡ i → - ∀f1. f2 ~⊚ f1 ≡ f → @⦃i1, f1⦄ ≡ i2. + ∀f1. f2 ~⊚ f1 ≡ f → @⦃i1, f1⦄ ≡ i2. #i elim i -i [2: #i #IH ] #i2 #i1 #f #f2 #Hf #Hf2 #f1 #Hf1 [ elim (at_inv_xxn … Hf) -Hf [1,3: * |*: // ] #g [ #j1 ] #Hg [ #H01 ] #H00 @@ -421,7 +428,7 @@ qed-. (* Properties with at *******************************************************) lemma coafter_uni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. f2 ~⊚ 𝐔❴i1❵ ≡ f → 𝐔❴i2❵ ~⊚ ⫱*[i2] f2 ≡ f. + ∀f. f2 ~⊚ 𝐔❴i1❵ ≡ f → 𝐔❴i2❵ ~⊚ ⫱*[i2] f2 ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct @@ -440,7 +447,7 @@ lemma coafter_uni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → qed. lemma coafter_uni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. 𝐔❴i2❵ ~⊚ ⫱*[i2] f2 ≡ f → f2 ~⊚ 𝐔❴i1❵ ≡ f. + ∀f. 𝐔❴i2❵ ~⊚ ⫱*[i2] f2 ≡ f → f2 ~⊚ 𝐔❴i1❵ ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct @@ -456,7 +463,7 @@ lemma coafter_uni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → qed-. lemma coafter_uni_succ_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. f2 ~⊚ 𝐔❴⫯i1❵ ≡ f → 𝐔❴⫯i2❵ ~⊚ ⫱*[⫯i2] f2 ≡ f. + ∀f. f2 ~⊚ 𝐔❴⫯i1❵ ≡ f → 𝐔❴⫯i2❵ ~⊚ ⫱*[⫯i2] f2 ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct @@ -476,7 +483,7 @@ lemma coafter_uni_succ_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → qed. lemma coafter_uni_succ_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 → - ∀f. 𝐔❴⫯i2❵ ~⊚ ⫱*[⫯i2] f2 ≡ f → f2 ~⊚ 𝐔❴⫯i1❵ ≡ f. + ∀f. 𝐔❴⫯i2❵ ~⊚ ⫱*[⫯i2] f2 ≡ f → f2 ~⊚ 𝐔❴⫯i1❵ ≡ f. #i2 elim i2 -i2 [ #i1 #f2 #Hf2 #f #Hf elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct -- 2.39.2