feat(library/data/{finset,list}/bigop.lean: generalize bigops from group to monoid
This commit is contained in:
parent
42f2fc973a
commit
6fce01385c
5 changed files with 7 additions and 9 deletions
|
@ -1,18 +1,16 @@
|
|||
/-
|
||||
Copyright (c) 2015 Microsoft Corporation. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: data.finset
|
||||
Author: Leonardo de Moura
|
||||
|
||||
Big operator for finite sets
|
||||
Big operator for finite sets.
|
||||
-/
|
||||
import algebra.group data.finset.basic data.list.bigop
|
||||
open algebra finset function binary quot subtype
|
||||
|
||||
namespace finset
|
||||
variables {A B : Type}
|
||||
variable [g : comm_group B]
|
||||
variable [g : comm_monoid B]
|
||||
include g
|
||||
|
||||
definition bigop (s : finset A) (f : A → B) : B :=
|
||||
|
|
|
@ -153,7 +153,7 @@ variables {A B : Type}
|
|||
definition product (s₁ : finset A) (s₂ : finset B) : finset (A × B) :=
|
||||
quot.lift_on₂ s₁ s₂
|
||||
(λ l₁ l₂,
|
||||
to_finset_of_nodup (list.product (elt_of l₁) (elt_of l₂))
|
||||
to_finset_of_nodup (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₂))
|
||||
|
||||
|
|
|
@ -12,7 +12,7 @@ open algebra function binary quot
|
|||
|
||||
namespace list
|
||||
variables {A B : Type}
|
||||
variable [g : group B]
|
||||
variable [g : monoid B]
|
||||
include g
|
||||
|
||||
definition mulf (f : A → B) : B → A → B :=
|
||||
|
@ -58,7 +58,7 @@ end list
|
|||
namespace list
|
||||
open perm
|
||||
variables {A B : Type}
|
||||
variable [g : comm_group B]
|
||||
variable [g : comm_monoid B]
|
||||
include g
|
||||
|
||||
theorem mulf_rcomm (f : A → B) : right_commutative (mulf f) :=
|
||||
|
|
|
@ -268,7 +268,7 @@ theorem zip_unzip : ∀ (l : list (A × B)), zip (pr₁ (unzip l)) (pr₂ (unzip
|
|||
definition flat (l : list (list A)) : list A :=
|
||||
foldl append nil l
|
||||
|
||||
/- cross product -/
|
||||
/- product -/
|
||||
section product
|
||||
|
||||
definition product : list A → list B → list (A × B)
|
||||
|
|
|
@ -130,7 +130,7 @@ end erase
|
|||
section disjoint
|
||||
variable {A : Type}
|
||||
|
||||
definition disjoint (l₁ l₂ : list A) : Prop := ∀ a, (a ∈ l₁ → a ∈ l₂ → false)
|
||||
definition disjoint (l₁ l₂ : list A) : Prop := ∀ ⦃a⦄, (a ∈ l₁ → a ∈ l₂ → false)
|
||||
|
||||
lemma disjoint_left {l₁ l₂ : list A} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₁ → a ∉ l₂ :=
|
||||
λ d a, d a
|
||||
|
|
Loading…
Reference in a new issue