refactor(library/data/{finset,list,fintype}: rename cross_product to product

This commit is contained in:
Jeremy Avigad 2015-05-08 13:46:14 +10:00
parent 68f7afa053
commit 42f2fc973a
5 changed files with 51 additions and 54 deletions

View file

@ -148,38 +148,38 @@ theorem all_inter_of_all_right {p : A → Prop} {s₁ : finset A} (s₂ : finset
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ h, list.all_inter_of_all_right _ h)
end all
section cross_product
section product
variables {A B : Type}
definition cross_product (s₁ : finset A) (s₂ : finset B) : finset (A × B) :=
definition product (s₁ : finset A) (s₂ : finset B) : finset (A × B) :=
quot.lift_on₂ s₁ s₂
(λ l₁ l₂,
to_finset_of_nodup (list.cross_product (elt_of l₁) (elt_of l₂))
(nodup_cross_product (has_property l₁) (has_property l₂)))
(λ v₁ v₂ w₁ w₂ p₁ p₂, quot.sound (perm_cross_product p₁ p₂))
to_finset_of_nodup (list.product (elt_of l₁) (elt_of l₂))
(nodup_product (has_property l₁) (has_property l₂)))
(λ v₁ v₂ w₁ w₂ p₁ p₂, quot.sound (perm_product p₁ p₂))
infix * := cross_product
infix * := product
theorem empty_cross_product (s : finset B) : @empty A * s = ∅ :=
theorem empty_product (s : finset B) : @empty A * s = ∅ :=
quot.induction_on s (λ l, rfl)
theorem mem_cross_product {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
theorem mem_product {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
: a ∈ s₁ → b ∈ s₂ → (a, b) ∈ s₁ * s₂ :=
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i₁ i₂, list.mem_cross_product i₁ i₂)
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i₁ i₂, list.mem_product i₁ i₂)
theorem mem_of_mem_cross_product_left {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
theorem mem_of_mem_product_left {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
: (a, b) ∈ s₁ * s₂ → a ∈ s₁ :=
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i, list.mem_of_mem_cross_product_left i)
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i, list.mem_of_mem_product_left i)
theorem mem_of_mem_cross_product_right {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
theorem mem_of_mem_product_right {a : A} {b : B} {s₁ : finset A} {s₂ : finset B}
: (a, b) ∈ s₁ * s₂ → b ∈ s₂ :=
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i, list.mem_of_mem_cross_product_right i)
quot.induction_on₂ s₁ s₂ (λ l₁ l₂ i, list.mem_of_mem_product_right i)
theorem cross_product_empty (s : finset A) : s * @empty B = ∅ :=
theorem product_empty (s : finset A) : s * @empty B = ∅ :=
ext (λ p,
match p with
| (a, b) := iff.intro
(λ i, absurd (mem_of_mem_cross_product_right i) !not_mem_empty)
(λ i, absurd (mem_of_mem_product_right i) !not_mem_empty)
(λ i, absurd i !not_mem_empty)
end)
end cross_product
end product
end finset

View file

@ -1,10 +1,9 @@
/-
Copyright (c) 2015 Leonardo de Moura. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
Finite type (type class)
Finite type (type class).
-/
import data.list data.bool
open list bool unit decidable option function
@ -26,11 +25,11 @@ fintype.mk [ff, tt]
definition fintype_product [instance] {A B : Type} : Π [h₁ : fintype A] [h₂ : fintype B], fintype (A × B)
| (fintype.mk e₁ u₁ c₁) (fintype.mk e₂ u₂ c₂) :=
fintype.mk
(cross_product e₁ e₂)
(nodup_cross_product u₁ u₂)
(product e₁ e₂)
(nodup_product u₁ u₂)
(λ p,
match p with
(a, b) := mem_cross_product (c₁ a) (c₂ b)
(a, b) := mem_product (c₁ a) (c₂ b)
end)
/- auxiliary function for finding 'a' s.t. f a ≠ g a -/

View file

@ -269,20 +269,20 @@ definition flat (l : list (list A)) : list A :=
foldl append nil l
/- cross product -/
section cross_product
section product
definition cross_product : list A → list B → list (A × B)
definition product : list A → list B → list (A × B)
| [] l₂ := []
| (a::l₁) l₂ := map (λ b, (a, b)) l₂ ++ cross_product l₁ l₂
| (a::l₁) l₂ := map (λ b, (a, b)) l₂ ++ product l₁ l₂
theorem nil_cross_product (l : list B) : cross_product (@nil A) l = []
theorem nil_product (l : list B) : product (@nil A) l = []
theorem cross_product_cons (a : A) (l₁ : list A) (l₂ : list B)
: cross_product (a::l₁) l₂ = map (λ b, (a, b)) l₂ ++ cross_product l₁ l₂
theorem product_cons (a : A) (l₁ : list A) (l₂ : list B)
: product (a::l₁) l₂ = map (λ b, (a, b)) l₂ ++ product l₁ l₂
theorem cross_product_nil : ∀ (l : list A), cross_product l (@nil B) = []
theorem product_nil : ∀ (l : list A), product l (@nil B) = []
| [] := rfl
| (a::l) := by rewrite [cross_product_cons, map_nil, cross_product_nil]
| (a::l) := by rewrite [product_cons, map_nil, product_nil]
theorem eq_of_mem_map_pair₁ {a₁ a : A} {b₁ : B} {l : list B} : (a₁, b₁) ∈ map (λ b, (a, b)) l → a₁ = a :=
assume ain,
@ -296,7 +296,7 @@ assert h₁ : pr2 (a₁, b₁) ∈ map pr2 (map (λ b, (a, b)) l), from mem_map
assert h₂ : b₁ ∈ map (λx, x) l, by rewrite [map_map at h₁, ↑pr2 at h₁]; exact h₁,
by rewrite [map_id at h₂]; exact h₂
theorem mem_cross_product {a : A} {b : B} : ∀ {l₁ l₂}, a ∈ l₁ → b ∈ l₂ → (a, b) ∈ cross_product l₁ l₂
theorem mem_product {a : A} {b : B} : ∀ {l₁ l₂}, a ∈ l₁ → b ∈ l₂ → (a, b) ∈ product l₁ l₂
| [] l₂ h₁ h₂ := absurd h₁ !not_mem_nil
| (x::l₁) l₂ h₁ h₂ :=
or.elim (eq_or_mem_of_mem_cons h₁)
@ -304,29 +304,29 @@ theorem mem_cross_product {a : A} {b : B} : ∀ {l₁ l₂}, a ∈ l₁ → b
assert aux : (a, b) ∈ map (λ b, (a, b)) l₂, from mem_map _ h₂,
by rewrite [-aeqx]; exact (mem_append_left _ aux))
(λ ainl₁ : a ∈ l₁,
have inl₁l₂ : (a, b) ∈ cross_product l₁ l₂, from mem_cross_product ainl₁ h₂,
have inl₁l₂ : (a, b) ∈ product l₁ l₂, from mem_product ainl₁ h₂,
mem_append_right _ inl₁l₂)
theorem mem_of_mem_cross_product_left {a : A} {b : B} : ∀ {l₁ l₂}, (a, b) ∈ cross_product l₁ l₂ → a ∈ l₁
theorem mem_of_mem_product_left {a : A} {b : B} : ∀ {l₁ l₂}, (a, b) ∈ product l₁ l₂ → a ∈ l₁
| [] l₂ h := absurd h !not_mem_nil
| (x::l₁) l₂ h :=
or.elim (mem_or_mem_of_mem_append h)
(λ ain : (a, b) ∈ map (λ b, (x, b)) l₂,
assert aeqx : a = x, from eq_of_mem_map_pair₁ ain,
by rewrite [aeqx]; exact !mem_cons)
(λ ain : (a, b) ∈ cross_product l₁ l₂,
have ainl₁ : a ∈ l₁, from mem_of_mem_cross_product_left ain,
(λ ain : (a, b) ∈ product l₁ l₂,
have ainl₁ : a ∈ l₁, from mem_of_mem_product_left ain,
mem_cons_of_mem _ ainl₁)
theorem mem_of_mem_cross_product_right {a : A} {b : B} : ∀ {l₁ l₂}, (a, b) ∈ cross_product l₁ l₂ → b ∈ l₂
theorem mem_of_mem_product_right {a : A} {b : B} : ∀ {l₁ l₂}, (a, b) ∈ product l₁ l₂ → b ∈ l₂
| [] l₂ h := absurd h !not_mem_nil
| (x::l₁) l₂ h :=
or.elim (mem_or_mem_of_mem_append h)
(λ abin : (a, b) ∈ map (λ b, (x, b)) l₂,
mem_of_mem_map_pair₁ abin)
(λ abin : (a, b) ∈ cross_product l₁ l₂,
mem_of_mem_cross_product_right abin)
end cross_product
(λ abin : (a, b) ∈ product l₁ l₂,
mem_of_mem_product_right abin)
end product
end list
attribute list.decidable_any [instance]

View file

@ -726,30 +726,30 @@ theorem perm_ext : ∀ {l₁ l₂ : list A}, nodup l₁ → nodup l₂ → (∀a
... = a₂::t₂ : by rewrite t₂_eq
end ext
/- cross_product -/
section cross_product
theorem perm_cross_product_left {l₁ l₂ : list A} (t₁ : list B) : l₁ ~ l₂ → (cross_product l₁ t₁) ~ (cross_product l₂ t₁) :=
/- product -/
section product
theorem perm_product_left {l₁ l₂ : list A} (t₁ : list B) : l₁ ~ l₂ → (product l₁ t₁) ~ (product l₂ t₁) :=
assume p : l₁ ~ l₂, perm.induction_on p
!perm.refl
(λ x l₁ l₂ p r, perm_app !perm.refl r)
(λ x y l,
let m₁ := map (λ b, (x, b)) t₁ in
let m₂ := map (λ b, (y, b)) t₁ in
let c := cross_product l t₁ in
let c := product l t₁ in
calc m₂ ++ (m₁ ++ c) = (m₂ ++ m₁) ++ c : by rewrite append.assoc
... ~ (m₁ ++ m₂) ++ c : perm_app !perm_app_comm !perm.refl
... = m₁ ++ (m₂ ++ c) : by rewrite append.assoc)
(λ l₁ l₂ l₃ p₁ p₂ r₁ r₂, trans r₁ r₂)
theorem perm_cross_product_right (l : list A) {t₁ t₂ : list B} : t₁ ~ t₂ → (cross_product l t₁) ~ (cross_product l t₂) :=
theorem perm_product_right (l : list A) {t₁ t₂ : list B} : t₁ ~ t₂ → (product l t₁) ~ (product l t₂) :=
list.induction_on l
(λ p, by rewrite [*nil_cross_product])
(λ p, by rewrite [*nil_product])
(λ a t r p,
perm_app (perm_map _ p) (r p))
theorem perm_cross_product {l₁ l₂ : list A} {t₁ t₂ : list B} : l₁ ~ l₂ → t₁ ~ t₂ → (cross_product l₁ t₁) ~ (cross_product l₂ t₂) :=
assume p₁ p₂, trans (perm_cross_product_left t₁ p₁) (perm_cross_product_right l₂ p₂)
end cross_product
theorem perm_product {l₁ l₂ : list A} {t₁ t₂ : list B} : l₁ ~ l₂ → t₁ ~ t₂ → (product l₁ t₁) ~ (product l₂ t₂) :=
assume p₁ p₂, trans (perm_product_left t₁ p₁) (perm_product_right l₂ p₂)
end product
/- filter -/
theorem perm_filter {l₁ l₂ : list A} {p : A → Prop} [decp : decidable_pred p] :

View file

@ -1,8 +1,6 @@
/-
Copyright (c) 2015 Leonardo de Moura. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Module: data.list.set
Authors: Leonardo de Moura
Set-like operations on lists
@ -396,12 +394,12 @@ definition decidable_nodup [instance] [h : decidable_eq A] : ∀ (l : list A), d
end
end
theorem nodup_cross_product : ∀ {l₁ : list A} {l₂ : list B}, nodup l₁ → nodup l₂ → nodup (cross_product l₁ l₂)
theorem nodup_product : ∀ {l₁ : list A} {l₂ : list B}, nodup l₁ → nodup l₂ → nodup (product l₁ l₂)
| [] l₂ n₁ n₂ := nodup_nil
| (a::l₁) l₂ n₁ n₂ :=
have nainl₁ : a ∉ l₁, from not_mem_of_nodup_cons n₁,
have n₃ : nodup l₁, from nodup_of_nodup_cons n₁,
have n₄ : nodup (cross_product l₁ l₂), from nodup_cross_product n₃ n₂,
have n₄ : nodup (product l₁ l₂), from nodup_product n₃ n₂,
have dgen : ∀ l, nodup l → nodup (map (λ b, (a, b)) l)
| [] h := nodup_nil
| (x::l) h :=
@ -412,11 +410,11 @@ theorem nodup_cross_product : ∀ {l₁ : list A} {l₂ : list B}, nodup l₁
assume pin, absurd (mem_of_mem_map_pair₁ pin) nxin,
nodup_cons npin dm,
have dm : nodup (map (λ b, (a, b)) l₂), from dgen l₂ n₂,
have dsj : disjoint (map (λ b, (a, b)) l₂) (cross_product l₁ l₂), from
have dsj : disjoint (map (λ b, (a, b)) l₂) (product l₁ l₂), from
λ p, match p with
| (a₁, b₁) :=
λ (i₁ : (a₁, b₁) ∈ map (λ b, (a, b)) l₂) (i₂ : (a₁, b₁) ∈ cross_product l₁ l₂),
have a₁inl₁ : a₁ ∈ l₁, from mem_of_mem_cross_product_left i₂,
λ (i₁ : (a₁, b₁) ∈ map (λ b, (a, b)) l₂) (i₂ : (a₁, b₁) ∈ product l₁ l₂),
have a₁inl₁ : a₁ ∈ l₁, from mem_of_mem_product_left i₂,
have a₁eqa : a₁ = a, from eq_of_mem_map_pair₁ i₁,
absurd (a₁eqa ▸ a₁inl₁) nainl₁
end,