theorem(library/data/countable): prove that the product of two countable types is countable

This commit is contained in:
Leonardo de Moura 2015-04-14 20:39:58 -07:00
parent 38b880b939
commit a7db8a2bac
2 changed files with 38 additions and 2 deletions

View file

@ -73,5 +73,41 @@ countable.mk
(λ s, pickle_sum s) (λ s, pickle_sum s)
(λ n, unpickle_sum n) (λ n, unpickle_sum n)
(λ s, unpickle_pickle_sum s) (λ s, unpickle_pickle_sum s)
end sum end sum
section prod
variables {A B : Type}
variables [h₁ : countable A] [h₂ : countable B]
include h₁ h₂
definition pickle_prod : A × B → nat
| (a, b) := mkpair (pickle a) (pickle b)
definition unpickle_prod (n : nat) : option (A × B) :=
match unpair n with
| (n₁, n₂) :=
match unpickle A n₁ with
| some a :=
match unpickle B n₂ with
| some b := some (a, b)
| none := none
end
| none := none
end
end
theorem unpickle_pickle_prod : ∀ p : A × B, unpickle_prod (pickle_prod p) = some p
| (a, b) :=
begin
esimp [pickle_prod, unpickle_prod, prod.cases_on],
rewrite [unpair_mkpair],
esimp,
rewrite [*countable.picklek]
end
definition countable_product [instance] {A B : Type} [h₁ : countable A] [h₂ : countable B] : countable (A × B) :=
countable.mk
pickle_prod
unpickle_prod
unpickle_pickle_prod
end prod

View file

@ -3,4 +3,4 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Author: Jeremy Avigad Author: Jeremy Avigad
-/ -/
import .basic .order .sub .div .bquant .sqrt import .basic .order .sub .div .bquant .sqrt .pairing .power