X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Fterm.ma;h=6e8370d497251de2cfd69bfba097e5418823bb77;hb=011cf6478141e69822a5b40933f2444d0522532f;hp=32f8272dcd217727db74222bd06063124189020d;hpb=fcd30dfead2fbc2889aa993fba0577dce8a90c88;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma index 32f8272dc..6e8370d49 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/grammar/term.ma @@ -81,24 +81,32 @@ lemma discr_tpair_xy_y: ∀I,V,T. ②{I} V. T = T → False. ] qed-. -lemma eq_false_inv_tpair: ∀I,V1,T1,V2,T2. - (②{I} V1. T1 = ②{I} V2. T2 → False) → - (V1 = V2 → False) ∨ (V1 = V2 ∧ (T1 = T2 → False)). +lemma eq_false_inv_tpair_sn: ∀I,V1,T1,V2,T2. + (②{I} V1. T1 = ②{I} V2. T2 → False) → + (V1 = V2 → False) ∨ (V1 = V2 ∧ (T1 = T2 → False)). #I #V1 #T1 #V2 #T2 #H elim (term_eq_dec V1 V2) /3 width=1/ #HV12 destruct @or_intror @conj // #HT12 destruct /2 width=1/ qed-. +lemma eq_false_inv_tpair_dx: ∀I,V1,T1,V2,T2. + (②{I} V1. T1 = ②{I} V2. T2 → False) → + (T1 = T2 → False) ∨ (T1 = T2 ∧ (V1 = V2 → False)). +#I #V1 #T1 #V2 #T2 #H +elim (term_eq_dec T1 T2) /3 width=1/ #HT12 destruct +@or_intror @conj // #HT12 destruct /2 width=1/ +qed-. + lemma eq_false_inv_beta: ∀V1,V2,W1,W2,T1,T2. (ⓐV1. ⓛW1. T1 = ⓐV2. ⓛW2 .T2 →False) → (W1 = W2 → False) ∨ (W1 = W2 ∧ (ⓓV1. T1 = ⓓV2. T2 → False)). #V1 #V2 #W1 #W2 #T1 #T2 #H -elim (eq_false_inv_tpair … H) -H +elim (eq_false_inv_tpair_sn … H) -H [ #HV12 elim (term_eq_dec W1 W2) /3 width=1/ #H destruct @or_intror @conj // #H destruct /2 width=1/ | * #H1 #H2 destruct - elim (eq_false_inv_tpair … H2) -H2 /3 width=1/ + elim (eq_false_inv_tpair_sn … H2) -H2 /3 width=1/ * #H #HT12 destruct @or_intror @conj // #H destruct /2 width=1/ ]