X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fgrammar%2Flenv.ma;h=80738ebbde5c366fa3ec33662696af5898d44aa0;hb=7d99a19985ae7ca20845d0a875e32f23ba06e536;hp=c34238632d50c18f97f9044bb0295291fa3a3b04;hpb=a2092f7ba4c7c566ea90653ff57e4623ab94d8d5;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma index c34238632..80738ebbd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/lenv.ma @@ -40,12 +40,22 @@ interpretation "abstraction (local environment)" (* Basic properties *********************************************************) -axiom eq_lenv_dec: ∀L1,L2:lenv. Decidable (L1 = L2). +lemma eq_lenv_dec: ∀L1,L2:lenv. Decidable (L1 = L2). +#L1 elim L1 -L1 [| #L1 #I1 #V1 #IHL1 ] * [2,4: #L2 #I2 #V2 ] +[1,4: @or_intror #H destruct +| elim (eq_bind2_dec I1 I2) #HI + [ elim (eq_term_dec V1 V2) #HV + [ elim (IHL1 L2) -IHL1 /2 width=1 by or_introl/ #HL ] + ] + @or_intror #H destruct /2 width=1 by/ +| /2 width=1 by or_introl/ +] +qed-. (* Basic inversion lemmas ***************************************************) -lemma destruct_lpair_lpair: ∀I1,I2,L1,L2,V1,V2. L1.ⓑ{I1}V1 = L2.ⓑ{I2}V2 → - ∧∧L1 = L2 & I1 = I2 & V1 = V2. +fact destruct_lpair_lpair_aux: ∀I1,I2,L1,L2,V1,V2. L1.ⓑ{I1}V1 = L2.ⓑ{I2}V2 → + ∧∧L1 = L2 & I1 = I2 & V1 = V2. #I1 #I2 #L1 #L2 #V1 #V2 #H destruct /2 width=1 by and3_intro/ qed-. @@ -53,6 +63,9 @@ lemma discr_lpair_x_xy: ∀I,V,L. L = L.ⓑ{I}V → ⊥. #I #V #L elim L -L [ #H destruct | #L #J #W #IHL #H - elim (destruct_lpair_lpair … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *) + elim (destruct_lpair_lpair_aux … H) -H #H1 #H2 #H3 destruct /2 width=1 by/ (**) (* destruct lemma needed *) ] qed-. + +lemma discr_lpair_xy_x: ∀I,V,L. L.ⓑ{I}V = L→ ⊥. +/2 width=4 by discr_lpair_x_xy/ qed-.