]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma
update in basic_2 and ground_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isfin.ma
index 8a127b35998fd15115db47270d5ca55214fc04cf..29f6eca5a79212539437b59770c5fdd265a59cbc 100644 (file)
@@ -18,7 +18,7 @@ include "ground_2/relocation/rtmap_fcla.ma".
 (* RELOCATION MAP ***********************************************************)
 
 definition isfin: predicate rtmap ≝
-                  Î»f. â\88\83n. ð\9d\90\82â¦\83fâ¦\84 â\89¡ n.
+                  Î»f. â\88\83n. ð\9d\90\82â¦\83fâ¦\84 â\89\98 n.
 
 interpretation "test for finite colength (rtmap)"
    'IsFinite f = (isfin f).
@@ -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-.