]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/syntax/theq_simple.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / theq_simple.ma
index 9ad5ed55081e8f3ed030a6d2af2a124ae85c7b19..d9f59e4a51ccc87b2ab5173bdc70ba39698259e0 100644 (file)
@@ -20,12 +20,12 @@ include "static_2/syntax/theq.ma".
 (* Properies with simple (neutral) terms ************************************)
 
 (* Basic_2A1: was: simple_tsts_repl_dx *)
-lemma simple_theq_repl_dx: ∀h,o,T1,T2. T1 ⩳[h, o] T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
-#h #o #T1 #T2 * -T1 -T2 //
+lemma simple_theq_repl_dx: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
+#T1 #T2 * -T1 -T2 //
 #I #V1 #V2 #T1 #T2 #H
 elim (simple_inv_pair … H) -H #J #H destruct //
 qed-.
 
 (* Basic_2A1: was: simple_tsts_repl_sn *)
-lemma simple_theq_repl_sn: ∀h,o,T1,T2. T1 ⩳[h, o] T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
-/3 width=5 by simple_theq_repl_dx, theq_sym/ qed-.
+lemma simple_theq_repl_sn: ∀T1,T2. T1 ⩳ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
+/3 width=3 by simple_theq_repl_dx, theq_sym/ qed-.