]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isfin.ma
index 29f6eca5a79212539437b59770c5fdd265a59cbc..f79431a17303bd306c160ee66c8e122e7d02f162 100644 (file)
@@ -26,8 +26,8 @@ interpretation "test for finite colength (rtmap)"
 (* Basic eliminators ********************************************************)
 
 lemma isfin_ind (R:predicate rtmap): (∀f.  𝐈⦃f⦄ → R f) →
-                                     (∀f. 𝐅⦃f⦄ → R f → R (↑f)) →
                                      (∀f. 𝐅⦃f⦄ → R f → R (⫯f)) →
+                                     (∀f. 𝐅⦃f⦄ → R f → R (↑f)) →
                                      ∀f. 𝐅⦃f⦄ → R f.
 #R #IH1 #IH2 #IH3 #f #H elim H -H
 #n #H elim H -f -n /3 width=2 by ex_intro/
@@ -35,11 +35,11 @@ qed-.
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma isfin_inv_push: â\88\80g. ð\9d\90\85â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g → 𝐅⦃f⦄.
+lemma isfin_inv_push: â\88\80g. ð\9d\90\85â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐅⦃f⦄.
 #g * /3 width=4 by fcla_inv_px, ex_intro/
 qed-.
 
-lemma isfin_inv_next: â\88\80g. ð\9d\90\85â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐅⦃f⦄.
+lemma isfin_inv_next: â\88\80g. ð\9d\90\85â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g → 𝐅⦃f⦄.
 #g * #n #H #f #H0 elim (fcla_inv_nx … H … H0) -g
 /2 width=2 by ex_intro/
 qed-.
@@ -56,23 +56,23 @@ lemma isfin_eq_repl_fwd: eq_repl_fwd … isfin.
 lemma isfin_isid: ∀f. 𝐈⦃f⦄ → 𝐅⦃f⦄.
 /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⦄.
+lemma isfin_push: â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83⫯f⦄.
 #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⦄.
+lemma isfin_next: â\88\80f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83â\86\91f⦄.
 #f * /3 width=2 by fcla_next, ex_intro/
 qed.
 
 (* Properties with iterated push ********************************************)
 
-lemma isfin_pushs: â\88\80n,f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83â\86\91*[n]f⦄.
+lemma isfin_pushs: â\88\80n,f. ð\9d\90\85â¦\83fâ¦\84 â\86\92 ð\9d\90\85â¦\83⫯*[n]f⦄.
 #n elim n -n /3 width=3 by isfin_push/
 qed.
 
 (* Inversion lemmas with iterated push **************************************)
 
-lemma isfin_inv_pushs: â\88\80n,g. ð\9d\90\85â¦\83â\86\91*[n]g⦄ → 𝐅⦃g⦄.
+lemma isfin_inv_pushs: â\88\80n,g. ð\9d\90\85â¦\83⫯*[n]g⦄ → 𝐅⦃g⦄.
 #n elim n -n /3 width=3 by isfin_inv_push/
 qed.