X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_isfin.ma;h=875378b987c00d56b339ed8ffeb79163b6a15165;hb=66962864d3703b8f3b44e95d32c03ed50ceee6f1;hp=8a127b35998fd15115db47270d5ca55214fc04cf;hpb=d64b4238ec803353f0a06f2aad25c173852b0526;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma index 8a127b359..875378b98 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma @@ -35,17 +35,15 @@ qed-. (* Basic inversion lemmas ***************************************************) +lemma isfin_inv_push: ∀g. 𝐅⦃g⦄ → ∀f. ↑f = g → 𝐅⦃f⦄. +#g * /3 width=4 by fcla_inv_px, ex_intro/ +qed-. + lemma isfin_inv_next: ∀g. 𝐅⦃g⦄ → ∀f. ⫯f = g → 𝐅⦃f⦄. #g * #n #H #f #H0 elim (fcla_inv_nx … H … H0) -g /2 width=2 by ex_intro/ qed-. -(* Basic forward lemmas *****************************************************) - -lemma isfin_fwd_push: ∀g. 𝐅⦃g⦄ → ∀f. ↑f = g → 𝐅⦃f⦄. -#g * /3 width=4 by fcla_inv_px, ex_intro/ -qed-. - (* Basic properties *********************************************************) lemma isfin_eq_repl_back: eq_repl_back … isfin. @@ -66,13 +64,33 @@ lemma isfin_next: ∀f. 𝐅⦃f⦄ → 𝐅⦃⫯f⦄. #f * /3 width=2 by fcla_next, ex_intro/ qed. +(* Properties with iterated push ********************************************) + +lemma isfin_pushs: ∀n,f. 𝐅⦃f⦄ → 𝐅⦃↑*[n]f⦄. +#n elim n -n /3 width=3 by isfin_push/ +qed. + +(* Inversion lemmas with iterated push **************************************) + +lemma isfin_inv_pushs: ∀n,g. 𝐅⦃↑*[n]g⦄ → 𝐅⦃g⦄. +#n elim n -n /3 width=3 by isfin_inv_push/ +qed. + +(* Properties with tail *****************************************************) + lemma isfin_tl: ∀f. 𝐅⦃f⦄ → 𝐅⦃⫱f⦄. #f elim (pn_split f) * #g #H #Hf destruct -/3 width=3 by isfin_fwd_push, isfin_inv_next/ +/3 width=3 by isfin_inv_push, isfin_inv_next/ qed. (* Inversion lemmas with tail ***********************************************) lemma isfin_inv_tl: ∀f. 𝐅⦃⫱f⦄ → 𝐅⦃f⦄. -#f elim (pn_split f) * /2 width=1 by isfin_next, isfin_push/ +#f elim (pn_split f) * /2 width=1 by isfin_next, isfin_push/ +qed-. + +(* Inversion lemmas with iterated tail **************************************) + +lemma isfin_inv_tls: ∀n,f. 𝐅⦃⫱*[n]f⦄ → 𝐅⦃f⦄. +#n elim n -n /3 width=1 by isfin_inv_tl/ qed-.