X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcir.ma;h=ba5ba7e7e8aeacef8f67be60d84b2f8cb2355143;hb=a02ba10c669642bd4b75a5b0ac9351c24ddb724a;hp=b315e29bc0fc6a8fc01454ffa268b990d142d592;hpb=1f30483032488ac4df2310b68fe8146e05524fec;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma index b315e29bc..ba5ba7e7e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma @@ -24,7 +24,7 @@ interpretation "irreducibility for context-sensitive reduction (term)" (* Basic inversion lemmas ***************************************************) -lemma cir_inv_delta: ∀G,L,K,V,i. ⇩[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐈⦃#i⦄ → ⊥. +lemma cir_inv_delta: ∀G,L,K,V,i. ⬇[i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ ➡ 𝐈⦃#i⦄ → ⊥. /3 width=3 by crr_delta/ qed-. lemma cir_inv_ri2: ∀I,G,L,V,T. ri2 I → ⦃G, L⦄ ⊢ ➡ 𝐈⦃②{I}V.T⦄ → ⊥. @@ -43,7 +43,7 @@ qed-. lemma cir_inv_appl: ∀G,L,V,T. ⦃G, L⦄ ⊢ ➡ 𝐈⦃ⓐV.T⦄ → ∧∧ ⦃G, L⦄ ⊢ ➡ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ ➡ 𝐈⦃T⦄ & 𝐒⦃T⦄. -#G #L #V #T #HVT @and3_intro /3 width=1/ +#G #L #V #T #HVT @and3_intro /3 width=1 by crr_appl_sn, crr_appl_dx/ generalize in match HVT; -HVT elim T -T // * // #a * #U #T #_ #_ #H elim H -H /2 width=1 by crr_beta, crr_theta/ qed-.