X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Ffrees_drops.ma;h=37c629351f1c30b64f19adb8ee78124871942f2a;hp=bc5a7056babb2ba2b28e75aa6627e699554963f0;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 diff --git a/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma b/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma index bc5a7056b..37c629351 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma @@ -21,8 +21,8 @@ include "static_2/static/frees_fqup.ma". (* Advanced properties ******************************************************) lemma frees_atom_drops: - ∀b,L,i. ⇩*[b,𝐔❴i❵] L ≘ ⋆ → - ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i]↑f. + ∀b,L,i. ⇩*[b,𝐔❨i❩] L ≘ ⋆ → + ∀f. 𝐈❪f❫ → L ⊢ 𝐅+❪#i❫ ≘ ⫯*[i]↑f. #b #L elim L -L /2 width=1 by frees_atom/ #L #I #IH * [ #H lapply (drops_fwd_isid … H ?) -H // #H destruct @@ -31,8 +31,8 @@ lemma frees_atom_drops: qed. lemma frees_pair_drops: - ∀f,K,V. K ⊢ 𝐅+⦃V⦄ ≘ f → - ∀i,I,L. ⇩*[i] L ≘ K.ⓑ{I}V → L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i] ↑f. + ∀f,K,V. K ⊢ 𝐅+❪V❫ ≘ f → + ∀i,I,L. ⇩*[i] L ≘ K.ⓑ[I]V → L ⊢ 𝐅+❪#i❫ ≘ ⫯*[i] ↑f. #f #K #V #Hf #i elim i -i [ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_pair/ | #i #IH #I #L #H elim (drops_inv_succ … H) -H /3 width=2 by frees_lref/ @@ -40,8 +40,8 @@ lemma frees_pair_drops: qed. lemma frees_unit_drops: - ∀f. 𝐈⦃f⦄ → ∀I,K,i,L. ⇩*[i] L ≘ K.ⓤ{I} → - L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i] ↑f. + ∀f. 𝐈❪f❫ → ∀I,K,i,L. ⇩*[i] L ≘ K.ⓤ[I] → + L ⊢ 𝐅+❪#i❫ ≘ ⫯*[i] ↑f. #f #Hf #I #K #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_unit/ | #i #IH #Y #H elim (drops_inv_succ … H) -H @@ -50,8 +50,8 @@ lemma frees_unit_drops: qed. lemma frees_lref_pushs: - ∀f,K,j. K ⊢ 𝐅+⦃#j⦄ ≘ f → - ∀i,L. ⇩*[i] L ≘ K → L ⊢ 𝐅+⦃#(i+j)⦄ ≘ ⫯*[i] f. + ∀f,K,j. K ⊢ 𝐅+❪#j❫ ≘ f → + ∀i,L. ⇩*[i] L ≘ K → L ⊢ 𝐅+❪#(i+j)❫ ≘ ⫯*[i] f. #f #K #j #Hf #i elim i -i [ #L #H lapply (drops_fwd_isid … H ?) -H // | #i #IH #L #H elim (drops_inv_succ … H) -H @@ -62,10 +62,10 @@ qed. (* Advanced inversion lemmas ************************************************) lemma frees_inv_lref_drops: - ∀L,i,f. L ⊢ 𝐅+⦃#i⦄ ≘ f → - ∨∨ ∃∃g. ⇩*[Ⓕ,𝐔❴i❵] L ≘ ⋆ & 𝐈⦃g⦄ & f = ⫯*[i] ↑g - | ∃∃g,I,K,V. K ⊢ 𝐅+⦃V⦄ ≘ g & ⇩*[i] L ≘ K.ⓑ{I}V & f = ⫯*[i] ↑g - | ∃∃g,I,K. ⇩*[i] L ≘ K.ⓤ{I} & 𝐈⦃g⦄ & f = ⫯*[i] ↑g. + ∀L,i,f. L ⊢ 𝐅+❪#i❫ ≘ f → + ∨∨ ∃∃g. ⇩*[Ⓕ,𝐔❨i❩] L ≘ ⋆ & 𝐈❪g❫ & f = ⫯*[i] ↑g + | ∃∃g,I,K,V. K ⊢ 𝐅+❪V❫ ≘ g & ⇩*[i] L ≘ K.ⓑ[I]V & f = ⫯*[i] ↑g + | ∃∃g,I,K. ⇩*[i] L ≘ K.ⓤ[I] & 𝐈❪g❫ & f = ⫯*[i] ↑g. #L elim L -L [ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H [ elim (frees_inv_atom … H) -H #f #Hf #H destruct @@ -86,9 +86,9 @@ qed-. (* Properties with generic slicing for local environments *******************) lemma frees_lifts: - ∀b,f1,K,T. K ⊢ 𝐅+⦃T⦄ ≘ f1 → + ∀b,f1,K,T. K ⊢ 𝐅+❪T❫ ≘ f1 → ∀f,L. ⇩*[b,f] L ≘ K → ∀U. ⇧*[f] T ≘ U → - ∀f2. f ~⊚ f1 ≘ f2 → L ⊢ 𝐅+⦃U⦄ ≘ f2. + ∀f2. f ~⊚ f1 ≘ f2 → L ⊢ 𝐅+❪U❫ ≘ f2. #b #f1 #K #T #H lapply (frees_fwd_isfin … H) elim H -f1 -K -T [ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3 lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2 @@ -144,8 +144,8 @@ lemma frees_lifts: qed-. lemma frees_lifts_SO: - ∀b,L,K. ⇩*[b,𝐔❴1❵] L ≘ K → ∀T,U. ⇧*[1] T ≘ U → - ∀f. K ⊢ 𝐅+⦃T⦄ ≘ f → L ⊢ 𝐅+⦃U⦄ ≘ ⫯f. + ∀b,L,K. ⇩*[b,𝐔❨1❩] L ≘ K → ∀T,U. ⇧*[1] T ≘ U → + ∀f. K ⊢ 𝐅+❪T❫ ≘ f → L ⊢ 𝐅+❪U❫ ≘ ⫯f. #b #L #K #HLK #T #U #HTU #f #Hf @(frees_lifts b … Hf … HTU) // (**) (* auto fails *) qed. @@ -153,44 +153,44 @@ qed. (* Forward lemmas with generic slicing for local environments ***************) lemma frees_fwd_coafter: - ∀b,f2,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f2 → + ∀b,f2,L,U. L ⊢ 𝐅+❪U❫ ≘ f2 → ∀f,K. ⇩*[b,f] L ≘ K → ∀T. ⇧*[f] T ≘ U → - ∀f1. K ⊢ 𝐅+⦃T⦄ ≘ f1 → f ~⊚ f1 ≘ f2. + ∀f1. K ⊢ 𝐅+❪T❫ ≘ f1 → f ~⊚ f1 ≘ f2. /4 width=11 by frees_lifts, frees_mono, coafter_eq_repl_back0/ qed-. (* Inversion lemmas with generic slicing for local environments *************) lemma frees_inv_lifts_ex: - ∀b,f2,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f2 → + ∀b,f2,L,U. L ⊢ 𝐅+❪U❫ ≘ f2 → ∀f,K. ⇩*[b,f] L ≘ K → ∀T. ⇧*[f] T ≘ U → - ∃∃f1. f ~⊚ f1 ≘ f2 & K ⊢ 𝐅+⦃T⦄ ≘ f1. + ∃∃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. 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 → + ∀b,f2,L,U. L ⊢ 𝐅+❪U❫ ≘ f2 → ∀f,K. ⇩*[b,f] L ≘ K → ∀T. ⇧*[f] T ≘ U → - ∀f1. f ~⊚ f1 ≘ f2 → K ⊢ 𝐅+⦃T⦄ ≘ f1. + ∀f1. f ~⊚ f1 ≘ f2 → K ⊢ 𝐅+❪T❫ ≘ f1. #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-. (* Note: this is used by rex_conf and might be modified *) lemma frees_inv_drops_next: - ∀f1,L1,T1. L1 ⊢ 𝐅+⦃T1⦄ ≘ f1 → - ∀I2,L2,V2,n. ⇩*[n] L1 ≘ L2.ⓑ{I2}V2 → + ∀f1,L1,T1. L1 ⊢ 𝐅+❪T1❫ ≘ f1 → + ∀I2,L2,V2,n. ⇩*[n] L1 ≘ L2.ⓑ[I2]V2 → ∀g1. ↑g1 = ⫱*[n] f1 → - ∃∃g2. L2 ⊢ 𝐅+⦃V2⦄ ≘ g2 & g2 ⊆ g1. + ∃∃g2. L2 ⊢ 𝐅+❪V2❫ ≘ g2 & g2 ⊆ g1. #f1 #L1 #T1 #H elim H -f1 -L1 -T1 [ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s lapply (isid_tls n … Hf1) -Hf1