X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Ffinset.ma;h=df8d0a87c66ae6783735f95a0c1b4791cd3f76f0;hb=a04fa03fcea0493e89b725960146cc0c06539583;hp=d72388dfb19953c525c6f43256f0796e44c8eec4;hpb=a15a68bdd7d5337f2d6f7573dbdb651c5d278cc4;p=helm.git diff --git a/matita/matita/lib/basics/finset.ma b/matita/matita/lib/basics/finset.ma index d72388dfb..df8d0a87c 100644 --- a/matita/matita/lib/basics/finset.ma +++ b/matita/matita/lib/basics/finset.ma @@ -64,16 +64,12 @@ lemma memb_enumn: ∀m,n,i:DeqNat. ∀h:n ≤ m. ∀k: i < m. n ≤ i → (¬ (memb (Nat_to m) (mk_Sig ?? i k) (enumnaux n m h))) = true. #m #n elim n -n // #n #Hind #i #ltm #k #ltni @sym_eq @noteq_to_eqnot @sym_not_eq % #H cases (orb_true_l … H) - [#H1 @(absurd … (\P H1)) % #Hfalse - cut (∀A,P,a,a1,h,h1.mk_Sig A P a h = mk_Sig A P a1 h1 → a = a1) - [#A #P #a #a1 #h #h1 #H destruct (H) %] #Hcut - lapply (Hcut nat (λi.iHfalse @le_n - |<(notb_notb (memb …)) >Hind normalize /2/ + [whd in ⊢ (??%?→?); #H1 @(absurd … ltni) @le_to_not_lt + >(eqb_true_to_eq … H1) @le_n + |<(notb_notb (memb …)) >Hind normalize /2 by lt_to_le, absurd/ ] qed. - lemma enumn_unique_aux: ∀n,m. ∀h:n ≤ m. uniqueb (Nat_to m) (enumnaux n m h) = true. #n elim n -n // #n #Hind #m #h @true_to_andb_true // @memb_enumn // qed. @@ -209,10 +205,28 @@ unification hint 0 ≔ C1,C2; definition enum_prod ≝ λA,B:DeqSet.λl1.λl2. compose ??? (mk_Prod A B) l1 l2. - -axiom enum_prod_unique: ∀A,B,l1,l2. + +lemma enum_prod_unique: ∀A,B,l1,l2. uniqueb A l1 = true → uniqueb B l2 = true → uniqueb ? (enum_prod A B l1 l2) = true. +#A #B #l1 elim l1 // + #a #tl #Hind #l2 #H1 #H2 @uniqueb_append + [@unique_map_inj [#x #y #Heq @(eq_f … \snd … Heq) | //] + |@Hind // @(andb_true_r … H1) + |#p #H3 cases (memb_map_to_exists … H3) #b * + #Hmemb #eqp (\b (refl ? a)) // + |@orb_true_r2 @Hind2 @H4 + ] + ] + ] + ] +qed. lemma enum_prod_complete:∀A,B:DeqSet.∀l1,l2. (∀a. memb A a l1 = true) → (∀b.memb B b l2 = true) →