]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma
bug fix in supclosure allows alternative definition of fpbs !
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop.ma
index d209bd4b0f094b93a5e2109e6a084306e458562c..94c6b0f733c2a3037932f8fdfd2c3acc6de847a7 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/lstar.ma".
 include "basic_2/notation/relations/rdrop_4.ma".
 include "basic_2/grammar/lenv_length.ma".
 include "basic_2/grammar/lenv_weight.ma".
@@ -20,7 +21,7 @@ include "basic_2/relocation/lift.ma".
 (* LOCAL ENVIRONMENT SLICING ************************************************)
 
 (* Basic_1: includes: drop_skip_bind *)
-inductive ldrop: nat → nat → relation lenv ≝
+inductive ldrop: relation4 nat nat lenv lenv ≝
 | ldrop_atom : ∀d. ldrop d 0 (⋆) (⋆)
 | ldrop_pair : ∀L,I,V. ldrop 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V)
 | ldrop_ldrop: ∀L1,L2,I,V,e. ldrop 0 e L1 L2 → ldrop 0 (e + 1) (L1. ⓑ{I} V) L2
@@ -234,6 +235,16 @@ lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R).
 ]
 qed.
 
+lemma l_deliftable_sn_llstar: ∀R. l_deliftable_sn R →
+                              ∀l. l_deliftable_sn (llstar … R l).
+#R #HR #l #L #U1 #U2 #H @(lstar_ind_r … l U2 H) -l -U2
+[ /2 width=3/
+| #l #U #U2 #_ #HU2 #IHU1 #K #d #e #HLK #T1 #HTU1
+  elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
+  elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5/
+]
+qed.
+
 (* Basic forvard lemmas *****************************************************)
 
 (* Basic_1: was: drop_S *)