]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isid.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isid.ma
index 5ad6aa83c32edb6c8c37599dad3974473f5e4f8c..7fea33f0021f8a370c1d37b89d181a2f08f42e10 100644 (file)
@@ -18,7 +18,7 @@ include "ground_2/relocation/rtmap_tls.ma".
 (* RELOCATION MAP ***********************************************************)
 
 coinductive isid: predicate rtmap ≝
-| isid_push: â\88\80f,g. isid f â\86\92 â\86\91f = g → isid g
+| isid_push: â\88\80f,g. isid f â\86\92 â«¯f = g → isid g
 .
 
 interpretation "test for identity (rtmap)"
@@ -26,24 +26,26 @@ interpretation "test for identity (rtmap)"
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma isid_inv_gen: â\88\80g. ð\9d\90\88â¦\83gâ¦\84 â\86\92 â\88\83â\88\83f. ð\9d\90\88â¦\83fâ¦\84 & â\86\91f = g.
+lemma isid_inv_gen: â\88\80g. ð\9d\90\88â\9dªgâ\9d« â\86\92 â\88\83â\88\83f. ð\9d\90\88â\9dªfâ\9d« & â«¯f = g.
 #g * -g
 #f #g #Hf * /2 width=3 by ex2_intro/
 qed-.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma isid_inv_push: â\88\80g. ð\9d\90\88â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g â\86\92 ð\9d\90\88â¦\83fâ¦\84.
+lemma isid_inv_push: â\88\80g. ð\9d\90\88â\9dªgâ\9d« â\86\92 â\88\80f. â«¯f = g â\86\92 ð\9d\90\88â\9dªfâ\9d«.
 #g #H elim (isid_inv_gen … H) -H
 #f #Hf * -g #g #H >(injective_push … H) -H //
 qed-.
 
-lemma isid_inv_next: â\88\80g. ð\9d\90\88â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → ⊥.
+lemma isid_inv_next: â\88\80g. ð\9d\90\88â\9dªgâ\9d« â\86\92 â\88\80f. â\86\91f = g → ⊥.
 #g #H elim (isid_inv_gen … H) -H
 #f #Hf * -g #g #H elim (discr_next_push … H)
 qed-.
 
-let corec isid_inv_eq_repl: ∀f1,f2. 𝐈⦃f1⦄ → 𝐈⦃f2⦄ → f1 ≗ f2 ≝ ?.
+(* Main inversion lemmas ****************************************************)
+
+corec theorem isid_inv_eq_repl: ∀f1,f2. 𝐈❪f1❫ → 𝐈❪f2❫ → f1 ≡ f2.
 #f1 #f2 #H1 #H2
 cases (isid_inv_gen … H1) -H1
 cases (isid_inv_gen … H2) -H2
@@ -52,7 +54,7 @@ qed-.
 
 (* Basic properties *********************************************************)
 
-let corec isid_eq_repl_back: eq_repl_back … isid ≝ ?.
+corec lemma isid_eq_repl_back: eq_repl_back … isid.
 #f1 #H cases (isid_inv_gen … H) -H
 #g1 #Hg1 #H1 #f2 #Hf cases (eq_inv_px … Hf … H1) -f1
 /3 width=3 by isid_push/
@@ -63,12 +65,39 @@ lemma isid_eq_repl_fwd: eq_repl_fwd … isid.
 
 (* Alternative definition ***************************************************)
 
-let corec eq_push_isid: ∀f. ↑f ≗ f → 𝐈⦃f⦄ ≝ ?.
+corec lemma eq_push_isid: ∀f. ⫯f ≡ f → 𝐈❪f❫.
 #f #H cases (eq_inv_px … H) -H /4 width=3 by isid_push, eq_trans/
 qed.
 
-let corec eq_push_inv_isid: ∀f. 𝐈⦃f⦄ → ↑f ≗ f ≝ ?.
+corec lemma eq_push_inv_isid: ∀f. 𝐈❪f❫ → ⫯f ≡ f.
 #f * -f
 #f #g #Hf #Hg @(eq_push … Hg) [2: @eq_push_inv_isid // | skip ]
 @eq_f //
 qed-.
+
+(* Properties with iterated push ********************************************)
+
+lemma isid_pushs: ∀n,f. 𝐈❪f❫ → 𝐈❪⫯*[n]f❫.
+#n elim n -n /3 width=3 by isid_push/
+qed.
+
+(* Inversion lemmas with iterated push **************************************)
+
+lemma isid_inv_pushs: ∀n,g. 𝐈❪⫯*[n]g❫ → 𝐈❪g❫.
+#n elim n -n /3 width=3 by isid_inv_push/
+qed.
+
+(* Properties with tail *****************************************************)
+
+lemma isid_tl: ∀f. 𝐈❪f❫ → 𝐈❪⫱f❫.
+#f cases (pn_split f) * #g * -f #H
+[ /2 width=3 by isid_inv_push/
+| elim (isid_inv_next … H) -H //
+]
+qed.
+
+(* Properties with iterated tail ********************************************)
+
+lemma isid_tls: ∀n,g. 𝐈❪g❫ → 𝐈❪⫱*[n]g❫.
+#n elim n -n /3 width=1 by isid_tl/
+qed.