]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/s_transition/fqu_length.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / s_transition / fqu_length.ma
index 712777ec45c7a54126c3d53ec8b2db7c2beb1fa2..467283ff1f6975a9fd6f0ec26db748b22aac0b88 100644 (file)
@@ -29,18 +29,3 @@ lemma fqu_fwd_length_lref1: ∀b,G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐[b] ⦃G2
                             |L2| < |L1|.
 /2 width=8 by fqu_fwd_length_lref1_aux/
 qed-.
-
-(* Inversion lemmas with length for local environments **********************)
-
-fact fqu_inv_eq_aux: ∀b,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐[b] ⦃G2, L2, T2⦄ →
-                     G1 = G2 → |L1| = |L2| → T1 = T2 → ⊥.
-#b #G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2
-[1: #I #G #L #V #_ #H elim (succ_inv_refl_sn … H)
-|6: #I #G #L #T #U #_ #_ #H elim (succ_inv_refl_sn … H)
-]
-/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x/
-qed-.
-
-lemma fqu_inv_eq: ∀b,G,L1,L2,T. ⦃G, L1, T⦄ ⊐[b] ⦃G, L2, T⦄ → |L1| = |L2| → ⊥.
-#b #G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *)
-qed-.