]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/list/list.ma
Generalize no more required for elim.
[helm.git] / helm / software / matita / library / list / list.ma
index ad44bd63a287846d917a08d278b57f01378fd205..f9f39bebf3b9dfde65c02733fbbb7955962deeaa 100644 (file)
@@ -178,8 +178,7 @@ lemma list_ind2 :
   (∀tl1,tl2,hd1,hd2. P tl1 tl2 â†’ P (hd1::tl1) (hd2::tl2)) â†’ 
   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| simplify in Hl; destruct Hl]]
+elim l1 in Hl l2 âŠ¢ % 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; simplify in Hl; destruct Hl; assumption;
 qed.