]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isfin.ma
index b73e34cea2fcba9b4fcd4d5dcd5022516adf5389..3a9d2656c7fbfacb4a982107d22334fc88582d38 100644 (file)
@@ -18,32 +18,30 @@ 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â\9dªfâ\9d« â\89\98 n.
 
 interpretation "test for finite colength (rtmap)"
    'IsFinite f = (isfin f).
 
 (* Basic eliminators ********************************************************)
 
-lemma isfin_ind (R:predicate rtmap): (â\88\80f.  ð\9d\90\88â¦\83fâ¦\84 → R f) →
-                                     (â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 R f â\86\92 R (â\86\91f)) →
-                                     (â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 R f â\86\92 R (⫯f)) →
-                                     â\88\80f. ð\9d\90\85â¦\83fâ¦\84 → R f.
+lemma isfin_ind (R:predicate rtmap): (â\88\80f.  ð\9d\90\88â\9dªfâ\9d« → R f) →
+                                     (â\88\80f. ð\9d\90\85â\9dªfâ\9d« â\86\92 R f â\86\92 R (⫯f)) →
+                                     (â\88\80f. ð\9d\90\85â\9dªfâ\9d« â\86\92 R f â\86\92 R (â\86\91f)) →
+                                     â\88\80f. ð\9d\90\85â\9dªfâ\9d« → R f.
 #R #IH1 #IH2 #IH3 #f #H elim H -H
 #n #H elim H -f -n /3 width=2 by ex_intro/
 qed-.
 
 (* Basic inversion lemmas ***************************************************)
 
-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/
+lemma isfin_inv_push: ∀g. 𝐅❪g❫ → ∀f. ⫯f = g → 𝐅❪f❫.
+#g * /3 width=4 by fcla_inv_px, 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/
+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 properties *********************************************************)
@@ -55,30 +53,44 @@ qed-.
 lemma isfin_eq_repl_fwd: eq_repl_fwd … isfin.
 /3 width=3 by isfin_eq_repl_back, eq_repl_sym/ qed-.
 
-lemma isfin_isid: â\88\80f. ð\9d\90\88â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83fâ¦\84.
+lemma isfin_isid: â\88\80f. ð\9d\90\88â\9dªfâ\9d« â\86\92 ð\9d\90\85â\9dªfâ\9d«.
 /3 width=2 by fcla_isid, ex_intro/ qed.
 
-lemma isfin_push: â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83â\86\91fâ¦\84.
+lemma isfin_push: â\88\80f. ð\9d\90\85â\9dªfâ\9d« â\86\92 ð\9d\90\85â\9dªâ«¯fâ\9d«.
 #f * /3 width=2 by fcla_push, ex_intro/
 qed.
 
-lemma isfin_next: â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83⫯fâ¦\84.
+lemma isfin_next: â\88\80f. ð\9d\90\85â\9dªfâ\9d« â\86\92 ð\9d\90\85â\9dªâ\86\91\9d«.
 #f * /3 width=2 by fcla_next, ex_intro/
 qed.
 
-lemma isfin_tl: ∀f. 𝐅⦃f⦄ → 𝐅⦃⫱f⦄.
+(* 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: â\88\80f. ð\9d\90\85â¦\83⫱fâ¦\84 â\86\92 ð\9d\90\85â¦\83fâ¦\84.
+lemma isfin_inv_tl: â\88\80f. ð\9d\90\85â\9dªâ«±fâ\9d« â\86\92 ð\9d\90\85â\9dªfâ\9d«.
 #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: â\88\80n,f. ð\9d\90\85â¦\83⫱*[n]fâ¦\84 â\86\92 ð\9d\90\85â¦\83fâ¦\84.
+lemma isfin_inv_tls: â\88\80n,f. ð\9d\90\85â\9dªâ«±*[n]fâ\9d« â\86\92 ð\9d\90\85â\9dªfâ\9d«.
 #n elim n -n /3 width=1 by isfin_inv_tl/
 qed-.