]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_isuni.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_isuni.ma
index 273793d9320c6ec9806c981a4e5ffcf12ce79bb6..38443bf59525e6bef7d792921897cfd695c65b1f 100644 (file)
@@ -19,7 +19,7 @@ include "ground_2/relocation/rtmap_isfin.ma".
 
 inductive isuni: predicate rtmap ≝
 | isuni_isid: ∀f. 𝐈⦃f⦄ → isuni f
-| isuni_next: â\88\80f. isuni f â\86\92 â\88\80g. â«¯f = g → isuni g
+| isuni_next: â\88\80f. isuni f â\86\92 â\88\80g. â\86\91f = g → isuni g
 .
 
 interpretation "test for uniformity (rtmap)"
@@ -27,26 +27,26 @@ interpretation "test for uniformity (rtmap)"
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma isuni_inv_push: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g → 𝐈⦃f⦄.
+lemma isuni_inv_push: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐈⦃f⦄.
 #g * -g /2 width=3 by isid_inv_push/
 #f #_ #g #H #x #Hx destruct elim (discr_push_next … Hx)
 qed-.
 
-lemma isuni_inv_next: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐔⦃f⦄.
+lemma isuni_inv_next: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g → 𝐔⦃f⦄.
 #g * -g #f #Hf
 [ #x #Hx elim (isid_inv_next … Hf … Hx)
 | #g #H #x #Hx destruct /2 width=1 by injective_push/
 ]
 qed-.
 
-lemma isuni_split: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 (â\88\83â\88\83f. ð\9d\90\88â¦\83fâ¦\84 & â\86\91f = g) â\88¨ (â\88\83â\88\83f.ð\9d\90\94â¦\83fâ¦\84 & â«¯f = g).
+lemma isuni_split: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 (â\88\83â\88\83f. ð\9d\90\88â¦\83fâ¦\84 & â«¯f = g) â\88¨ (â\88\83â\88\83f.ð\9d\90\94â¦\83fâ¦\84 & â\86\91f = g).
 #g #H elim (pn_split g) * #f #Hf
 /4 width=3 by isuni_inv_next, isuni_inv_push, or_introl, or_intror, ex2_intro/
 qed-.
 
 (* basic forward lemmas *****************************************************)
 
-lemma isuni_fwd_push: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â\86\91f = g → 𝐔⦃f⦄.
+lemma isuni_fwd_push: â\88\80g. ð\9d\90\94â¦\83gâ¦\84 â\86\92 â\88\80f. â«¯f = g → 𝐔⦃f⦄.
 /3 width=3 by isuni_inv_push, isuni_isid/ qed-.
 
 (* Forward lemmas with test for finite colength *****************************)