(****** DeqSet: a set with a decidbale equality ******)
-record DeqSet : Type[1] ≝ { carr :> Type[0];
+record DeqSet : Type[1] ≝ {
+ carr :> Type[0];
eqb: carr → carr → bool;
eqb_true: ∀x,y. (eqb x y = true) ↔ (x = y)
}.
-
+
notation "a == b" non associative with precedence 45 for @{ 'eqb $a $b }.
interpretation "eqb" 'eqb a b = (eqb ? a b).
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 //
+ |@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) →
]
]
qed.
+
(**************************** iterators ******************************)
let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝
#H >H normalize [@Hind //] >H normalize @Hind //
qed.
+lemma uniqueb_append: ∀A,l1,l2. uniqueb A l1 = true → uniqueb A l2 =true →
+ (∀a. memb A a l1 =true → ¬ memb A a l2 =true) → uniqueb A (l1@l2) = true.
+#A #l1 elim l1 [normalize //] #a #tl #Hind #l2 #Hatl #Hl2
+#Hmem normalize cut (memb A a (tl@l2)=false)
+ [2:#Hcut >Hcut normalize @Hind //
+ [@(andb_true_r … Hatl) |#x #Hmemx @Hmem @orb_true_r2 //]
+ |@(noteq_to_eqnot ? true) % #Happend cases (memb_append … Happend)
+ [#H1 @(absurd … H1) @sym_not_eq @eqnot_to_noteq
+ @sym_eq @(andb_true_l … Hatl)
+ |#H @(absurd … H) @Hmem normalize >(\b (refl ? a)) //
+ ]
+ ]
+qed.
+
+lemma memb_map_to_exists: ∀A,B:DeqSet.∀f:A→B.∀l,b.
+ memb ? b (map ?? f l) = true → ∃a. memb ? a l = true ∧ f a = b.
+#A #B #f #l elim l
+ [#b normalize #H destruct (H)
+ |#a #tl #Hind #b #H cases (orb_true_l … H)
+ [#eqb @(ex_intro … a) <(\P eqb) % //
+ |#memb cases (Hind … memb) #a * #mema #eqb
+ @(ex_intro … a) /3/
+ ]
+ ]
+qed.
+
lemma memb_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l,a. injective A B f →
memb ? (f a) (map ?? f l) = true → memb ? a l = true.
#A #B #f #l #a #injf elim l
\ / GNU General Public License Version 2
V_____________________________________________________________*)
-include "basics/finset.ma".
+include "basics/lists/list.ma".
record Vector (A:Type[0]) (n:nat): Type[0] ≝
{ vec :> list A;