X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdecidable_kit%2Ffgraph.ma;h=5b28071dcff455d2cc37d02a3f603861d05842e8;hb=063523ae5f8da7e6458232f4afb6744ec86dc8bd;hp=a024ada3d031b259fc823128bfbbaa16150067f8;hpb=32d8d8d419e0b910435da275361bb55d49bc43a9;p=helm.git diff --git a/helm/software/matita/library/decidable_kit/fgraph.ma b/helm/software/matita/library/decidable_kit/fgraph.ma index a024ada3d..5b28071dc 100644 --- a/helm/software/matita/library/decidable_kit/fgraph.ma +++ b/helm/software/matita/library/decidable_kit/fgraph.ma @@ -35,7 +35,7 @@ apply prove_reflect; intros; [1: generalize in match H; rewrite > (b2pT ? ? (eqP (list_eqType d2) ? ?) H2); intros; clear H H2; rewrite < (pirrel ? ? ? H1 H3 (eqType_decidable nat_eqType)); reflexivity - |2: unfold Not; intros (H3); destruct H3; rewrite > Hcut in H2; + |2: unfold Not; intros (H3); destruct H3; rewrite > (cmp_refl (list_eqType d2)) in H2; destruct H2;] qed. @@ -140,7 +140,7 @@ definition mkpermr := lemma mem_multes : ∀d:finType.∀e:d.∀l:list_eqType (list_eqType d).∀s:list_eqType d. mem ? s (multes ? e l) = - match s in list with [ Nil ⇒ false | (Cons e1 tl) ⇒ andb (cmp ? e e1) (mem ? tl l)]. + match s in list with [ nil ⇒ false | (cons e1 tl) ⇒ andb (cmp ? e e1) (mem ? tl l)]. intros (d e l s); elim l (hd tl IH); [cases s;simplify;[2: rewrite > andbC] reflexivity] simplify; rewrite > IH; cases s; simplify; [reflexivity] unfold list_eqType; simplify; @@ -160,7 +160,7 @@ lemma orb_refl : ∀x.orb x x = x. intros (x); cases x; reflexivity; qed. lemma mem_multss : ∀d:finType.∀s:list_eqType d.∀l:list_eqType (list_eqType d).∀x:list_eqType d. mem ? x (multss ? s l) = - match x in list with [ Nil ⇒ false | (Cons he tl) ⇒ andb (mem ? he s) (mem ? tl l)]. + match x in list with [ nil ⇒ false | (cons he tl) ⇒ andb (mem ? he s) (mem ? tl l)]. intros (d s l x); elim s; [cases x] simplify; try reflexivity; rewrite > mem_concat; rewrite > H; clear H; rewrite > mem_multes; cases x; simplify; [reflexivity] unfold list_eqType; simplify; apply (cmpP ? t s1); intros (E);