X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fdecidable_kit%2Flist_aux.ma;h=077c40477386f30feac3310e4ae1d06ef2529b7a;hb=8ee0e6f729105eaf1907de0baef22e170b0d17b3;hp=d641d9018aaec1560ea81b387ceb80b078ba5730;hpb=7c84225a8e472e754f1baf8f8b37f8627c8da6fa;p=helm.git diff --git a/matita/library/decidable_kit/list_aux.ma b/matita/library/decidable_kit/list_aux.ma index d641d9018..077c40477 100644 --- a/matita/library/decidable_kit/list_aux.ma +++ b/matita/library/decidable_kit/list_aux.ma @@ -49,9 +49,9 @@ lemma list_ind2 : P l1 l2. intros (T1 T2 l1 l2 P Hl Pnil Pcons); generalize in match Hl; clear Hl; generalize in match l2; clear l2; -elim l1 1 (l2 x1); [ cases l2; intros (Hl); [assumption| destruct Hl]] +elim l1 1 (l2 x1); [ cases l2; intros (Hl); [assumption| simplify in Hl; destruct Hl]] intros 3 (tl1 IH l2); cases l2; [1: simplify; intros 1 (Hl); destruct Hl] -intros 1 (Hl); apply Pcons; apply IH; destruct Hl; assumption; +intros 1 (Hl); apply Pcons; apply IH; simplify in Hl; destruct Hl; assumption; qed. lemma eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l. @@ -71,7 +71,7 @@ lemma lcmp_length : lcmp ? l1 l2 = true → length ? l1 = length ? l2. intros 2 (d l1); elim l1 1 (l2 x1); [1: cases l2; simplify; intros; [reflexivity|destruct H] -|2: intros 3 (tl1 IH l2); cases (l2); intros; [1:destruct H] +|2: intros 3 (tl1 IH l2); cases (l2); intros; [1:simplify in H; destruct H] simplify; (* XXX la apply non fa simplify? *) apply congr_S; apply (IH l); (* XXX qualcosa di enorme e' rotto! la regola di convertibilita?! *) @@ -90,19 +90,15 @@ cases c; intros (H); [ apply reflect_true | apply reflect_false ] rewrite > (IH H2); rewrite > (b2pT ? ? (eqP d ? ?) H1); reflexivity | generalize in match H; clear H; generalize in match l2; clear l2; elim l1 1 (l1 x1); - [ cases l1; [intros; destruct H | unfold Not; intros; destruct H1;] + [ cases l1; simplify; [intros; destruct H | unfold Not; intros; destruct H1;] | intros 3 (tl1 IH l2); cases l2; [ unfold Not; intros; destruct H1; | simplify; intros; cases (b2pT ? ? (andbPF ? ?) (p2bT ? ? (negbP ?) H)); clear H; [ intros; lapply (b2pF ? ? (eqP d ? ?) H1) as H'; clear H1; - (* XXX destruct H; *) - change in H' with (Not ((match (x1::tl1) with [nil⇒x1|(cons x1 _) ⇒ x1]) = s)); - rewrite > H in H'; simplify in H'; apply H'; reflexivity; - | intros; lapply (IH ? H1) as H'; - (* XXX destruct H1 *) - change in H' with (Not ((match (x1::tl1) with [nil⇒tl1|(cons _ l) ⇒ l]) = l)); - rewrite > H in H'; simplify in H'; apply H'; reflexivity;]]]] + destruct H; apply H'; reflexivity; + | intros; lapply (IH ? H1) as H'; destruct H; + apply H'; reflexivity;]]]] qed. definition list_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (lcmpP d).