]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma
- more results on relocation
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isfin.ma
index b73e34cea2fcba9b4fcd4d5dcd5022516adf5389..875378b987c00d56b339ed8ffeb79163b6a15165 100644 (file)
@@ -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,9 +64,23 @@ 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 ***********************************************)
@@ -77,7 +89,7 @@ lemma isfin_inv_tl: āˆ€f. š…ā¦ƒā«±fā¦„ ā†’ š…ā¦ƒfā¦„.
 #f elim (pn_split f) * /2 width=1 by isfin_next, isfin_push/
 qed-.
 
-(* Inversion lemmas with tls ********************************************************)
+(* 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/