X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Frtmap_isfin.ma;h=f79431a17303bd306c160ee66c8e122e7d02f162;hb=cc6fcb70ca4f3cf01205ed722d75a2fdb2aaf779;hp=875378b987c00d56b339ed8ffeb79163b6a15165;hpb=66962864d3703b8f3b44e95d32c03ed50ceee6f1;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 875378b98..f79431a17 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isfin.ma @@ -18,7 +18,7 @@ include "ground_2/relocation/rtmap_fcla.ma". (* RELOCATION MAP ***********************************************************) definition isfin: predicate rtmap ≝ - λf. ∃n. 𝐂⦃f⦄ ≡ n. + λf. ∃n. 𝐂⦃f⦄ ≘ n. interpretation "test for finite colength (rtmap)" 'IsFinite f = (isfin f). @@ -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: ∀g. 𝐅⦃g⦄ → ∀f. ↑f = g → 𝐅⦃f⦄. +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⦄. +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-. @@ -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: ∀f. 𝐅⦃f⦄ → 𝐅⦃↑f⦄. +lemma isfin_push: ∀f. 𝐅⦃f⦄ → 𝐅⦃⫯f⦄. #f * /3 width=2 by fcla_push, ex_intro/ qed. -lemma isfin_next: ∀f. 𝐅⦃f⦄ → 𝐅⦃⫯f⦄. +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⦄. +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⦄. +lemma isfin_inv_pushs: ∀n,g. 𝐅⦃⫯*[n]g⦄ → 𝐅⦃g⦄. #n elim n -n /3 width=3 by isfin_inv_push/ qed.