]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/basics/finset.ma
lfsx_lfpxs completed!
[helm.git] / matita / matita / lib / basics / finset.ma
index babbf997c84f837a40d45d5a4481eea620418122..df8d0a87c66ae6783735f95a0c1b4791cd3f76f0 100644 (file)
@@ -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.i<m) i n ? ? Hfalse) #Hfalse @(absurd … ltni)
-    @le_to_not_lt >Hfalse @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 <eqp @(not_to_not ? (memb ? a tl = true))
+    [2: @sym_not_eq @eqnot_to_noteq @sym_eq @(andb_true_l … H1)
+    |elim tl 
+      [normalize //
+      |#a1 #tl1 #Hind2 #H4 cases (memb_append … H4) -H4 #H4
+        [cases (memb_map_to_exists … H4) #b1 * #memb1 #H destruct (H)
+         normalize >(\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) →
@@ -233,3 +247,26 @@ unification hint  0 ≔ C1,C2;
 (* ---------------------------------------- *) ⊢ 
     T1×T2 ≡ FinSetcarr X.
 
+(* graph of a function *)
+
+definition graph_of ≝ λA,B.λf:A→B. 
+  Σp:A×B.f (\fst p) = \snd p.
+
+definition graph_enum ≝ λA,B:FinSet.λf:A→B. 
+  filter ? (λp.f (\fst p) == \snd p) (enum (FinProd A B)).
+  
+lemma graph_enum_unique : ∀A,B,f.
+  uniqueb ? (graph_enum A B f) = true.
+#A #B #f @uniqueb_filter @(enum_unique (FinProd A B))
+qed.
+
+lemma graph_enum_correct: ∀A,B:FinSet.∀f:A→B. ∀a,b.
+memb ? 〈a,b〉 (graph_enum A B f) = true → f a = b.
+#A #B #f #a #b #membp @(\P ?) @(filter_true … membp)
+qed.
+
+lemma graph_enum_complete: ∀A,B:FinSet.∀f:A→B. ∀a,b.
+f a = b → memb ? 〈a,b〉 (graph_enum A B f) = true.
+#A #B #f #a #b #eqf @memb_filter_l [@(\b eqf)]
+@enum_prod_complete //
+qed.