]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isid.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isid.ma
index 65e111bbf9d3e209bbb6516965ff215c934dd339..4fa424d494f1f535e3eedbba964023dbb8f046f5 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,26 +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â¦\83gâ¦\84 â\86\92 â\88\83â\88\83f. ð\9d\90\88â¦\83fâ¦\84 & â«¯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 → 𝐈⦃f⦄.
+lemma isid_inv_push: â\88\80g. ð\9d\90\88â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐈⦃f⦄.
 #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â¦\83gâ¦\84 â\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-.
 
 (* Main inversion lemmas ****************************************************)
 
-corec theorem isid_inv_eq_repl: â\88\80f1,f2. ð\9d\90\88â¦\83f1â¦\84 â\86\92 ð\9d\90\88â¦\83f2â¦\84 â\86\92 f1 â\89\97 f2.
+corec theorem isid_inv_eq_repl: â\88\80f1,f2. ð\9d\90\88â¦\83f1â¦\84 â\86\92 ð\9d\90\88â¦\83f2â¦\84 â\86\92 f1 â\89¡ f2.
 #f1 #f2 #H1 #H2
 cases (isid_inv_gen … H1) -H1
 cases (isid_inv_gen … H2) -H2
@@ -65,11 +65,11 @@ lemma isid_eq_repl_fwd: eq_repl_fwd … isid.
 
 (* Alternative definition ***************************************************)
 
-corec lemma eq_push_isid: â\88\80f. â\86\91f â\89\97 f → 𝐈⦃f⦄.
+corec lemma eq_push_isid: â\88\80f. â«¯f â\89¡ f → 𝐈⦃f⦄.
 #f #H cases (eq_inv_px … H) -H /4 width=3 by isid_push, eq_trans/
 qed.
 
-corec lemma eq_push_inv_isid: â\88\80f. ð\9d\90\88â¦\83fâ¦\84 â\86\92 â\86\91f â\89\97 f.
+corec lemma eq_push_inv_isid: â\88\80f. ð\9d\90\88â¦\83fâ¦\84 â\86\92 â«¯f â\89¡ f.
 #f * -f
 #f #g #Hf #Hg @(eq_push … Hg) [2: @eq_push_inv_isid // | skip ]
 @eq_f //
@@ -77,13 +77,13 @@ qed-.
 
 (* Properties with iterated push ********************************************)
 
-lemma isid_pushs: â\88\80n,f. ð\9d\90\88â¦\83fâ¦\84 â\86\92 ð\9d\90\88â¦\83â\86\91*[n]f⦄.
+lemma isid_pushs: â\88\80n,f. ð\9d\90\88â¦\83fâ¦\84 â\86\92 ð\9d\90\88â¦\83⫯*[n]f⦄.
 #n elim n -n /3 width=3 by isid_push/
 qed.
 
 (* Inversion lemmas with iterated push **************************************)
 
-lemma isid_inv_pushs: â\88\80n,g. ð\9d\90\88â¦\83â\86\91*[n]g⦄ → 𝐈⦃g⦄.
+lemma isid_inv_pushs: â\88\80n,g. ð\9d\90\88â¦\83⫯*[n]g⦄ → 𝐈⦃g⦄.
 #n elim n -n /3 width=3 by isid_inv_push/
 qed.