]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/steps/rtc_ist.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / steps / rtc_ist.ma
index 4b4c8ac8fcf4d8812f2e38a6b9bb899ca4e08cc6..3f6f70dae955840cadc0f8ce51956703cf83656a 100644 (file)
@@ -25,32 +25,32 @@ interpretation "test for t-transition counter (rtc)"
 
 (* Basic properties *********************************************************)
 
-lemma isr_00: ð\9d\90\93â¦\830,ð\9d\9f\98ð\9d\9f\98â¦\84.
+lemma isr_00: ð\9d\90\93â\9dª0,ð\9d\9f\98ð\9d\9f\98â\9d«.
 // qed.
 
-lemma ist_01: ð\9d\90\93â¦\831,ð\9d\9f\98ð\9d\9f\99â¦\84.
+lemma ist_01: ð\9d\90\93â\9dª1,ð\9d\9f\98ð\9d\9f\99â\9d«.
 // qed.
 
 (* Basic inversion properties ***********************************************)
 
-lemma ist_inv_00: â\88\80n. ð\9d\90\93â¦\83n,ð\9d\9f\98ð\9d\9f\98â¦\84 → 0 = n.
+lemma ist_inv_00: â\88\80n. ð\9d\90\93â\9dªn,ð\9d\9f\98ð\9d\9f\98â\9d« → 0 = n.
 #n #H destruct //
 qed-.
 
-lemma ist_inv_01: â\88\80n. ð\9d\90\93â¦\83n,ð\9d\9f\98ð\9d\9f\99â¦\84 → 1 = n.
+lemma ist_inv_01: â\88\80n. ð\9d\90\93â\9dªn,ð\9d\9f\98ð\9d\9f\99â\9d« → 1 = n.
 #n #H destruct //
 qed-.
 
-lemma ist_inv_10: â\88\80n. ð\9d\90\93â¦\83n,ð\9d\9f\99ð\9d\9f\98â¦\84 → ⊥.
+lemma ist_inv_10: â\88\80n. ð\9d\90\93â\9dªn,ð\9d\9f\99ð\9d\9f\98â\9d« → ⊥.
 #h #H destruct
 qed-.
 
 (* Main inversion properties ************************************************)
 
-theorem ist_inj: â\88\80n1,n2,c. ð\9d\90\93â¦\83n1,câ¦\84 â\86\92 ð\9d\90\93â¦\83n2,câ¦\84 → n1 = n2.
+theorem ist_inj: â\88\80n1,n2,c. ð\9d\90\93â\9dªn1,câ\9d« â\86\92 ð\9d\90\93â\9dªn2,câ\9d« → n1 = n2.
 #n1 #n2 #c #H1 #H2 destruct //
 qed-.
 
-theorem ist_mono: â\88\80n,c1,c2. ð\9d\90\93â¦\83n,c1â¦\84 â\86\92 ð\9d\90\93â¦\83n,c2â¦\84 → c1 = c2.
+theorem ist_mono: â\88\80n,c1,c2. ð\9d\90\93â\9dªn,c1â\9d« â\86\92 ð\9d\90\93â\9dªn,c2â\9d« → c1 = c2.
 #n #c1 #c2 #H1 #H2 destruct //
 qed-.