feat(library/data/list/perm): add perm_cross_product theorem
This commit is contained in:
parent
41ddc97e0d
commit
4c827293a8
2 changed files with 26 additions and 1 deletions
|
@ -275,7 +275,7 @@ definition cross_product : list A → list B → list (A × B)
|
||||||
| [] l₂ := []
|
| [] l₂ := []
|
||||||
| (a::l₁) l₂ := map (λ b, (a, b)) l₂ ++ cross_product l₁ l₂
|
| (a::l₁) l₂ := map (λ b, (a, b)) l₂ ++ cross_product l₁ l₂
|
||||||
|
|
||||||
theorem nil_cross_product_nil (l : list B) : cross_product (@nil A) l = []
|
theorem nil_cross_product (l : list B) : cross_product (@nil A) l = []
|
||||||
|
|
||||||
theorem cross_product_cons (a : A) (l₁ : list A) (l₂ : list B)
|
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₂
|
: cross_product (a::l₁) l₂ = map (λ b, (a, b)) l₂ ++ cross_product l₁ l₂
|
||||||
|
|
|
@ -726,4 +726,29 @@ theorem perm_ext : ∀ {l₁ l₂ : list A}, nodup l₁ → nodup l₂ → (∀a
|
||||||
... ~ s₁++(a₁::s₂) : !perm_middle
|
... ~ s₁++(a₁::s₂) : !perm_middle
|
||||||
... = a₂::t₂ : by rewrite t₂_eq
|
... = a₂::t₂ : by rewrite t₂_eq
|
||||||
end ext
|
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₁) :=
|
||||||
|
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
|
||||||
|
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₂) :=
|
||||||
|
list.induction_on l
|
||||||
|
(λ p, by rewrite [*nil_cross_product]; exact !perm.refl)
|
||||||
|
(λ a t r p,
|
||||||
|
perm_app (perm_map _ p) (r p))
|
||||||
|
|
||||||
|
theorem perm_cross_product {l₁ l₂ t₁ t₂ : list A} : 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
|
||||||
end perm
|
end perm
|
||||||
|
|
Loading…
Add table
Reference in a new issue