X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdecidable_kit%2Flist_aux.ma;h=a5bf2edb7588ca139325fa040fb2bf75ac148f3c;hb=052140f5d87dabd49f798b5cf42e35e1df411db3;hp=a38d5eb83cc499063b1349003b2904e911925b10;hpb=3c4660604bb23d862c72f64e6b5b5974260eea28;p=helm.git diff --git a/helm/software/matita/library/decidable_kit/list_aux.ma b/helm/software/matita/library/decidable_kit/list_aux.ma index a38d5eb83..a5bf2edb7 100644 --- a/helm/software/matita/library/decidable_kit/list_aux.ma +++ b/helm/software/matita/library/decidable_kit/list_aux.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/decidable_kit/list_aux/". - include "list/list.ma". include "decidable_kit/eqtype.ma". include "nat/plus.ma". @@ -58,8 +56,7 @@ cases c; intros (H); [ apply reflect_true | apply reflect_false ] apply (list_ind2 ? ? ? ? ? Hl); [1: intros; reflexivity] simplify; intros (tl1 tl2 hd1 hd2 IH H); cases (b2pT ? ? (andbP ? ?) H); 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); +| elim l1 in H l2 ⊢ % 1 (l1 x1); [ cases l1; simplify; [intros; destruct H | unfold Not; intros; destruct H1;] | intros 3 (tl1 IH l2); cases l2; [ unfold Not; intros; destruct H1; @@ -71,4 +68,4 @@ cases c; intros (H); [ apply reflect_true | apply reflect_false ] apply H'; reflexivity;]]]] qed. -definition list_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (lcmpP d). \ No newline at end of file +definition list_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (lcmpP d).