]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_istot.ma
- improved lfpr_lfpr
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / nstream_istot.ma
index 0001b90b9c06cff094db74fc2b49b3d54da71815..39211d7f2c74ece3b91c7f155788ca883719e581 100644 (file)
@@ -18,7 +18,7 @@ include "ground_2/relocation/rtmap_istot.ma".
 
 (* RELOCATION N-STREAM ******************************************************)
 
-let rec apply (i: nat) on i: rtmap → nat ≝ ?.
+rec definition apply (i: nat) on i: rtmap → nat ≝ ?.
 * #n #f cases i -i
 [ @n
 | #i lapply (apply i f) -apply -i -f
@@ -48,7 +48,8 @@ lemma at_istot: ∀f. 𝐓⦃f⦄.
 /2 width=2 by ex_intro/ qed.
 
 lemma at_plus2: ∀f,i1,i,n,m. @⦃i1, n@f⦄ ≡ i → @⦃i1, (m+n)@f⦄ ≡ m+i.
-#f #i1 #i #n #m #H elim m -m /2 width=5 by at_next/
+#f #i1 #i #n #m #H elim m -m //
+#m <plus_S1 /2 width=5 by at_next/ (**) (* full auto fails *)
 qed.
 
 (* Specific inversion lemmas on at ******************************************)
@@ -98,8 +99,6 @@ lemma apply_S1: ∀f,i. (⫯f)@❴i❵ = ⫯(f@❴i❵).
 qed.
 
 (* Main inversion lemmas ****************************************************)
-(*
-lemma apply_inj_aux: ∀f1,f2,j1,j2,i1,i2. f1@❴i1❵ = j1 → f2@❴i2❵ = j2 →
-                     j1 = j2 → f1 ≗ f2 → i1 = i2.
-/2 width=6 by at_inj/ qed-.
-*)
+
+theorem apply_inj: ∀f,i1,i2,j. f@❴i1❵ = j → f@❴i2❵ = j → i1 = i2.
+/2 width=4 by at_inj/ qed-.