]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_drops.ma
- the trace is explicit in all auto tactics with depth > 1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsubc_drops.ma
index f76065b8001500bad8affe8546ffa20a049b1480..df89d984f5342225469af09ad40e7b5a936e0b30 100644 (file)
@@ -20,11 +20,11 @@ include "basic_2/computation/lsubc_drop.ma".
 
 (* Basic_1: was: csubc_drop1_conf_rev *)
 lemma drops_lsubc_trans: ∀RR,RS,RP. gcp RR RS RP →
-                         ∀G,L1,K1,des. ⬇*[Ⓕ, des] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
-                         ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, des] L2 ≡ K2.
-#RR #RS #RP #Hgcp #G #L1 #K1 #des #H elim H -L1 -K1 -des
+                         ∀G,L1,K1,cs. ⬇*[Ⓕ, cs] L1 ≡ K1 → ∀K2. G ⊢ K1 ⫃[RP] K2 →
+                         ∃∃L2. G ⊢ L1 ⫃[RP] L2 & ⬇*[Ⓕ, cs] L2 ≡ K2.
+#RR #RS #RP #Hgcp #G #L1 #K1 #cs #H elim H -L1 -K1 -cs
 [ /2 width=3 by drops_nil, ex2_intro/
-| #L1 #L #K1 #des #l #m #_ #HLK1 #IHL #K2 #HK12
+| #L1 #L #K1 #cs #l #m #_ #HLK1 #IHL #K2 #HK12
   elim (drop_lsubc_trans … Hgcp … HLK1 … HK12) -Hgcp -K1 #K #HLK #HK2
   elim (IHL … HLK) -IHL -HLK /3 width=5 by drops_cons, ex2_intro/
 ]