X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Ffrees_fqup.ma;h=7046b8ce95aea63ed81a227d0d01c15646b2ae6f;hb=b4b5f03ffca4f250a1dc02f277b70e4f33ac8a9b;hp=8b182b560bc49d076061ae783fee7633ff120beb;hpb=3ef251397627da80aeea0cf08b053a4bc781ef88;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/frees_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/static/frees_fqup.ma index 8b182b560..7046b8ce9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/frees_fqup.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/frees_fqup.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "ground_2/relocation/rtmap_id.ma". +include "basic_2/relocation/drops.ma". include "basic_2/s_computation/fqup_weight.ma". include "basic_2/static/frees.ma". @@ -20,7 +20,7 @@ include "basic_2/static/frees.ma". (* Advanced properties ******************************************************) -(* Notes: this replaces lemma 1400 concluding the "big tree" theorem *) +(* Note: this replaces lemma 1400 concluding the "big tree" theorem *) lemma frees_total: ∀L,T. ∃f. L ⊢ 𝐅*⦃T⦄ ≡ f. #L #T @(fqup_wf_ind_eq … (⋆) L T) -L -T #G0 #L0 #T0 #IH #G #L * @@ -45,3 +45,59 @@ lemma frees_total: ∀L,T. ∃f. L ⊢ 𝐅*⦃T⦄ ≡ f. ] ] qed-. + +(* Properties with plus-iterated supclosure *********************************) + +lemma frees_drops_next: ∀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. +#f1 #L1 #T1 #H elim H -f1 -L1 -T1 +[ #f1 #I1 #Hf1 #I2 #L2 #V2 #n #HL12 + elim (drops_inv_atom1 … HL12) -HL12 #H destruct +| #f1 #I1 #L1 #V1 #s #_ #IH #I2 #L2 #V2 * + [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1) + | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12 + #HL12 #g1 (injective_next … Hgf1) -g1 + /2 width=3 by sle_refl, ex2_intro/ + | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12 + #HL12 #g1 tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2 + /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/ + ] +| #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1 + lapply (sor_tls … Hf1 n) -Hf1