]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground/relocation/pr_after_ist.ma
partial update in static_2
[helm.git] / matita / matita / contribs / lambdadelta / ground / relocation / pr_after_ist.ma
index f25cd9b512f3b5de781fe02b5b4638992622e9f2..2128885217bc56f8596871d7caf95527197613d6 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 include "ground/relocation/pr_pat_lt.ma".
+include "ground/relocation/pr_nat.ma".
 include "ground/relocation/pr_ist.ma".
 include "ground/relocation/pr_after_pat.ma".
 
@@ -52,6 +53,14 @@ lemma pr_after_des_ist_pat:
 /3 width=8 by pr_after_des_pat, ex2_intro/
 qed-.
 
+lemma pr_after_des_ist_nat:
+      ∀f1,l1,l2. @↑❪l1, f1❫ ≘ l2 → ∀f2. 𝐓❪f2❫ → ∀f. f2 ⊚ f1 ≘ f →
+      ∃∃l. @↑❪l2, f2❫ ≘ l & @↑❪l1, f❫ ≘ l.
+#f1 #l1 #l2 #H1 #f2 #H2 #f #Hf
+elim (pr_after_des_ist_pat … H1 … H2 … Hf) -f1 -H2
+/2 width=3 by ex2_intro/
+qed-.
+
 (* Inversions with pr_ist ***************************************************)
 
 (*** after_inv_istot *)