]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/static/lsuba_aaa.ma
More work on inserting UnsafeCoerce in argument applications only when needed.
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / static / lsuba_aaa.ma
index 4f45ce712960cb2c5fd00f8d7f629bb13d064036..66e802aaed6e8348ef63c69c2deba38a170b267a 100644 (file)
@@ -19,7 +19,7 @@ include "basic_2/static/lsuba_ldrop.ma".
 
 (* Properties concerning atomic arity assignment ****************************)
 
-lemma lsuba_aaa_conf: ∀L1,V,A. L1 ⊢ V ÷ A → ∀L2. L1 ÷⊑ L2 → L2 ⊢ V ÷ A.
+lemma lsuba_aaa_conf: ∀L1,V,A. L1 ⊢ V ⁝ A → ∀L2. L1 ⁝⊑ L2 → L2 ⊢ V ⁝ A.
 #L1 #V #A #H elim H -L1 -V -A
 [ //
 | #I #L1 #K1 #V1 #B #i #HLK1 #HV1B #IHV1 #L2 #HL12
@@ -36,7 +36,7 @@ lemma lsuba_aaa_conf: ∀L1,V,A. L1 ⊢ V ÷ A → ∀L2. L1 ÷⊑ L2 → L2 ⊢
 ]
 qed.
 
-lemma lsuba_aaa_trans: ∀L2,V,A. L2 ⊢ V ÷ A → ∀L1. L1 ÷⊑ L2 → L1 ⊢ V ÷ A.
+lemma lsuba_aaa_trans: ∀L2,V,A. L2 ⊢ V ⁝ A → ∀L1. L1 ⁝⊑ L2 → L1 ⊢ V ⁝ A.
 #L2 #V #A #H elim H -L2 -V -A
 [ //
 | #I #L2 #K2 #V2 #B #i #HLK2 #HV2B #IHV2 #L1 #HL12