From 926796df5884453d8f0cf9f294d7776d469ef45b Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Mon, 6 Jun 2016 20:20:44 +0000 Subject: [PATCH] work in progress on frees_drops --- .../lambdadelta/basic_2/relocation/drops.ma | 12 ++ .../lambdadelta/basic_2/relocation/lifts.ma | 30 +++- .../lambdadelta/basic_2/static/frees_drops.ma | 163 +++++++++++++++++- .../ground_2/relocation/rtmap_at.ma | 5 + .../ground_2/relocation/rtmap_coafter.ma | 77 ++++++++- .../ground_2/relocation/rtmap_isfin.ma | 8 +- 6 files changed, 276 insertions(+), 19 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index 8d22bcf00..9a4cfca83 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -282,6 +282,18 @@ lemma drops_inv_TF: ∀f,I,L,K,V. ⬇*[Ⓕ, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ (* Advanced inversion lemmas ************************************************) +lemma drops_inv_atom2: ∀b,L,f. ⬇*[b,f] L ≡ ⋆ → + ∃∃n,f1. ⬇*[b,𝐔❴n❵] L ≡ ⋆ & 𝐔❴n❵ ⊚ f1 ≡ f. +#b #L elim L -L +[ /3 width=4 by drops_atom, after_isid_sn, ex2_2_intro/ +| #L #I #V #IH #f #H elim (pn_split f) * #g #H0 destruct + [ elim (drops_inv_skip1 … H) -H #K #W #_ #_ #H destruct + | lapply (drops_inv_drop1 … H) -H #HL + elim (IH … HL) -IH -HL /3 width=8 by drops_drop, after_next, ex2_2_intro/ + ] +] +qed-. + (* Basic_2A1: includes: drop_inv_gen *) lemma drops_inv_gen: ∀b,f,I,L,K,V. ⬇*[b, f] L ≡ K.ⓑ{I}V → 𝐔⦃f⦄ → ⬇*[Ⓣ, f] L ≡ K.ⓑ{I}V. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma index a621b819f..be42c0ebf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma @@ -210,6 +210,28 @@ lemma lifts_inv_flat2: ∀f:rtmap. ∀I,V2,T2,X. ⬆*[f] X ≡ ⓕ{I}V2.T2 → (* Advanced inversion lemmas ************************************************) +lemma lifts_inv_atom1: ∀f,I,Y. ⬆*[f] ⓪{I} ≡ Y → + ∨∨ ∃∃s. I = Sort s & Y = ⋆s + | ∃∃i,j. @⦃i, f⦄ ≡ j & I = LRef i & Y = #j + | ∃∃l. I = GRef l & Y = §l. +#f * #n #Y #H +[ lapply (lifts_inv_sort1 … H) +| elim (lifts_inv_lref1 … H) +| lapply (lifts_inv_gref1 … H) +] -H /3 width=5 by or3_intro0, or3_intro1, or3_intro2, ex3_2_intro, ex2_intro/ +qed-. + +lemma lifts_inv_atom2: ∀f,I,X. ⬆*[f] X ≡ ⓪{I} → + ∨∨ ∃∃s. X = ⋆s & I = Sort s + | ∃∃i,j. @⦃i, f⦄ ≡ j & X = #i & I = LRef j + | ∃∃l. X = §l & I = GRef l. +#f * #n #X #H +[ lapply (lifts_inv_sort2 … H) +| elim (lifts_inv_lref2 … H) +| lapply (lifts_inv_gref2 … H) +] -H /3 width=5 by or3_intro0, or3_intro1, or3_intro2, ex3_2_intro, ex2_intro/ +qed-. + (* Basic_2A1: includes: lift_inv_pair_xy_x *) lemma lifts_inv_pair_xy_x: ∀f,I,V,T. ⬆*[f] ②{I}V.T ≡ V → ⊥. #f #J #V elim V -V @@ -249,14 +271,6 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma lifts_fwd_atom2: ∀f,I,X. ⬆*[f] X ≡ ⓪{I} → ∃J. X = ⓪{J}. -#f * #n #X #H -[ lapply (lifts_inv_sort2 … H) -| elim (lifts_inv_lref2 … H) -| lapply (lifts_inv_gref2 … H) -] -H /2 width=2 by ex_intro/ -qed-. - (* Basic_2A1: includes: lift_inv_O2 *) lemma lifts_fwd_isid: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐈⦃f⦄ → T1 = T2. #f #T1 #T2 #H elim H -f -T1 -T2 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 fd4ec0ce1..aa4f34998 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma @@ -14,19 +14,19 @@ include "ground_2/relocation/rtmap_pushs.ma". include "ground_2/relocation/rtmap_coafter.ma". -include "basic_2/relocation/lifts_lifts.ma". -include "basic_2/relocation/drops.ma". +include "basic_2/relocation/drops_drops.ma". include "basic_2/static/frees.ma". (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************) (* Advanced properties ******************************************************) -lemma frees_lref_atom: ∀L,i. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ → L ⊢ 𝐅*⦃#i⦄ ≡ 𝐈𝐝. -#L elim L -L /2 width=1 by frees_atom/ +lemma frees_lref_atom: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≡ ⋆ → + ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≡ f. +#b #L elim L -L /2 width=1 by frees_atom/ #L #I #V #IH * [ #H lapply (drops_fwd_isid … H ?) -H // #H destruct -| /4 width=3 by frees_lref, drops_inv_drop1/ +| /5 width=3 by frees_eq_repl_back, frees_lref, drops_inv_drop1, eq_push_inv_isid/ ] qed. @@ -58,6 +58,91 @@ qed-. (* Properties with generic slicing for local environments *******************) +axiom coafter_inv_xpx: ∀g2,f1,g. g2 ~⊚ ↑f1 ≡ g → ∀n. @⦃0, g2⦄ ≡ n → + ∃∃f2,f. f2 ~⊚ f1 ≡ f & ⫱*[n]g2 = ↑f2 & ⫱*[n]g = ↑f. +(* +#g2 #g1 #g #Hg #n #Hg2 +lapply (coafter_tls … Hg2 … Hg) -Hg #Hg +lapply (at_pxx_tls … Hg2) -Hg2 #H +elim (at_inv_pxp … H) -H [ |*: // ] #f2 #H2 +elim (coafter_inv_pxx … Hg … H2) -Hg * #f1 #f #Hf #H1 #H0 destruct +