]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma
- commit of the "s_computation" component ...
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_after.ma
index 3db5d524f4dab642e516046b610acd92942913de..09bab5e109f5d2665cd9377585216bf4ef7aff61 100644 (file)
@@ -250,7 +250,7 @@ corec theorem after_trans2: ∀f1,f0,f4. f1 ⊚ f0 ≡ f4 →
 ]
 qed-.
 
-(* Main inversion lemmas on after *******************************************)
+(* Main inversion lemmas ****************************************************)
 
 corec theorem after_mono: ∀f1,f2,x,y. f1 ⊚ f2 ≡ x → f1 ⊚ f2 ≡ y → x ≗ y.
 #f1 #f2 #x #y * -f1 -f2 -x
@@ -313,6 +313,31 @@ qed-.
 lemma after_inv_isid3: ∀f1,f2,f. f1 ⊚ f2 ≡ f → 𝐈⦃f⦄ → 𝐈⦃f1⦄ ∧ 𝐈⦃f2⦄.
 /3 width=4 by after_fwd_isid2, after_fwd_isid1, conj/ qed-.
 
+(* Properties on isuni ******************************************************)
+
+lemma after_isid_isuni: ∀f1,f2. 𝐈⦃f2⦄ → 𝐔⦃f1⦄ → f1 ⊚ ⫯f2 ≡ ⫯f1.
+#f1 #f2 #Hf2 #H elim H -H
+/5 width=7 by isid_after_dx, after_eq_repl_back_2, after_next, after_push, eq_push_inv_isid/
+qed.
+
+lemma after_uni_next2: ∀f2. 𝐔⦃f2⦄ → ∀f1,f. ⫯f2 ⊚ f1 ≡ f → f2 ⊚ ⫯f1 ≡ f.
+#f2 #H elim H -f2
+[ #f2 #Hf2 #f1 #f #Hf
+  elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
+  /4 width=7 by after_isid_inv_sn, isid_after_sn, after_eq_repl_back_0, eq_next/
+| #f2 #_ #g2 #H2 #IH #f1 #f #Hf
+  elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct
+  /3 width=5 by after_next/
+]
+qed.
+
+(* Properties on uni ********************************************************)
+
+lemma after_uni: ∀n1,n2. 𝐔❴n1❵ ⊚ 𝐔❴n2❵ ≡ 𝐔❴n1+n2❵.
+@nat_elim2
+/4 width=5 by after_uni_next2, isid_after_sn, isid_after_dx, after_next/
+qed.
+
 (* Forward lemmas on at *****************************************************)
 
 lemma after_at_fwd: ∀i,i1,f. @⦃i1, f⦄ ≡ i → ∀f2,f1. f2 ⊚ f1 ≡ f →
@@ -374,6 +399,43 @@ lemma after_fwd_at1: ∀i,i2,i1,f,f2. @⦃i1, f⦄ ≡ i → @⦃i2, f2⦄ ≡ i
 ]
 qed-.
 
+(* Properties with at *******************************************************)
+
+lemma after_isuni_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 →
+                      ∀f. f2 ⊚ 𝐔❴i1❵ ≡ f → 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f.
+#i2 elim i2 -i2
+[ #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
+  lapply (after_isid_inv_dx … Hf ?) -Hf
+  /3 width=3 by isid_after_sn, after_eq_repl_back_0/
+| #i2 #IH #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
+  [ #g2 #j1 #Hg2 #H1 #H2 destruct
+    elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H destruct
+    /3 width=5 by after_next/
+  | #g2 #Hg2 #H2 destruct
+    elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
+    /3 width=5 by after_next/
+  ]
+]
+qed.
+
+lemma after_isuni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 →
+                      ∀f. 𝐔❴i2❵ ⊚ ⫱*[i2] f2 ≡ f → f2 ⊚ 𝐔❴i1❵ ≡ f.
+#i2 elim i2 -i2
+[ #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
+  lapply (after_isid_inv_sn … Hf ?) -Hf
+  /3 width=3 by isid_after_dx, after_eq_repl_back_0/
+| #i2 #IH #i1 #f2 #Hf2 #f #Hf
+  elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
+  elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
+  [ #g2 #j1 #Hg2 #H1 #H2 destruct /3 width=7 by after_push/
+  | #g2 #Hg2 #H2 destruct /3 width=5 by after_next/
+  ]
+]
+qed-.
+
 (* Forward lemmas on istot **************************************************)
 
 lemma after_istot_fwd: ∀f2,f1,f. f2 ⊚ f1 ≡ f → 𝐓⦃f2⦄ → 𝐓⦃f1⦄ → 𝐓⦃f⦄.